src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML
author blanchet
Thu, 25 Nov 2010 14:13:48 +0100
changeset 40694 77435a7853d1
parent 40670 c059d550afec
child 40941 a3e6f8634a11
permissions -rw-r--r--
reverted c059d550afec -- the triviality check had apparently nothing to do with spontaneous Interrupt exceptions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32564
378528d2f7eb standard headers and text sections;
wenzelm
parents: 32551
diff changeset
     1
(*  Title:      HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML
378528d2f7eb standard headers and text sections;
wenzelm
parents: 32551
diff changeset
     2
    Author:     Jasmin Blanchette and Sascha Boehme and Tobias Nipkow, TU Munich
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
     3
*)
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
     4
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
     5
structure Mirabelle_Sledgehammer : MIRABELLE_ACTION =
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
     6
struct
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
     7
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
     8
val proverK = "prover"
32541
cea1716eb106 timeout option for ATPs
boehmes
parents: 32536
diff changeset
     9
val prover_timeoutK = "prover_timeout"
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    10
val keepK = "keep"
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    11
val full_typesK = "full_types"
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    12
val minimizeK = "minimize"
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    13
val minimize_timeoutK = "minimize_timeout"
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    14
val metis_ftK = "metis_ft"
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    15
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    16
fun sh_tag id = "#" ^ string_of_int id ^ " sledgehammer: "
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    17
fun minimize_tag id = "#" ^ string_of_int id ^ " minimize (sledgehammer): "
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    18
fun reconstructor_tag reconstructor id =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    19
  "#" ^ string_of_int id ^ " " ^ (!reconstructor) ^ " (sledgehammer): "
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    20
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    21
val separator = "-----"
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    22
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    23
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    24
datatype sh_data = ShData of {
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    25
  calls: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    26
  success: int,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    27
  nontriv_calls: int,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    28
  nontriv_success: int,
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
    29
  lemmas: int,
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    30
  max_lems: int,
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    31
  time_isa: int,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    32
  time_prover: int,
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    33
  time_prover_fail: int}
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    34
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    35
datatype re_data = ReData of {
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    36
  calls: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    37
  success: int,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    38
  nontriv_calls: int,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    39
  nontriv_success: int,
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
    40
  proofs: int,
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    41
  time: int,
32550
d57c7a2d927c logging number of metis lemmas
nipkow
parents: 32549
diff changeset
    42
  timeout: int,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    43
  lemmas: int * int * int,
39341
d2b981a0429a indicate triviality in the list of proved things
blanchet
parents: 39340
diff changeset
    44
  posns: (Position.T * bool) list
32550
d57c7a2d927c logging number of metis lemmas
nipkow
parents: 32549
diff changeset
    45
  }
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    46
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    47
datatype min_data = MinData of {
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
    48
  succs: int,
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
    49
  ab_ratios: int
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    50
  }
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    51
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    52
fun make_sh_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    53
      (calls,success,nontriv_calls,nontriv_success,lemmas,max_lems,time_isa,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    54
       time_prover,time_prover_fail) =
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    55
  ShData{calls=calls, success=success, nontriv_calls=nontriv_calls,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    56
         nontriv_success=nontriv_success, lemmas=lemmas, max_lems=max_lems,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    57
         time_isa=time_isa, time_prover=time_prover,
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    58
         time_prover_fail=time_prover_fail}
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    59
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
    60
fun make_min_data (succs, ab_ratios) =
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
    61
  MinData{succs=succs, ab_ratios=ab_ratios}
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    62
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    63
fun make_re_data (calls,success,nontriv_calls,nontriv_success,proofs,time,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    64
                  timeout,lemmas,posns) =
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    65
  ReData{calls=calls, success=success, nontriv_calls=nontriv_calls,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    66
         nontriv_success=nontriv_success, proofs=proofs, time=time,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    67
         timeout=timeout, lemmas=lemmas, posns=posns}
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    68
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    69
val empty_sh_data = make_sh_data (0, 0, 0, 0, 0, 0, 0, 0, 0)
35871
c93bda4fdf15 remove the iteration counter from Sledgehammer's minimizer
blanchet
parents: 35867
diff changeset
    70
