src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML
author blanchet
Wed, 16 Nov 2011 16:35:19 +0100
changeset 45520 2b1dde0b1c30
parent 45519 cd6e78cb6ee8
child 45706 418846ea4f99
permissions -rw-r--r--
thread in additional options to minimizer
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"
43626
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43569
diff changeset
    11
val type_encK = "type_enc"
44089
10287833549f added "sound" option to Mirabelle
blanchet
parents: 43904
diff changeset
    12
val soundK = "sound"
42725
64dea91bbe0e added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents: 42642
diff changeset
    13
val slicingK = "slicing"
45514
973bb7846505 parse lambda translation option in Metis
blanchet
parents: 45512
diff changeset
    14
val lam_transK = "lam_trans"
42725
64dea91bbe0e added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents: 42642
diff changeset
    15
val e_weight_methodK = "e_weight_method"
44099
5e14f591515e support local HOATPs
blanchet
parents: 44090
diff changeset
    16
val force_sosK = "force_sos"
41752
949eaf045e00 added option to Mirabelle Sledgehammer
blanchet
parents: 41742
diff changeset
    17
val max_relevantK = "max_relevant"
44431
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
    18
val max_callsK = "max_calls"
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    19
val minimizeK = "minimize"
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    20
val minimize_timeoutK = "minimize_timeout"
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    21
val metis_ftK = "metis_ft"
41357
ae76960d86a2 added "sledgehammer_tac" as possible reconstructor in Mirabelle
blanchet
parents: 41338
diff changeset
    22
val reconstructorK = "reconstructor"
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    23
44448
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
    24
val preplay_timeout = "4"
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
    25
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    26
fun sh_tag id = "#" ^ string_of_int id ^ " sledgehammer: "
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    27
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
    28
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
    29
  "#" ^ string_of_int id ^ " " ^ (!reconstructor) ^ " (sledgehammer): "
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    30
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    31
val separator = "-----"
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    32
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    33
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    34
datatype sh_data = ShData of {
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    35
  calls: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    36
  success: int,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    37
  nontriv_calls: int,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    38
  nontriv_success: int,
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
    39
  lemmas: int,
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    40
  max_lems: int,
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    41
  time_isa: int,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    42
  time_prover: int,
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    43
  time_prover_fail: int}
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    44
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    45
datatype re_data = ReData of {
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    46
  calls: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    47
  success: int,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    48
  nontriv_calls: int,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    49
  nontriv_success: int,
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
    50
  proofs: int,
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    51
  time: int,
32550
d57c7a2d927c logging number of metis lemmas
nipkow
parents: 32549
diff changeset
    52
  timeout: int,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    53
  lemmas: int * int * int,
39341
d2b981a0429a indicate triviality in the list of proved things
blanchet
parents: 39340
diff changeset
    54
  posns: (Position.T * bool) list
32550
d57c7a2d927c logging number of metis lemmas
nipkow
parents: 32549
diff changeset
    55
  }
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    56
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    57
datatype min_data = MinData of {
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
    58
  succs: int,
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
    59
  ab_ratios: int
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    60
  }
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    61
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    62
fun make_sh_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    63
      (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
    64
       time_prover,time_prover_fail) =
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    65
  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
    66
         nontriv_success=nontriv_success, lemmas=lemmas, max_lems=max_lems,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    67
         time_isa=time_isa, time_prover=time_prover,
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    68
         time_prover_fail=time_prover_fail}
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    69
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
    70