val empty_min_data = make_min_data (0, 0)
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    71
val empty_re_data = make_re_data (0, 0, 0, 0, 0, 0, 0, (0,0,0), [])
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    72
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    73
fun tuple_of_sh_data (ShData {calls, success, nontriv_calls, nontriv_success,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    74
                              lemmas, max_lems, time_isa,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    75
  time_prover, time_prover_fail}) = (calls, success, nontriv_calls,
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    76
  nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    77
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
    78
fun tuple_of_min_data (MinData {succs, ab_ratios}) = (succs, ab_ratios)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    79
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    80
fun tuple_of_re_data (ReData {calls, success, nontriv_calls, nontriv_success,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    81
  proofs, time, timeout, lemmas, posns}) = (calls, success, nontriv_calls,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    82
  nontriv_success, proofs, time, timeout, lemmas, posns)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    83
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    84
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    85
datatype reconstructor_mode =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    86
  Unminimized | Minimized | UnminimizedFT | MinimizedFT
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    87
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    88
datatype data = Data of {
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    89
  sh: sh_data,
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    90
  min: min_data,
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    91
  re_u: re_data, (* reconstructor with unminimized set of lemmas *)
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    92
  re_m: re_data, (* reconstructor with minimized set of lemmas *)
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    93
  re_uft: re_data, (* reconstructor with unminimized set of lemmas and fully-typed *)
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    94
  re_mft: re_data, (* reconstructor with minimized set of lemmas and fully-typed *)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    95
  mini: bool   (* with minimization *)
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    96
  }
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    97
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    98
fun make_data (sh, min, re_u, re_m, re_uft, re_mft, mini) =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    99
  Data {sh=sh, min=min, re_u=re_u, re_m=re_m, re_uft=re_uft, re_mft=re_mft,
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   100
    mini=mini}
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   101
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   102
val empty_data = make_data (empty_sh_data, empty_min_data,
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   103
  empty_re_data, empty_re_data, empty_re_data, empty_re_data, false)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   104
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   105
fun map_sh_data f (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   106
  let val sh' = make_sh_data (f (tuple_of_sh_data sh))
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   107
  in make_data (sh', min, re_u, re_m, re_uft, re_mft, mini) end
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   108
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   109
fun map_min_data f (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   110
  let val min' = make_min_data (f (tuple_of_min_data min))
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   111
  in make_data (sh, min', re_u, re_m, re_uft, re_mft, mini) end
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   112
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   113
fun map_re_data f m (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   114
  let
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   115
    fun map_me g Unminimized   (u, m, uft, mft) = (g u, m, uft, mft)
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   116
      | map_me g Minimized     (u, m, uft, mft) = (u, g m, uft, mft)
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   117
      | map_me g UnminimizedFT (u, m, uft, mft) = (u, m, g uft, mft)
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   118
      | map_me g MinimizedFT   (u, m, uft, mft) = (u, m, uft, g mft)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   119
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   120
    val f' = make_re_data o f o tuple_of_re_data
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   121
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   122
    val (re_u', re_m', re_uft', re_mft') =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   123
      map_me f' m (re_u, re_m, re_uft, re_mft)
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   124
  in make_data (sh, min, re_u', re_m', re_uft', re_mft', mini) end
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   125
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   126
fun set_mini mini (Data {sh, min, re_u, re_m, re_uft, re_mft, ...}) =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   127
  make_data (sh, min, re_u, re_m, re_uft, re_mft, mini)
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   128
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   129
fun inc_max (n:int) (s,sos,m) = (s+n, sos + n*n, Int.max(m,n));
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   130
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   131
val inc_sh_calls =  map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   132
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   133
    => (calls + 1, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   134
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   135
val inc_sh_success = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   136
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   137
    => (calls, success + 1, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail))
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   138
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   139
val inc_sh_nontriv_calls =  map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   140
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   141
    => (calls, success, nontriv_calls + 1, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail))
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   142
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   143
val inc_sh_nontriv_success = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   144
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   145
    => (calls, success, nontriv_calls, nontriv_success + 1, lemmas,max_lems, time_isa, time_prover, time_prover_fail))
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
   146
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   147
fun inc_sh_lemmas n = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   148
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   149
    => (calls,success,nontriv_calls, nontriv_success, lemmas+n,max_lems,time_isa,time_prover,time_prover_fail))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   150
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   151
fun inc_sh_max_lems n = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   152
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   153
    => (calls,success,nontriv_calls, nontriv_success, lemmas,Int.max(max_lems,n),time_isa,time_prover,time_prover_fail))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   154
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   155
fun inc_sh_time_isa t = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   156
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   157
    => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa + t,time_prover,time_prover_fail))
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   158
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   159
fun inc_sh_time_prover t = map_sh_data
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   160
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   161
    => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover + t,time_prover_fail))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   162
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   163
fun inc_sh_time_prover_fail t = map_sh_data
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   164
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   165
    => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail + t))
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   166
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   167
val inc_min_succs = map_min_data
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
   168
  (fn (succs,ab_ratios) => (succs+1, ab_ratios))
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   169
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   170
fun inc_min_ab_ratios r = map_min_data
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
   171
  (fn (succs, ab_ratios) => (succs, ab_ratios+r))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   172
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   173
val inc_reconstructor_calls = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   174
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   175
    => (calls + 1, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   176
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   177
val inc_reconstructor_success = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   178
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   179
    => (calls, success + 1, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   180
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   181
val inc_reconstructor_nontriv_calls = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   182
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   183
    => (calls, success, nontriv_calls + 1, nontriv_success, proofs, time, timeout, lemmas,posns))
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   184
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   185
val inc_reconstructor_nontriv_success = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   186
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   187
    => (calls, success, nontriv_calls, nontriv_success + 1, proofs, time, timeout, lemmas,posns))
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   188
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   189
val inc_reconstructor_proofs = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   190
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   191
    => (calls, success, nontriv_calls, nontriv_success, proofs + 1, time, timeout, lemmas,posns))
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   192
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   193
fun inc_reconstructor_time m t = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   194
 (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   195
  => (calls, success, nontriv_calls, nontriv_success, proofs, time + t, timeout, lemmas,posns)) m
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   196
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   197
val inc_reconstructor_timeout = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   198
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   199
    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout + 1, lemmas,posns))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   200
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   201
fun inc_reconstructor_lemmas m n = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   202
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   203
    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, inc_max n lemmas, posns)) m
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   204
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   205
fun inc_reconstructor_posns m pos = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   206
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   207
    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas, pos::posns)) m
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   208
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   209
local
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   210
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   211
val str = string_of_int
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   212
val str3 = Real.fmt (StringCvt.FIX (SOME 3))
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   213
fun percentage a b = string_of_int (a * 100 div b)
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   214
fun time t = Real.fromInt t / 1000.0
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   215
fun avg_time t n =
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   216
  if n > 0 then (Real.fromInt t / 1000.0) / Real.fromInt n else 0.0
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   217
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   218
fun log_sh_data log
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   219
    (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail) =
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   220
 (log ("Total number of sledgehammer calls: " ^ str calls);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   221
  log ("Number of successful sledgehammer calls: " ^ str success);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   222
  log ("Number of sledgehammer lemmas: " ^ str lemmas);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   223
  log ("Max number of sledgehammer lemmas: " ^ str max_lems);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   224
  log ("Success rate: " ^ percentage success calls ^ "%");
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   225
  log ("Total number of nontrivial sledgehammer calls: " ^ str nontriv_calls);
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   226
  log ("Number of successful nontrivial sledgehammer calls: " ^ str nontriv_success);
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   227
  log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time time_isa));
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   228
  log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_prover));
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   229
  log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_prover_fail));
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   230
  log ("Average time for sledgehammer calls (Isabelle): " ^
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   231
    str3 (avg_time time_isa calls));
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   232
  log ("Average time for successful sledgehammer calls (ATP): " ^
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   233
    str3 (avg_time time_prover success));
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   234
  log ("Average time for failed sledgehammer calls (ATP): " ^
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   235
    str3 (avg_time time_prover_fail (calls - success)))
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   236
  )
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   237
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   238
39341
d2b981a0429a indicate triviality in the list of proved things
blanchet
parents: 39340
diff changeset
   239
fun str_of_pos (pos, triv) =
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   240
  let val str0 = string_of_int o the_default 0
39341
d2b981a0429a indicate triviality in the list of proved things
blanchet
parents: 39340
diff changeset
   241
  in
d2b981a0429a indicate triviality in the list of proved things
blanchet
parents: 39340
diff changeset
   242
    str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos) ^
d2b981a0429a indicate triviality in the list of proved things
blanchet
parents: 39340
diff changeset
   243
    (if triv then "[T]" else "")
d2b981a0429a indicate triviality in the list of proved things
blanchet
parents: 39340
diff changeset
   244
  end
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   245
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   246
fun log_re_data log tag sh_calls (re_calls, re_success, re_nontriv_calls,
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   247
     re_nontriv_success, re_proofs, re_time, re_timeout,
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   248
    (lemmas, lems_sos, lems_max), re_posns) =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   249
 (log ("Total number of " ^ tag ^ "reconstructor calls: " ^ str re_calls);
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   250
  log ("Number of successful " ^ tag ^ "reconstructor calls: " ^ str re_success ^
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   251
    " (proof: " ^ str re_proofs ^ ")");
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   252
  log ("Number of " ^ tag ^ "reconstructor timeouts: " ^ str re_timeout);
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   253
  log ("Success rate: " ^ percentage re_success sh_calls ^ "%");
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   254
  log ("Total number of nontrivial " ^ tag ^ "reconstructor calls: " ^ str re_nontriv_calls);
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   255
  log ("Number of successful nontrivial " ^ tag ^ "reconstructor calls: " ^ str re_nontriv_success ^
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   256
    " (proof: " ^ str re_proofs ^ ")");
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   257
  log ("Number of successful " ^ tag ^ "reconstructor lemmas: " ^ str lemmas);
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   258
  log ("SOS of successful " ^ tag ^ "reconstructor lemmas: " ^ str lems_sos);
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   259
  log ("Max number of successful " ^ tag ^ "reconstructor lemmas: " ^ str lems_max);
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   260
  log ("Total time for successful " ^ tag ^ "reconstructor calls: " ^ str3 (time re_time));
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   261
  log ("Average time for successful " ^ tag ^ "reconstructor calls: " ^
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   262
    str3 (avg_time re_time re_success));
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   263
  if tag=""
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   264
  then log ("Proved: " ^ space_implode " " (map str_of_pos re_posns))
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   265
  else ()
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   266
 )
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   267
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
   268