fun make_min_data (succs, ab_ratios) =
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
    71
  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
    72
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    73
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
    74
                  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
    75
  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
    76
         nontriv_success=nontriv_success, proofs=proofs, time=time,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    77
         timeout=timeout, lemmas=lemmas, posns=posns}
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    78
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    79
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
    80
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
    81
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
    82
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    83
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
    84
                              lemmas, max_lems, time_isa,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    85
  time_prover, time_prover_fail}) = (calls, success, nontriv_calls,
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    86
  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
    87
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
    88
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
    89
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    90
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
    91
  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
    92
  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
    93
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    94
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    95
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
    96
  Unminimized | Minimized | UnminimizedFT | MinimizedFT
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    97
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    98
datatype data = Data of {
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    99
  sh: sh_data,
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   100
  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
   101
  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
   102
  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
   103
  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
   104
  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
   105
  mini: bool   (* with minimization *)
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   106
  }
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   107
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   108
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
   109
  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
   110
    mini=mini}
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   111
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   112
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
   113
  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
   114
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   115
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
   116
  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
   117
  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
   118
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   119
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
   120
  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
   121
  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
   122
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   123
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
   124
  let
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   125
    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
   126
      | 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
   127
      | 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
   128
      | 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
   129
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   130
    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
   131
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   132
    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
   133
      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
   134
  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
   135
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   136
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
   137
  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
   138
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   139
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
   140
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   141
val inc_sh_calls =  map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   142
  (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
   143
    => (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
   144
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   145
val inc_sh_success = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   146
  (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
   147
    => (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
   148
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   149
val inc_sh_nontriv_calls =  map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   150
  (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
   151
    => (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
   152
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   153
val inc_sh_nontriv_success = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   154
  (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
   155
    => (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
   156
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   157
fun inc_sh_lemmas n = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   158
  (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
   159
    => (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
   160
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   161
fun inc_sh_max_lems n = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   162
  (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
   163
    => (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
   164
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   165
fun inc_sh_time_isa t = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   166
  (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
   167
    => (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
   168
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   169
fun inc_sh_time_prover t = map_sh_data
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   170
  (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
   171
    => (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
   172
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   173
fun inc_sh_time_prover_fail t = map_sh_data
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   174
  (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
   175
    => (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
   176
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   177
val inc_min_succs = map_min_data
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
   178
  (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
   179
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   180
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
   181
  (fn (succs, ab_ratios) => (succs, ab_ratios+r))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   182
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   183
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
   184
  (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
   185
    => (calls + 1, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   186
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   187
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
   188
  (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
   189
    => (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
   190
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   191
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
   192
  (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
   193
    => (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
   194
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   195
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
   196
  (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
   197
    => (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
   198
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   199
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
   200
  (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
   201
    => (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
   202
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   203
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
   204
 (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
   205
  => (calls, success, nontriv_calls, nontriv_success, proofs, time + t, timeout, lemmas,posns)) m
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   206
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   207
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
   208
  (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
   209
    => (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
   210
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   211
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
   212
  (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
   213
    => (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
   214
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   215
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
   216
  (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
   217
    => (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
   218
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   219
val str0 = string_of_int o the_default 0
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   220
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   221
local
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   222
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   223
val str = string_of_int
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   224
val str3 = Real.fmt (StringCvt.FIX (SOME 3))
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   225
fun percentage a b = string_of_int (a * 100 div b)
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   226
fun time t = Real.fromInt t / 1000.0
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   227
fun avg_time t n =
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   228
  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
   229
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   230
fun log_sh_data log
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   231
    (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
   232
 (log ("Total number of sledgehammer calls: " ^ str calls);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   233
  log ("Number of successful sledgehammer calls: " ^ str success);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   234
  log ("Number of sledgehammer lemmas: " ^ str lemmas);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   235
  log ("Max number of sledgehammer lemmas: " ^ str max_lems);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   236
  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
   237
  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
   238
  log ("Number of successful nontrivial sledgehammer calls: " ^ str nontriv_success);
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   239
  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
   240
  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
   241
  log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_prover_fail));
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   242
  log ("Average time for sledgehammer calls (Isabelle): " ^
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   243
    str3 (avg_time time_isa calls));
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   244
  log ("Average time for successful sledgehammer calls (ATP): " ^
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   245
    str3 (avg_time time_prover success));
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   246
  log ("Average time for failed sledgehammer calls (ATP): " ^
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   247
    str3 (avg_time time_prover_fail (calls - success)))
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   248
  )
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   249
39341
d2b981a0429a indicate triviality in the list of proved things
blanchet
parents: 39340
diff changeset
   250
fun str_of_pos (pos, triv) =
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   251
  str0 (Position.line_of pos) (* ^ ":" ^ str0 (Position.offset_of pos) *) ^
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   252
  (if triv then "[T]" else "")
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   253
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   254
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
   255
     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
   256
    (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
   257
 (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
   258
  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
   259
    " (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
   260
  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
   261
  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
   262
  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
   263
  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
   264
    " (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
   265
  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
   266
  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
   267
  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
   268
  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
   269
  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
   270
    str3 (avg_time re_time re_success));
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   271
  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
   272
  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
   273
  else ()
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   274
 )
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   275
35866
513074557e06 move the Sledgehammer Isar commands together into one file;
blanchet
parents: 35830
diff changeset
   276
fun log_min_data log (succs, ab_ratios) =
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   277
  (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
   278
   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
   279
  )
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   280
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   281
in
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   282
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   283
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
   284
  let
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   285
    val ShData {calls=sh_calls, ...} = sh
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   286
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   287
    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
   288
    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
   289
      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
   290
    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
   291
      (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
   292
  in
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   293
    if sh_calls > 0
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   294
    then
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   295
     (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
   296
      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
   297
      log "";
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   298
      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
   299
      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
   300
      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
   301
        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
   302
         (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
   303
          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
   304
          app_if re_m (fn () =>
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   305
            (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
   306
             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
   307
    else ()
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   308
  end
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   309
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   310
end
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   311
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   312
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   313
(* Warning: we implicitly assume single-threaded execution here! *)
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   314
val data = Unsynchronized.ref ([] : (int * data) list)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   315
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   316
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
   317
fun done id ({log, ...}: Mirabelle.done_args) =
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   318
  AList.lookup (op =) (!data) id
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   319
  |> Option.map (log_data id log)
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   320
  |> K ()
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   321
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   322
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
   323
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   324
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   325
fun get_prover ctxt args =
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   326
  let
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   327
    fun default_prover_name () =
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40065
diff changeset
   328
      hd (#provers (Sledgehammer_Isar.default_params ctxt []))
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   329
      handle Empty => error "No ATP available."
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 40983
diff changeset
   330
    fun get_prover name =
43021
5910dd009d0e handle non-auto try case of Sledgehammer better
blanchet
parents: 43020
diff changeset
   331
      (name, Sledgehammer_Run.get_minimizing_prover ctxt
5910dd009d0e handle non-auto try case of Sledgehammer better
blanchet
parents: 43020
diff changeset
   332
                Sledgehammer_Provers.Normal name)
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   333
  in
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   334
    (case AList.lookup (op =) args proverK of
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   335
      SOME name => get_prover name
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   336
    | NONE => get_prover (default_prover_name ()))
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   337
  end
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   338
43088
0a97f3295642 compile
blanchet
parents: 43064
diff changeset
   339
type locality = ATP_Translate.locality
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38700
diff changeset
   340
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   341
(* hack *)
41357
ae76960d86a2 added "sledgehammer_tac" as possible reconstructor in Mirabelle
blanchet
parents: 41338
diff changeset
   342
fun reconstructor_from_msg args msg =
ae76960d86a2 added "sledgehammer_tac" as possible reconstructor in Mirabelle
blanchet
parents: 41338
diff changeset
   343
  (case AList.lookup (op =) args reconstructorK of
ae76960d86a2 added "sledgehammer_tac" as possible reconstructor in Mirabelle
blanchet
parents: 41338
diff changeset
   344
    SOME name => name
ae76960d86a2 added "sledgehammer_tac" as possible reconstructor in Mirabelle
blanchet
parents: 41338
diff changeset
   345
  | NONE =>
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   346
    if String.isSubstring "metis (" msg then
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   347
      msg |> Substring.full
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   348
          |> Substring.position "metis ("
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   349
          |> snd |> Substring.position ")"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   350
          |> fst |> Substring.string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   351
          |> suffix ")"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   352
    else if String.isSubstring "metis" msg then
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   353
      "metis"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   354
    else
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   355
      "smt")
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   356
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   357
local
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   358
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   359
datatype sh_result =
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38700
diff changeset
   360
  SH_OK of int * int * (string * locality) list |
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   361
  SH_FAIL of int * int |
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   362
  SH_ERROR
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   363
45514
973bb7846505 parse lambda translation option in Metis
blanchet
parents: 45512
diff changeset
   364
fun run_sh prover_name prover type_enc sound max_relevant slicing lam_trans
45512
a6cce8032fff rename configuration option to more reasonable length
blanchet
parents: 45371
diff changeset
   365
        e_weight_method force_sos hard_timeout timeout dir pos st =
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   366
  let
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   367
    val {context = ctxt, facts = chained_ths, goal} = Proof.goal st
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   368
    val i = 1
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   369
    fun set_file_name (SOME dir) =
41337
263fe1670067 mechanism to keep SMT input and output files around in Mirabelle
blanchet
parents: 41276
diff changeset
   370
        Config.put Sledgehammer_Provers.dest_dir dir
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   371
        #> Config.put Sledgehammer_Provers.problem_prefix
44424
2434dd7519e8 clearer separator in generated file names
blanchet
parents: 44099
diff changeset
   372
          ("prob_" ^ str0 (Position.line_of pos) ^ "__")
41337
263fe1670067 mechanism to keep SMT input and output files around in Mirabelle
blanchet
parents: 41276
diff changeset
   373
        #> Config.put SMT_Config.debug_files
43088
0a97f3295642 compile
blanchet
parents: 43064
diff changeset
   374
          (dir ^ "/" ^ Name.desymbolize false (ATP_Util.timestamp ()) ^ "_"
41338
ffd730fcf0ac avoid weird symbols in path
blanchet
parents: 41337
diff changeset
   375
          ^ serial_string ())
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   376
      | set_file_name NONE = I
39321
23951979a362 make Mirabelle happy
blanchet
parents: 39005
diff changeset
   377
    val st' =
23951979a362 make Mirabelle happy
blanchet
parents: 39005
diff changeset
   378
      st |> Proof.map_context
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   379
                (set_file_name dir
42725
64dea91bbe0e added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents: 42642
diff changeset
   380
                 #> (Option.map (Config.put ATP_Systems.e_weight_method)
64dea91bbe0e added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents: 42642
diff changeset
   381
                       e_weight_method |> the_default I)
44099
5e14f591515e support local HOATPs
blanchet
parents: 44090
diff changeset
   382
                 #> (Option.map (Config.put ATP_Systems.force_sos)
44636
9a8de0397f65 always measure time for ATPs -- auto minimization relies on it
blanchet
parents: 44634
diff changeset
   383
                       force_sos |> the_default I))
42642
blanchet
parents: 42638
diff changeset
   384
    val params as {relevance_thresholds, max_relevant, slicing, ...} =
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40065
diff changeset
   385
      Sledgehammer_Isar.default_params ctxt
40554
ff446d5e9a62 turn on Sledgehammer verbosity so we can track down crashes
blanchet
parents: 40526
diff changeset
   386
          [("verbose", "true"),
43626
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43569
diff changeset
   387
           ("type_enc", type_enc),
44089
10287833549f added "sound" option to Mirabelle
blanchet
parents: 43904
diff changeset
   388
           ("sound", sound),
45514
973bb7846505 parse lambda translation option in Metis
blanchet
parents: 45512
diff changeset
   389
           ("lam_trans", lam_trans |> the_default "smart"),
44448
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   390
           ("preplay_timeout", preplay_timeout),
41752
949eaf045e00 added option to Mirabelle Sledgehammer
blanchet
parents: 41742
diff changeset
   391
           ("max_relevant", max_relevant),
42725
64dea91bbe0e added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents: 42642
diff changeset
   392
           ("slicing", slicing),
44448
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   393
           ("timeout", string_of_int timeout),
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   394
           ("preplay_timeout", preplay_timeout)]
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   395
    val default_max_relevant =
42443
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   396
      Sledgehammer_Provers.default_max_relevant_for_prover ctxt slicing
724e612ba248 implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents: 42361
diff changeset
   397
        prover_name
42952
96f62b77748f tuning -- the "appropriate" terminology is inspired from TPTP
blanchet
parents: 42944
diff changeset
   398
    val is_appropriate_prop =
96f62b77748f tuning -- the "appropriate" terminology is inspired from TPTP
blanchet
parents: 42944
diff changeset
   399
      Sledgehammer_Provers.is_appropriate_prop_for_prover ctxt prover_name
40369
53dca3bd4250 use the SMT integration's official list of built-ins
blanchet
parents: 40301
diff changeset
   400
    val is_built_in_const =
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 40983
diff changeset
   401
      Sledgehammer_Provers.is_built_in_const_for_prover ctxt prover_name
40941
a3e6f8634a11 replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents: 40694
diff changeset
   402
    val relevance_fudge =
41087
d7b5fd465198 split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents: 40983
diff changeset
   403
      Sledgehammer_Provers.relevance_fudge_for_prover ctxt prover_name
40070
bdb890782d4a replaced references with proper record that's threaded through
blanchet
parents: 40069
diff changeset
   404
    val relevance_override = {add = [], del = [], only = false}
43088
0a97f3295642 compile
blanchet
parents: 43064
diff changeset
   405
    val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal i
32574
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   406
    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
   407
      (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
   408
        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
   409
      | SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs))
42953
26111aafab12 detect inappropriate problems and crashes better in Waldmeister
blanchet
parents: 42952
diff changeset
   410
    fun failed failure =
45371
blanchet
parents: 44768
diff changeset
   411
      ({outcome = SOME failure, used_facts = [], run_time = Time.zeroTime,
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   412
        preplay =
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   413
          K (ATP_Reconstruct.Failed_to_Play Sledgehammer_Provers.plain_metis),
43261
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43228
diff changeset
   414
        message = K "", message_tail = ""}, ~1)
45371
blanchet
parents: 44768
diff changeset
   415
    val ({outcome, used_facts, run_time, preplay, message, message_tail}
blanchet
parents: 44768
diff changeset
   416
         : Sledgehammer_Provers.prover_result,
41275
3897af72c731 move relevance filter into hard timeout
blanchet
parents: 41274
diff changeset
   417
        time_isa) = time_limit (Mirabelle.cpu_time (fn () =>
3897af72c731 move relevance filter into hard timeout
blanchet
parents: 41274
diff changeset
   418
      let
42953
26111aafab12 detect inappropriate problems and crashes better in Waldmeister
blanchet
parents: 42952
diff changeset
   419
        val _ = if is_appropriate_prop concl_t then ()
26111aafab12 detect inappropriate problems and crashes better in Waldmeister
blanchet
parents: 42952
diff changeset
   420
                else raise Fail "inappropriate"
44625
4a1132815a70 more tuning
blanchet
parents: 44591
diff changeset
   421
        val ho_atp = Sledgehammer_Provers.is_ho_atp ctxt prover_name
41275
3897af72c731 move relevance filter into hard timeout
blanchet
parents: 41274
diff changeset
   422
        val facts =
44625
4a1132815a70 more tuning
blanchet
parents: 44591
diff changeset
   423
          Sledgehammer_Filter.nearly_all_facts ctxt ho_atp relevance_override
4a1132815a70 more tuning
blanchet
parents: 44591
diff changeset
   424
            chained_ths hyp_ts concl_t
43351
b19d95b4d736 compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
blanchet
parents: 43261
diff changeset
   425
          |> filter (is_appropriate_prop o prop_of o snd)
44625
4a1132815a70 more tuning
blanchet
parents: 44591
diff changeset
   426
          |> Sledgehammer_Filter.relevant_facts ctxt relevance_thresholds
43351
b19d95b4d736 compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
blanchet
parents: 43261
diff changeset
   427
                 (the_default default_max_relevant max_relevant)
b19d95b4d736 compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
blanchet
parents: 43261
diff changeset
   428
                 is_built_in_const relevance_fudge relevance_override
b19d95b4d736 compute the set of base facts only once (instead of three times in parallel) -- this saves about .5 s of CPU time, albeit much less clock wall time
blanchet
parents: 43261
diff changeset
   429
                 chained_ths hyp_ts concl_t
41275
3897af72c731 move relevance filter into hard timeout
blanchet
parents: 41274
diff changeset
   430
        val problem =
3897af72c731 move relevance filter into hard timeout
blanchet
parents: 41274
diff changeset
   431
          {state = st', goal = goal, subgoal = i,
3897af72c731 move relevance filter into hard timeout
blanchet
parents: 41274
diff changeset
   432
           subgoal_count = Sledgehammer_Util.subgoal_count st,
3897af72c731 move relevance filter into hard timeout
blanchet
parents: 41274
diff changeset
   433
           facts = facts |> map Sledgehammer_Provers.Untranslated_Fact,
41741
839d1488045f renamed field
blanchet
parents: 41491
diff changeset
   434
           smt_filter = NONE}
45520
2b1dde0b1c30 thread in additional options to minimizer
blanchet
parents: 45519
diff changeset
   435
      in prover params (K (K (K ""))) problem end)) ()
42953
26111aafab12 detect inappropriate problems and crashes better in Waldmeister
blanchet
parents: 42952
diff changeset
   436
      handle TimeLimit.TimeOut => failed ATP_Proof.TimedOut
26111aafab12 detect inappropriate problems and crashes better in Waldmeister
blanchet
parents: 42952
diff changeset
   437
           | Fail "inappropriate" => failed ATP_Proof.Inappropriate
45371
blanchet
parents: 44768
diff changeset
   438
    val time_prover = run_time |> Time.toMilliseconds
43261
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43228
diff changeset
   439
    val msg = message (preplay ()) ^ message_tail
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   440
  in
36405
05a8d79eb338 compile
blanchet
parents: 36373
diff changeset
   441
    case outcome of
43052
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   442
      NONE => (msg, SH_OK (time_isa, time_prover, used_facts))
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   443
    | SOME _ => (msg, SH_FAIL (time_isa, time_prover))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   444
  end
37994
b04307085a09 make TPTP generator accept full first-order formulas
blanchet
parents: 37631
diff changeset
   445
  handle ERROR msg => ("error: " ^ msg, SH_ERROR)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   446
32454
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   447
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
   448
  let
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36959
diff changeset
   449
    val lex = Keyword.get_lexicons
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42179
diff changeset
   450
    val get = maps (Proof_Context.get_fact ctxt o fst)
32454
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   451
  in
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   452
    Source.of_string name
40526
ca3c6b1bfcdb total Symbol.source;
wenzelm
parents: 40417
diff changeset
   453
    |> 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
   454
    |> 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
   455
    |> 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
   456
    |> 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
   457
    |> 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
   458
  end
32452
d84edd022efe apply metis with found theorems in case sledgehammer was successful
boehmes
parents: 32434
diff changeset
   459
32498
1132c7c13f36 Mirabelle: actions are responsible for handling exceptions,
boehmes
parents: 32496
diff changeset
   460
in
1132c7c13f36 Mirabelle: actions are responsible for handling exceptions,
boehmes
parents: 32496
diff changeset
   461
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   462
fun run_sledgehammer trivial args reconstructor named_thms id
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   463
      ({pre=st, log, pos, ...}: Mirabelle.run_args) =
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   464
  let
39340
3998dc0bf82b indicate which goals are trivial
blanchet
parents: 39337
diff changeset
   465
    val triv_str = if trivial then "[T] " else ""
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   466
    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
   467
    val _ = if trivial then () else change_data id inc_sh_nontriv_calls
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   468
    val (prover_name, prover) = get_prover (Proof.context_of st) args
43626
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43569
diff changeset
   469
    val type_enc = AList.lookup (op =) args type_encK |> the_default "smart"
44089
10287833549f added "sound" option to Mirabelle
blanchet
parents: 43904
diff changeset
   470
    val sound = AList.lookup (op =) args soundK |> the_default "false"
41752
949eaf045e00 added option to Mirabelle Sledgehammer
blanchet
parents: 41742
diff changeset
   471
    val max_relevant = AList.lookup (op =) args max_relevantK |> the_default "smart"
42725
64dea91bbe0e added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents: 42642
diff changeset
   472
    val slicing = AList.lookup (op =) args slicingK |> the_default "true"
45514
973bb7846505 parse lambda translation option in Metis
blanchet
parents: 45512
diff changeset
   473
    val lam_trans = AList.lookup (op =) args lam_transK
42725
64dea91bbe0e added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents: 42642
diff changeset
   474
    val e_weight_method = AList.lookup (op =) args e_weight_methodK
44099
5e14f591515e support local HOATPs
blanchet
parents: 44090
diff changeset
   475
    val force_sos = AList.lookup (op =) args force_sosK
42725
64dea91bbe0e added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents: 42642
diff changeset
   476
      |> Option.map (curry (op <>) "false")
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   477
    val dir = AList.lookup (op =) args keepK
32541
cea1716eb106 timeout option for ATPs
boehmes
parents: 32536
diff changeset
   478
    val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30)
41268
56b7e277fd7d higher hard timeout
blanchet
parents: 41266
diff changeset
   479
    (* always use a hard timeout, but give some slack so that the automatic
56b7e277fd7d higher hard timeout
blanchet
parents: 41266
diff changeset
   480
       minimizer has a chance to do its magic *)
56b7e277fd7d higher hard timeout
blanchet
parents: 41266
diff changeset
   481
    val hard_timeout = SOME (2 * timeout)
41155
85da8cbb4966 added support for "type_sys" option to Mirabelle
blanchet
parents: 41154
diff changeset
   482
    val (msg, result) =
45514
973bb7846505 parse lambda translation option in Metis
blanchet
parents: 45512
diff changeset
   483
      run_sh prover_name prover type_enc sound max_relevant slicing lam_trans
45512
a6cce8032fff rename configuration option to more reasonable length
blanchet
parents: 45371
diff changeset
   484
        e_weight_method force_sos hard_timeout timeout dir pos st
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   485
  in
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   486
    case result of
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   487
      SH_OK (time_isa, time_prover, names) =>
38700
fcb0fe4c2f27 make Mirabelle happy
blanchet
parents: 38590
diff changeset
   488
        let
44487
0989d8deab69 include chained facts for minimizer, otherwise it won't work
blanchet
parents: 44462
diff changeset
   489
          fun get_thms (name, loc) =
0989d8deab69 include chained facts for minimizer, otherwise it won't work
blanchet
parents: 44462
diff changeset
   490
            SOME ((name, loc), thms_of_name (Proof.context_of st) name)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   491
        in
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   492
          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
   493
          if trivial then () else change_data id inc_sh_nontriv_success;
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   494
          change_data id (inc_sh_lemmas (length names));
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   495
          change_data id (inc_sh_max_lems (length names));
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   496
          change_data id (inc_sh_time_isa time_isa);
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   497
          change_data id (inc_sh_time_prover time_prover);
41357
ae76960d86a2 added "sledgehammer_tac" as possible reconstructor in Mirabelle
blanchet
parents: 41338
diff changeset
   498
          reconstructor := reconstructor_from_msg args msg;
38826
f42f425edf24 drop chained facts
blanchet
parents: 38752
diff changeset
   499
          named_thms := SOME (map_filter get_thms names);
39340
3998dc0bf82b indicate which goals are trivial
blanchet
parents: 39337
diff changeset
   500
          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
   501
            string_of_int time_prover ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   502
        end
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   503
    | SH_FAIL (time_isa, time_prover) =>
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   504
        let
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   505
          val _ = change_data id (inc_sh_time_isa time_isa)
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   506
          val _ = change_data id (inc_sh_time_prover_fail time_prover)
39340
3998dc0bf82b indicate which goals are trivial
blanchet
parents: 39337
diff changeset
   507
        in log (sh_tag id ^ triv_str ^ "failed: " ^ msg) end
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   508
    | SH_ERROR => log (sh_tag id ^ "failed: " ^ msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   509
  end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   510
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   511
end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   512
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   513
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
   514
        ({pre=st, log, ...}: Mirabelle.run_args) =
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   515
  let
40069
6f7bf79b1506 fixed signature of "is_smt_solver_installed";
blanchet
parents: 40065
diff changeset
   516
    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
   517
    val n0 = length (these (!named_thms))
42444
8e5438dc70bb cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents: 42443
diff changeset
   518
    val (prover_name, _) = get_prover ctxt args
43626
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43569
diff changeset
   519
    val type_enc = AList.lookup (op =) args type_encK |> the_default "smart"
44089
10287833549f added "sound" option to Mirabelle
blanchet
parents: 43904
diff changeset
   520
    val sound = AList.lookup (op =) args soundK |> the_default "false"
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   521
    val timeout =
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   522
      AList.lookup (op =) args minimize_timeoutK
40627
becf5d5187cc renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents: 40554
diff changeset
   523
      |> Option.map (fst o read_int o raw_explode)  (* FIXME Symbol.explode (?) *)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   524
      |> the_default 5
43064
b6e61d22fa61 made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
blanchet
parents: 43052
diff changeset
   525
    val params = Sledgehammer_Isar.default_params ctxt
41155
85da8cbb4966 added support for "type_sys" option to Mirabelle
blanchet
parents: 41154
diff changeset
   526
      [("provers", prover_name),
85da8cbb4966 added support for "type_sys" option to Mirabelle
blanchet
parents: 41154
diff changeset
   527
       ("verbose", "true"),
43626
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43569
diff changeset
   528
       ("type_enc", type_enc),
44089
10287833549f added "sound" option to Mirabelle
blanchet
parents: 43904
diff changeset
   529
       ("sound", sound),
44448
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   530
       ("timeout", string_of_int timeout),
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   531
       ("preplay_timeout", preplay_timeout)]
37587
466031e057a0 compile
blanchet
parents: 37578
diff changeset
   532
    val minimize =
41742
11e862c68b40 automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents: 41741
diff changeset
   533
      Sledgehammer_Minimize.minimize_facts prover_name params
43064
b6e61d22fa61 made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
blanchet
parents: 43052
diff changeset
   534
          true 1 (Sledgehammer_Util.subgoal_count st)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   535
    val _ = log separator
43261
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43228
diff changeset
   536
    val (used_facts, (preplay, message, message_tail)) =
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43228
diff changeset
   537
      minimize st (these (!named_thms))
a4aeb26a6362 make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
blanchet
parents: 43228
diff changeset
   538
    val msg = message (preplay ()) ^ message_tail
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   539
  in
43052
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   540
    case used_facts of
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   541
      SOME named_thms' =>
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   542
        (change_data id inc_min_succs;
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   543
         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
   544
         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
   545
         then log (minimize_tag id ^ "already minimal")
41357
ae76960d86a2 added "sledgehammer_tac" as possible reconstructor in Mirabelle
blanchet
parents: 41338
diff changeset
   546
         else (reconstructor := reconstructor_from_msg args msg;
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   547
               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
   548
               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
   549
        )
43052
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   550
    | NONE => log (minimize_tag id ^ "failed: " ^ msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   551
  end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   552
44542
3f5fd3635281 beef up sledgehammer_tac in Mirabelle some more
blanchet
parents: 44487
diff changeset
   553
fun override_params prover type_enc timeout =
3f5fd3635281 beef up sledgehammer_tac in Mirabelle some more
blanchet
parents: 44487
diff changeset
   554
  [("provers", prover),
44449
b622a98b79fb don't select facts when using sledgehammer_tac for reconstruction
blanchet
parents: 44448
diff changeset
   555
   ("max_relevant", "0"),
44542
3f5fd3635281 beef up sledgehammer_tac in Mirabelle some more
blanchet
parents: 44487
diff changeset
   556
   ("type_enc", type_enc),
44430
c6f0490d432f beef up "sledgehammer_tac" reconstructor
blanchet
parents: 44424
diff changeset
   557
   ("sound", "true"),
44461
5e19eecb0e1c specify timeout for "sledgehammer_tac"
blanchet
parents: 44449
diff changeset
   558
   ("slicing", "false"),
5e19eecb0e1c specify timeout for "sledgehammer_tac"
blanchet
parents: 44449
diff changeset
   559
   ("timeout", timeout |> Time.toSeconds |> string_of_int)]
44430
c6f0490d432f beef up "sledgehammer_tac" reconstructor
blanchet
parents: 44424
diff changeset
   560
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   561
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
   562
    ({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) =
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   563
  let
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   564
    fun do_reconstructor named_thms ctxt =
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   565
      let
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   566
        val ref_of_str =
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   567
          suffix ";" #> Outer_Syntax.scan Position.none #> Parse_Spec.xthm
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   568
          #> fst
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   569
        val thms = named_thms |> maps snd
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   570
        val facts = named_thms |> map (ref_of_str o fst o fst)
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   571
        val relevance_override = {add = facts, del = [], only = true}
44566
bf8331161ad9 split timeout among ATPs in and add Metis to the mix as backup
blanchet
parents: 44542
diff changeset
   572
        fun my_timeout time_slice =
bf8331161ad9 split timeout among ATPs in and add Metis to the mix as backup
blanchet
parents: 44542
diff changeset
   573
          timeout |> Time.toReal |> curry Real.* time_slice |> Time.fromReal
bf8331161ad9 split timeout among ATPs in and add Metis to the mix as backup
blanchet
parents: 44542
diff changeset
   574
        fun sledge_tac time_slice prover type_enc =
44542
3f5fd3635281 beef up sledgehammer_tac in Mirabelle some more
blanchet
parents: 44487
diff changeset
   575
          Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt
44566
bf8331161ad9 split timeout among ATPs in and add Metis to the mix as backup
blanchet
parents: 44542
diff changeset
   576
            (override_params prover type_enc (my_timeout time_slice))
bf8331161ad9 split timeout among ATPs in and add Metis to the mix as backup
blanchet
parents: 44542
diff changeset
   577
            relevance_override
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   578
      in
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   579
        if !reconstructor = "sledgehammer_tac" then
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
   580
          sledge_tac 0.2 ATP_Systems.z3_tptpN "mono_simple"
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
   581
          ORELSE' sledge_tac 0.2 ATP_Systems.eN "mono_guards??"
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
   582
          ORELSE' sledge_tac 0.2 ATP_Systems.vampireN "mono_guards??"
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
   583
          ORELSE' sledge_tac 0.2 ATP_Systems.spassN "poly_tags"
45514
973bb7846505 parse lambda translation option in Metis
blanchet
parents: 45512
diff changeset
   584
          ORELSE' Metis_Tactic.metis_tac [] ATP_Translate.combinatorsN ctxt thms
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   585
        else if !reconstructor = "smt" then
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   586
          SMT_Solver.smt_tac ctxt thms
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   587
        else if full then
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   588
          Metis_Tactic.metis_tac [ATP_Reconstruct.full_typesN]
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   589
            ATP_Reconstruct.metis_default_lam_trans ctxt thms
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   590
        else if String.isPrefix "metis (" (!reconstructor) then
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   591
          let
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   592
            val (type_encs, lam_trans) =
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   593
              !reconstructor
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   594
              |> Outer_Syntax.scan Position.start
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   595
              |> filter Token.is_proper |> tl
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   596
              |> Metis_Tactic.parse_metis_options |> fst
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   597
              |>> the_default [ATP_Reconstruct.partial_typesN]
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   598
              ||> the_default ATP_Reconstruct.metis_default_lam_trans
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   599
          in Metis_Tactic.metis_tac type_encs lam_trans ctxt thms end
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   600
        else if !reconstructor = "metis" then
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   601
          Metis_Tactic.metis_tac [] ATP_Reconstruct.metis_default_lam_trans ctxt
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   602
            thms
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   603
        else
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   604
          K all_tac
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   605
      end
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   606
    fun apply_reconstructor named_thms =
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   607
      Mirabelle.can_apply timeout (do_reconstructor named_thms) st
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   608
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   609
    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
   610
      | 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
   611
          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
   612
          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
   613
          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
   614
          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
   615
          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
   616
          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
   617
          else ();
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   618
          "succeeded (" ^ string_of_int t ^ ")")
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   619
    fun timed_reconstructor named_thms =
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   620
      (with_time (Mirabelle.cpu_time apply_reconstructor named_thms), true)
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   621
      handle TimeLimit.TimeOut => (change_data id (inc_reconstructor_timeout m);
34052
b2e6245fb3da only invoke metisFT if metis failed
boehmes
parents: 34035
diff changeset
   622
               ("timeout", false))
b2e6245fb3da only invoke metisFT if metis failed
boehmes
parents: 34035
diff changeset
   623
           | ERROR msg => ("error: " ^ msg, false)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   624
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   625
    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
   626
    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
   627
    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
   628
            else change_data id (inc_reconstructor_nontriv_calls m)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   629
  in
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   630
    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
   631
    |> 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
   632
    |>> log o prefix (reconstructor_tag reconstructor id)
34052
b2e6245fb3da only invoke metisFT if metis failed
boehmes
parents: 34035
diff changeset
   633
    |> snd
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   634
  end
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   635
41276
285aea0c153c two layers of timeouts seem to be less reliable than a single layer
blanchet
parents: 41275
diff changeset
   636
val try_timeout = seconds 5.0
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   637
44431
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   638
(* crude hack *)
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   639
val num_sledgehammer_calls = Unsynchronized.ref 0
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   640
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   641
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
   642
  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
   643
    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
   644
    then () else
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   645
    let
44431
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   646
      val max_calls =
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   647
        AList.lookup (op =) args max_callsK |> the_default "10000000"
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   648
        |> Int.fromString |> the
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   649
      val _ = num_sledgehammer_calls := !num_sledgehammer_calls + 1;
44448
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   650
    in
44431
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   651
      if !num_sledgehammer_calls > max_calls then ()
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   652
      else
44448
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   653
        let
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   654
          val reconstructor = Unsynchronized.ref ""
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   655
          val named_thms =
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   656
            Unsynchronized.ref (NONE : ((string * locality) * thm list) list option)
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   657
          val minimize = AList.defined (op =) args minimizeK
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   658
          val metis_ft = AList.defined (op =) args metis_ftK
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   659
          val trivial =
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   660
            Try_Methods.try_methods (SOME try_timeout) ([], [], [], []) pre
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   661
            handle TimeLimit.TimeOut => false
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   662
          fun apply_reconstructor m1 m2 =
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   663
            if metis_ft
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   664
            then
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   665
              if not (Mirabelle.catch_result (reconstructor_tag reconstructor) false
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   666
                  (run_reconstructor trivial false m1 name reconstructor
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   667
                       (these (!named_thms))) id st)
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   668
              then
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   669
                (Mirabelle.catch_result (reconstructor_tag reconstructor) false
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   670
                  (run_reconstructor trivial true m2 name reconstructor
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   671
                       (these (!named_thms))) id st; ())
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   672
              else ()
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   673
            else
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   674
              (Mirabelle.catch_result (reconstructor_tag reconstructor) false
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   675
                (run_reconstructor trivial false m1 name reconstructor
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   676
                     (these (!named_thms))) id st; ())
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   677
        in
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   678
          change_data id (set_mini minimize);
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   679
          Mirabelle.catch sh_tag (run_sledgehammer trivial args reconstructor
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   680
                                                   named_thms) id st;
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   681
          if is_some (!named_thms)
44431
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   682
          then
44448
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   683
           (apply_reconstructor Unminimized UnminimizedFT;
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   684
            if minimize andalso not (null (these (!named_thms)))
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   685
            then
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   686
             (Mirabelle.catch minimize_tag
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   687
                  (run_minimize args reconstructor named_thms) id st;
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   688
              apply_reconstructor Minimized MinimizedFT)
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   689
            else ())
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   690
          else ()
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   691
        end
35592
768d17f54125 use regular Proof.goal (which is what methods usually see) and prevent sledgehammer from crashing;
wenzelm
parents: 34052
diff changeset
   692
    end
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   693
  end
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   694
32511
43d2ac4aa2de added option full_typed for sledgehammer action
boehmes
parents: 32510
diff changeset
   695
fun invoke args =
43569
b342cd125533 removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
blanchet
parents: 43351
diff changeset
   696
  Mirabelle.register (init, sledgehammer_action args, done)
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   697
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   698
end