fun log_min_data log (succs, ab_ratios) =
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   269
  (log ("Number of successful minimizations: " ^ string_of_int succs);
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
   270
   log ("After/before ratios: " ^ string_of_int ab_ratios)
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   271
  )
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   272
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   273
in
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   274
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   275
fun log_data id log (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   276
  let
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   277
    val ShData {calls=sh_calls, ...} = sh
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   278
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   279
    fun app_if (ReData {calls, ...}) f = if calls > 0 then f () else ()
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   280
    fun log_re tag m =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   281
      log_re_data log tag sh_calls (tuple_of_re_data m)
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   282
    fun log_reconstructor (tag1, m1) (tag2, m2) = app_if m1 (fn () =>
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   283
      (log_re tag1 m1; log ""; app_if m2 (fn () => log_re tag2 m2)))
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   284
  in
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   285
    if sh_calls > 0
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   286
    then
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   287
     (log ("\n\n\nReport #" ^ string_of_int id ^ ":\n");
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   288
      log_sh_data log (tuple_of_sh_data sh);
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   289
      log "";
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   290
      if not mini
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   291
      then log_reconstructor ("", re_u) ("fully-typed ", re_uft)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   292
      else
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   293
        app_if re_u (fn () =>
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   294
         (log_reconstructor ("unminimized ", re_u) ("unminimized fully-typed ", re_uft);
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   295
          log "";
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   296
          app_if re_m (fn () =>
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   297
            (log_min_data log (tuple_of_min_data min); log "";
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   298
             log_reconstructor ("", re_m) ("fully-typed ", re_mft))))))
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   299
    else ()
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   300
  end
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   301
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   302
end
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   303
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   304
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   305
(* Warning: we implicitly assume single-threaded execution here! *)
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   306
val data = Unsynchronized.ref ([] : (int * data) list)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   307
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   308
fun init id thy = (Unsynchronized.change data (cons (id, empty_data)); thy)
32567
de411627a985 explicitly export type abbreviations (as usual in SML97);
wenzelm
parents: 32564
diff changeset
   309
fun done id ({log, ...}: Mirabelle.done_args) =
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   310
  AList.lookup (op =) (!data) id
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   311
  |> Option.map (log_data id log)
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   312
  |> K ()
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   313
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   314
fun change_data id f = (Unsynchronized.change data (AList.map_entry (op =) id f); ())
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   315
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   316
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40065
diff changeset
   317
fun get_prover ctxt args =
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   318
  let
40200
870818d2b56b remove needless context argument;
blanchet
parents: 40184
diff changeset
   319
    val thy = ProofContext.theory_of ctxt
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   320
    fun default_prover_name () =
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40065
diff changeset
   321
      hd (#provers (Sledgehammer_Isar.default_params ctxt []))
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   322
      handle Empty => error "No ATP available."
40200
870818d2b56b remove needless context argument;
blanchet
parents: 40184
diff changeset
   323
    fun get_prover name = (name, Sledgehammer.get_prover thy false name)
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   324
  in
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   325
    (case AList.lookup (op =) args proverK of
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   326
      SOME name => get_prover name
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   327
    | NONE => get_prover (default_prover_name ()))
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   328
  end
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   329
38988
483879af0643 finished renaming
blanchet
parents: 38826
diff changeset
   330
type locality = Sledgehammer_Filter.locality
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38700
diff changeset
   331
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   332
(* hack *)
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   333
fun reconstructor_from_msg msg =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   334
  if String.isSubstring "metis" msg then "metis" else "smt"
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   335
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   336
local
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   337
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   338
datatype sh_result =
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38700
diff changeset
   339
  SH_OK of int * int * (string * locality) list |
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   340
  SH_FAIL of int * int |
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   341
  SH_ERROR
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   342
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   343
fun run_sh prover_name prover hard_timeout timeout dir st =
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   344
  let
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   345
    val {context = ctxt, facts = chained_ths, goal} = Proof.goal st
35971
4f24a4e9af4a make Mirabelle happy again
blanchet
parents: 35871
diff changeset
   346
    val thy = ProofContext.theory_of ctxt
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   347
    val i = 1
39321
23951979a362 make Mirabelle happy
blanchet
parents: 39005
diff changeset
   348
    fun change_dir (SOME dir) = Config.put Sledgehammer.dest_dir dir
23951979a362 make Mirabelle happy
blanchet
parents: 39005
diff changeset
   349
      | change_dir NONE = I
23951979a362 make Mirabelle happy
blanchet
parents: 39005
diff changeset
   350
    val st' =
23951979a362 make Mirabelle happy
blanchet
parents: 39005
diff changeset
   351
      st |> Proof.map_context
23951979a362 make Mirabelle happy
blanchet
parents: 39005
diff changeset
   352
                (change_dir dir
23951979a362 make Mirabelle happy
blanchet
parents: 39005
diff changeset
   353
                 #> Config.put Sledgehammer.measure_run_time true)
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   354
    val params as {full_types, relevance_thresholds, max_relevant, ...} =
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40065
diff changeset
   355
      Sledgehammer_Isar.default_params ctxt
40554
ff446d5e9a62 turn on Sledgehammer verbosity so we can track down crashes
blanchet
parents: 40526
diff changeset
   356
          [("verbose", "true"),
ff446d5e9a62 turn on Sledgehammer verbosity so we can track down crashes
blanchet
parents: 40526
diff changeset
   357
           ("timeout", Int.toString timeout)]
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   358
    val default_max_relevant =
40063
d086e3699e78 bring ATPs and SMT solvers more in line with each other
blanchet
parents: 40062
diff changeset
   359
      Sledgehammer.default_max_relevant_for_prover thy prover_name
40369
53dca3bd4250 use the SMT integration's official list of built-ins
blanchet
parents: 40301
diff changeset
   360
    val is_built_in_const =
53dca3bd4250 use the SMT integration's official list of built-ins
blanchet
parents: 40301
diff changeset
   361
      Sledgehammer.is_built_in_const_for_prover ctxt prover_name
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   362
    val relevance_fudge = Sledgehammer.relevance_fudge_for_prover prover_name
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   363
    val relevance_override = {add = [], del = [], only = false}
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   364
    val (_, hyp_ts, concl_t) = Sledgehammer_Util.strip_subgoal goal i
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
   365
    val facts =
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   366
      Sledgehammer_Filter.relevant_facts ctxt full_types relevance_thresholds
40369
53dca3bd4250 use the SMT integration's official list of built-ins
blanchet
parents: 40301
diff changeset
   367
          (the_default default_max_relevant max_relevant) is_built_in_const
40071
658a37c80b53 generalized the relevance filter so that it takes the list of "irrelevant constants" as argument (since the ATP and SMT preprocessing are different)
blanchet
parents: 40070
diff changeset
   368
          relevance_fudge relevance_override chained_ths hyp_ts concl_t
39004
f1b465f889b5 translate the axioms to FOF once and for all ATPs
blanchet
parents: 39003
diff changeset
   369
    val problem =
39321
23951979a362 make Mirabelle happy
blanchet
parents: 39005
diff changeset
   370
      {state = st', goal = goal, subgoal = i,
40065
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
   371
       subgoal_count = Sledgehammer_Util.subgoal_count st,
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
   372
       facts = facts |> map Sledgehammer.Untranslated, only = true}
32574
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   373
    val time_limit =
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   374
      (case hard_timeout of
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   375
        NONE => I
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   376
      | SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs))
40374
443b426e05ea make Mirabelle work correctly if the prover (e.g. the SMT solver) returns no timing information
blanchet
parents: 40372
diff changeset
   377
    val ({outcome, message, used_facts, run_time_in_msecs, ...}
40061
71cc5aac8b76 generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents: 40060
diff changeset
   378
         : Sledgehammer.prover_result,
35971
4f24a4e9af4a make Mirabelle happy again
blanchet
parents: 35871
diff changeset
   379
        time_isa) = time_limit (Mirabelle.cpu_time
38103
4339b1acfd4f make Mirabelle happy
blanchet
parents: 38084
diff changeset
   380
                      (prover params (K ""))) problem
40374
443b426e05ea make Mirabelle work correctly if the prover (e.g. the SMT solver) returns no timing information
blanchet
parents: 40372
diff changeset
   381
    val time_prover = run_time_in_msecs |> the_default ~1
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   382
  in
36405
05a8d79eb338 compile
blanchet
parents: 36373
diff changeset
   383
    case outcome of
40204
da97d75e20e6 standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents: 40200
diff changeset
   384
      NONE => (message, SH_OK (time_isa, time_prover, used_facts))
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   385
    | SOME _ => (message, SH_FAIL (time_isa, time_prover))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   386
  end
37994
b04307085a09 make TPTP generator accept full first-order formulas
blanchet
parents: 37631
diff changeset
   387
  handle ERROR msg => ("error: " ^ msg, SH_ERROR)
32574
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   388
       | TimeLimit.TimeOut => ("timeout", SH_ERROR)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   389
32454
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   390
fun thms_of_name ctxt name =
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   391
  let
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36959
diff changeset
   392
    val lex = Keyword.get_lexicons
32454
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   393
    val get = maps (ProofContext.get_fact ctxt o fst)
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   394
  in
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   395
    Source.of_string name
40526
ca3c6b1bfcdb total Symbol.source;
wenzelm
parents: 40417
diff changeset
   396
    |> Symbol.source
36959
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36954
diff changeset
   397
    |> Token.source {do_recover=SOME false} lex Position.start
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36954
diff changeset
   398
    |> Token.source_proper
f5417836dbea renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
wenzelm
parents: 36954
diff changeset
   399
    |> Source.source Token.stopper (Parse_Spec.xthms1 >> get) NONE
32454
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   400
    |> Source.exhaust
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   401
  end
32452
d84edd022efe apply metis with found theorems in case sledgehammer was successful
boehmes
parents: 32434
diff changeset
   402
32498
1132c7c13f36 Mirabelle: actions are responsible for handling exceptions,
boehmes
parents: 32496
diff changeset
   403
in
1132c7c13f36 Mirabelle: actions are responsible for handling exceptions,
boehmes
parents: 32496
diff changeset
   404
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   405
fun run_sledgehammer trivial args reconstructor named_thms id ({pre=st, log, ...}: Mirabelle.run_args) =
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   406
  let
39340
3998dc0bf82b indicate which goals are trivial
blanchet
parents: 39337
diff changeset
   407
    val triv_str = if trivial then "[T] " else ""
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   408
    val _ = change_data id inc_sh_calls
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   409
    val _ = if trivial then () else change_data id inc_sh_nontriv_calls
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40065
diff changeset
   410
    val (prover_name, prover) = get_prover (Proof.context_of st) args
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   411
    val dir = AList.lookup (op =) args keepK
32541
cea1716eb106 timeout option for ATPs
boehmes
parents: 32536
diff changeset
   412
    val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30)
40417
a29b2fee592b always use a hard timeout in Mirabelle
blanchet
parents: 40416
diff changeset
   413
    val hard_timeout = SOME timeout (* always use a hard timeout *)
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   414
    val (msg, result) = run_sh prover_name prover hard_timeout timeout dir st
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   415
  in
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   416
    case result of
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   417
      SH_OK (time_isa, time_prover, names) =>
38700
fcb0fe4c2f27 make Mirabelle happy
blanchet
parents: 38590
diff changeset
   418
        let
39377
blanchet
parents: 39366
diff changeset
   419
          fun get_thms (_, Sledgehammer_Filter.Chained) = NONE
38826
f42f425edf24 drop chained facts
blanchet
parents: 38752
diff changeset
   420
            | get_thms (name, loc) =
f42f425edf24 drop chained facts
blanchet
parents: 38752
diff changeset
   421
              SOME ((name, loc), thms_of_name (Proof.context_of st) name)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   422
        in
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   423
          change_data id inc_sh_success;
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   424
          if trivial then () else change_data id inc_sh_nontriv_success;
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   425
          change_data id (inc_sh_lemmas (length names));
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   426
          change_data id (inc_sh_max_lems (length names));
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   427
          change_data id (inc_sh_time_isa time_isa);
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   428
          change_data id (inc_sh_time_prover time_prover);
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   429
          reconstructor := reconstructor_from_msg msg;
38826
f42f425edf24 drop chained facts
blanchet
parents: 38752
diff changeset
   430
          named_thms := SOME (map_filter get_thms names);
39340
3998dc0bf82b indicate which goals are trivial
blanchet
parents: 39337
diff changeset
   431
          log (sh_tag id ^ triv_str ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   432
            string_of_int time_prover ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   433
        end
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   434
    | SH_FAIL (time_isa, time_prover) =>
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   435
        let
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   436
          val _ = change_data id (inc_sh_time_isa time_isa)
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   437
          val _ = change_data id (inc_sh_time_prover_fail time_prover)
39340
3998dc0bf82b indicate which goals are trivial
blanchet
parents: 39337
diff changeset
   438
        in log (sh_tag id ^ triv_str ^ "failed: " ^ msg) end
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   439
    | SH_ERROR => log (sh_tag id ^ "failed: " ^ msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   440
  end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   441
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   442
end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   443
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   444
fun run_minimize args reconstructor named_thms id
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   445
        ({pre=st, log, ...}: Mirabelle.run_args) =
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   446
  let
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40065
diff changeset
   447
    val ctxt = Proof.context_of st
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   448
    val n0 = length (these (!named_thms))
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40065
diff changeset
   449
    val (prover_name, _) = get_prover ctxt args
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   450
    val timeout =
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   451
      AList.lookup (op =) args minimize_timeoutK
40627
becf5d5187cc renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents: 40554
diff changeset
   452
      |> Option.map (fst o read_int o raw_explode)  (* FIXME Symbol.explode (?) *)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   453
      |> the_default 5
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40065
diff changeset
   454
    val params = Sledgehammer_Isar.default_params ctxt
40554
ff446d5e9a62 turn on Sledgehammer verbosity so we can track down crashes
blanchet
parents: 40526
diff changeset
   455
      [("verbose", "true"),
ff446d5e9a62 turn on Sledgehammer verbosity so we can track down crashes
blanchet
parents: 40526
diff changeset
   456
       ("provers", prover_name),
ff446d5e9a62 turn on Sledgehammer verbosity so we can track down crashes
blanchet
parents: 40526
diff changeset
   457
       ("timeout", Int.toString timeout)]
37587
466031e057a0 compile
blanchet
parents: 37578
diff changeset
   458
    val minimize =
40065
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
   459
      Sledgehammer_Minimize.minimize_facts params 1
1e4c7185f3f9 remove more needless code ("run_smt_solvers");
blanchet
parents: 40063
diff changeset
   460
                                           (Sledgehammer_Util.subgoal_count st)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   461
    val _ = log separator
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   462
  in
35971
4f24a4e9af4a make Mirabelle happy again
blanchet
parents: 35871
diff changeset
   463
    case minimize st (these (!named_thms)) of
35871
c93bda4fdf15 remove the iteration counter from Sledgehammer's minimizer
blanchet
parents: 35867
diff changeset
   464
      (SOME named_thms', msg) =>
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   465
        (change_data id inc_min_succs;
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   466
         change_data id (inc_min_ab_ratios ((100 * length named_thms') div n0));
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   467
         if length named_thms' = n0
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   468
         then log (minimize_tag id ^ "already minimal")
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   469
         else (reconstructor := reconstructor_from_msg msg;
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   470
               named_thms := SOME named_thms';
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   471
               log (minimize_tag id ^ "succeeded:\n" ^ msg))
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   472
        )
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   473
    | (NONE, msg) => log (minimize_tag id ^ "failed: " ^ msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   474
  end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   475
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   476
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   477
fun run_reconstructor trivial full m name reconstructor named_thms id
32567
de411627a985 explicitly export type abbreviations (as usual in SML97);
wenzelm
parents: 32564
diff changeset
   478
    ({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) =
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   479
  let
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   480
    fun do_reconstructor thms ctxt =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   481
      (if !reconstructor = "smt" then SMT_Solver.smt_tac
40416
6461fc0f9f47 use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents: 40374
diff changeset
   482
       else if full then Metis_Tactics.metisFT_tac
6461fc0f9f47 use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents: 40374
diff changeset
   483
       else Metis_Tactics.metis_tac) ctxt thms
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   484
    fun apply_reconstructor thms =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   485
      Mirabelle.can_apply timeout (do_reconstructor thms) st
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   486
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   487
    fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   488
      | with_time (true, t) = (change_data id (inc_reconstructor_success m);
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   489
          if trivial then ()
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   490
          else change_data id (inc_reconstructor_nontriv_success m);
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   491
          change_data id (inc_reconstructor_lemmas m (length named_thms));
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   492
          change_data id (inc_reconstructor_time m t);
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   493
          change_data id (inc_reconstructor_posns m (pos, trivial));
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   494
          if name = "proof" then change_data id (inc_reconstructor_proofs m)
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   495
          else ();
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   496
          "succeeded (" ^ string_of_int t ^ ")")
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   497
    fun timed_reconstructor thms =
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   498
      (with_time (Mirabelle.cpu_time apply_reconstructor thms), true)
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   499
      handle TimeLimit.TimeOut => (change_data id (inc_reconstructor_timeout m);
34052
b2e6245fb3da only invoke metisFT if metis failed
boehmes
parents: 34035
diff changeset
   500
               ("timeout", false))
b2e6245fb3da only invoke metisFT if metis failed
boehmes
parents: 34035
diff changeset
   501
           | ERROR msg => ("error: " ^ msg, false)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   502
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   503
    val _ = log separator
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   504
    val _ = change_data id (inc_reconstructor_calls m)
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   505
    val _ = if trivial then ()
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   506
            else change_data id (inc_reconstructor_nontriv_calls m)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   507
  in
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   508
    maps snd named_thms
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   509
    |> timed_reconstructor
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   510
    |>> log o prefix (reconstructor_tag reconstructor id)
34052
b2e6245fb3da only invoke metisFT if metis failed
boehmes
parents: 34035
diff changeset
   511
    |> snd
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   512
  end
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   513
40301
bf39a257b3d3 simplified some time constants;
wenzelm
parents: 40204
diff changeset
   514
val try_inner_timeout = seconds 5.0
bf39a257b3d3 simplified some time constants;
wenzelm
parents: 40204
diff changeset
   515
val try_outer_timeout = seconds 30.0
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   516
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   517
fun sledgehammer_action args id (st as {pre, name, ...}: Mirabelle.run_args) =
35592
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   518
  let val goal = Thm.major_prem_of (#goal (Proof.goal pre)) in
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   519
    if can Logic.dest_conjunction goal orelse can Logic.dest_equals goal
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   520
    then () else
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   521
    let
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   522
      val reconstructor = Unsynchronized.ref ""
38700
fcb0fe4c2f27 make Mirabelle happy
blanchet
parents: 38590
diff changeset
   523
      val named_thms =
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38700
diff changeset
   524
        Unsynchronized.ref (NONE : ((string * locality) * thm list) list option)
40416
6461fc0f9f47 use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents: 40374
diff changeset
   525
      val ctxt = Proof.context_of pre
6461fc0f9f47 use "smt" (rather than "metis") to reconstruct SMT proofs
blanchet
parents: 40374
diff changeset
   526
      val (prover_name, _) = get_prover ctxt args
35592
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   527
      val minimize = AList.defined (op =) args minimizeK
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   528
      val metis_ft = AList.defined (op =) args metis_ftK
40694
77435a7853d1 reverted c059d550afec -- the triviality check had apparently nothing to do with spontaneous Interrupt exceptions
blanchet
parents: 40670
diff changeset
   529
      val trivial =
40670
c059d550afec disable triviality checking -- it might be the source of the spurious Interrupt exceptions that make it almost impossible to run Judgement Day
blanchet
parents: 40667
diff changeset
   530
       TimeLimit.timeLimit try_outer_timeout
39363
1c4e38d635e1 added a timeout around "try" call in Mirabelle
blanchet
parents: 39341
diff changeset
   531
                                   (Try.invoke_try (SOME try_inner_timeout)) pre
1c4e38d635e1 added a timeout around "try" call in Mirabelle
blanchet
parents: 39341
diff changeset
   532
                    handle TimeLimit.TimeOut => false
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   533
      fun apply_reconstructor m1 m2 =
35592
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   534
        if metis_ft
34052
b2e6245fb3da only invoke metisFT if metis failed
boehmes
parents: 34035
diff changeset
   535
        then
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   536
          if not (Mirabelle.catch_result (reconstructor_tag reconstructor) false
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   537
              (run_reconstructor trivial false m1 name reconstructor
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   538
                   (these (!named_thms))) id st)
35592
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   539
          then
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   540
            (Mirabelle.catch_result (reconstructor_tag reconstructor) false
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   541
              (run_reconstructor trivial true m2 name reconstructor
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   542
                   (these (!named_thms))) id st; ())
35592
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   543
          else ()
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   544
        else
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   545
          (Mirabelle.catch_result (reconstructor_tag reconstructor) false
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   546
            (run_reconstructor trivial false m1 name reconstructor
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   547
                 (these (!named_thms))) id st; ())
35592
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   548
    in 
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   549
      change_data id (set_mini minimize);
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   550
      Mirabelle.catch sh_tag (run_sledgehammer trivial args reconstructor
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   551
                                               named_thms) id st;
35592
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   552
      if is_some (!named_thms)
32612
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   553
      then
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   554
       (apply_reconstructor Unminimized UnminimizedFT;
35592
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   555
        if minimize andalso not (null (these (!named_thms)))
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   556
        then
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   557
         (Mirabelle.catch minimize_tag
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   558
              (run_minimize args reconstructor named_thms) id st;
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   559
          apply_reconstructor Minimized MinimizedFT)
35592
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   560
        else ())
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   561
      else ()
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   562
    end
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   563
  end
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   564
32511
43d2ac4aa2de added option full_typed for sledgehammer action
boehmes
parents: 32510
diff changeset
   565
fun invoke args =
32515
e7c0d3c0494a Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents: 32511
diff changeset
   566
  let
36373
66af0a49de39 move some sledgehammer stuff out of "atp_manager.ML"
blanchet
parents: 36294
diff changeset
   567
    val _ = Sledgehammer_Isar.full_types := AList.defined (op =) args full_typesK
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   568
  in Mirabelle.register (init, sledgehammer_action args, done) end
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   569
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   570
end