src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML
author nipkow
Sun, 18 Oct 2009 16:25:59 +0200
changeset 32991 b6ba8adc14c2
parent 32985 fbc642835ecb
parent 32990 717680b14041
child 33016 b73b74fe23c3
permissions -rw-r--r--
merged
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"
32574
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
    10
val prover_hard_timeoutK = "prover_hard_timeout"
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    11
val keepK = "keep"
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    12
val full_typesK = "full_types"
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    13
val minimizeK = "minimize"
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    14
val minimize_timeoutK = "minimize_timeout"
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    15
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    16
fun sh_tag id = "#" ^ string_of_int id ^ " sledgehammer: "
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    17
fun minimize_tag id = "#" ^ string_of_int id ^ " minimize (sledgehammer): "
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    18
fun metis_tag id = "#" ^ string_of_int id ^ " metis (sledgehammer): "
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    19
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    20
val separator = "-----"
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
    21
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    22
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    23
datatype sh_data = ShData of {
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    24
  calls: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    25
  success: int,
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
    26
  lemmas: int,
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    27
  max_lems: int,
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    28
  time_isa: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    29
  time_atp: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    30
  time_atp_fail: int}
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    31
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    32
datatype me_data = MeData of {
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    33
  calls: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    34
  success: int,
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
    35
  proofs: int,
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    36
  time: int,
32550
d57c7a2d927c logging number of metis lemmas
nipkow
parents: 32549
diff changeset
    37
  timeout: int,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    38
  lemmas: int * int * int,
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
    39
  posns: Position.T list
32550
d57c7a2d927c logging number of metis lemmas
nipkow
parents: 32549
diff changeset
    40
  }
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    41
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    42
datatype min_data = MinData of {
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
    43
  succs: int,
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
    44
  ab_ratios: int,
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
    45
  it_ratios: int
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    46
  }
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    47
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    48
(* The first me_data component is only used if "minimize" is on.
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    49
   Then it records how metis behaves with un-minimized lemmas.
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    50
*)
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    51
datatype data = Data of sh_data * me_data * min_data * me_data
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    52
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    53
fun make_sh_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    54
      (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail) =
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    55
  ShData{calls=calls, success=success, lemmas=lemmas, max_lems=max_lems,
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    56
         time_isa=time_isa, time_atp=time_atp, time_atp_fail=time_atp_fail}
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    57
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
    58
fun make_min_data (succs, ab_ratios, it_ratios) =
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
    59
  MinData{succs=succs, ab_ratios=ab_ratios, it_ratios=it_ratios}
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    60
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    61
fun make_me_data (calls,success,proofs,time,timeout,lemmas,posns) =
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    62
  MeData{calls=calls, success=success, proofs=proofs, time=time,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    63
         timeout=timeout, lemmas=lemmas, posns=posns}
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    64
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    65
val empty_data =
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    66
  Data(make_sh_data (0, 0, 0, 0, 0, 0, 0),
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    67
       make_me_data(0, 0, 0, 0, 0, (0,0,0), []),
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
    68
       MinData{succs=0, ab_ratios=0, it_ratios=0},
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    69
       make_me_data(0, 0, 0, 0, 0, (0,0,0), []))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    70
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    71
fun map_sh_data f
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    72
    (Data(ShData{calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail},
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    73
          meda0, minda, meda)) =
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    74
  Data (make_sh_data (f (calls,success,lemmas,max_lems,
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    75
                         time_isa,time_atp,time_atp_fail)),
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    76
        meda0, minda, meda)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    77
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    78
fun map_min_data f
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
    79
  (Data(shda, meda0, MinData{succs,ab_ratios,it_ratios}, meda)) =
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
    80
  Data(shda, meda0, make_min_data(f(succs,ab_ratios,it_ratios)), meda)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    81
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    82
fun map_me_data0 f (Data (shda, MeData{calls,success,proofs,time,timeout,lemmas,posns}, minda, meda)) =
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    83
  Data(shda, make_me_data(f (calls,success,proofs,time,timeout,lemmas,posns)), minda, meda)
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    84
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    85
fun map_me_data f (Data (shda, meda0, minda, MeData{calls,success,proofs,time,timeout,lemmas,posns})) =
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    86
  Data(shda, meda0, minda, make_me_data(f (calls,success,proofs,time,timeout,lemmas,posns)))
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    87
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    88
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
    89
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    90
val inc_sh_calls =  map_sh_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    91
  (fn (calls, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    92
    => (calls + 1, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    93
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    94
val inc_sh_success = map_sh_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    95
  (fn (calls, success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    96
    => (calls, success + 1, lemmas,max_lems, time_isa, time_atp, time_atp_fail))
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
    97
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    98
fun inc_sh_lemmas n = map_sh_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    99
  (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   100
    => (calls,success,lemmas+n,max_lems,time_isa,time_atp,time_atp_fail))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   101
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   102
fun inc_sh_max_lems n = map_sh_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   103
  (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   104
    => (calls,success,lemmas,Int.max(max_lems,n),time_isa,time_atp,time_atp_fail))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   105
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   106
fun inc_sh_time_isa t = map_sh_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   107
  (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   108
    => (calls,success,lemmas,max_lems,time_isa + t,time_atp,time_atp_fail))
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   109
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   110
fun inc_sh_time_atp t = map_sh_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   111
  (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   112
    => (calls,success,lemmas,max_lems,time_isa,time_atp + t,time_atp_fail))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   113
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   114
fun inc_sh_time_atp_fail t = map_sh_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   115
  (fn (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail)
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   116
    => (calls,success,lemmas,max_lems,time_isa,time_atp,time_atp_fail + t))
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   117
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   118
val inc_min_succs = map_min_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   119
  (fn (succs,ab_ratios,it_ratios) => (succs+1, ab_ratios, it_ratios))
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   120
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   121
fun inc_min_ab_ratios r = map_min_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   122
  (fn (succs, ab_ratios, it_ratios) => (succs, ab_ratios+r, it_ratios))
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   123
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   124
fun inc_min_it_ratios r = map_min_data
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   125
  (fn (succs, ab_ratios, it_ratios) => (succs, ab_ratios, it_ratios+r))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   126
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   127
val inc_metis_calls = map_me_data
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   128
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   129
    => (calls + 1, success, proofs, time, timeout, lemmas,posns))
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   130
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   131
val inc_metis_success = map_me_data
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   132
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   133
    => (calls, success + 1, proofs, time, timeout, lemmas,posns))
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   134
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   135
val inc_metis_proofs = map_me_data
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   136
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   137
    => (calls, success, proofs + 1, time, timeout, lemmas,posns))
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   138
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   139
fun inc_metis_time t = map_me_data
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   140
 (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   141
  => (calls, success, proofs, time + t, timeout, lemmas,posns))
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   142
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   143
val inc_metis_timeout = map_me_data
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   144
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   145
    => (calls, success, proofs, time, timeout + 1, lemmas,posns))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   146
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   147
fun inc_metis_lemmas n = map_me_data
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   148
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   149
    => (calls, success, proofs, time, timeout, inc_max n lemmas, posns))
32550
d57c7a2d927c logging number of metis lemmas
nipkow
parents: 32549
diff changeset
   150
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   151
fun inc_metis_posns pos = map_me_data
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   152
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   153
    => (calls, success, proofs, time, timeout, lemmas, pos::posns))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   154
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   155
val inc_metis_calls0 = map_me_data0 
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   156
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   157
    => (calls + 1, success, proofs, time, timeout, lemmas,posns))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   158
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   159
val inc_metis_success0 = map_me_data0
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   160
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   161
    => (calls, success + 1, proofs, time, timeout, lemmas,posns))
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   162
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   163
val inc_metis_proofs0 = map_me_data0
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   164
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   165
    => (calls, success, proofs + 1, time, timeout, lemmas,posns))
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   166
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   167
fun inc_metis_time0 t = map_me_data0
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   168
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   169
    => (calls, success, proofs, time + t, timeout, lemmas,posns))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   170
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   171
val inc_metis_timeout0 = map_me_data0
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   172
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   173
    => (calls, success, proofs, time, timeout + 1, lemmas,posns))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   174
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   175
fun inc_metis_lemmas0 n = map_me_data0
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   176
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   177
    => (calls, success, proofs, time, timeout, inc_max n lemmas, posns))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   178
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   179
fun inc_metis_posns0 pos = map_me_data0
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   180
  (fn (calls,success,proofs,time,timeout,lemmas,posns)
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   181
    => (calls, success, proofs, time, timeout, lemmas, pos::posns))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   182
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   183
local
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   184
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   185
val str = string_of_int
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   186
val str3 = Real.fmt (StringCvt.FIX (SOME 3))
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   187
fun percentage a b = string_of_int (a * 100 div b)
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   188
fun time t = Real.fromInt t / 1000.0
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   189
fun avg_time t n =
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   190
  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
   191
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   192
fun log_sh_data log calls success lemmas max_lems time_isa time_atp time_atp_fail =
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   193
 (log ("Total number of sledgehammer calls: " ^ str calls);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   194
  log ("Number of successful sledgehammer calls: " ^ str success);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   195
  log ("Number of sledgehammer lemmas: " ^ str lemmas);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   196
  log ("Max number of sledgehammer lemmas: " ^ str max_lems);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   197
  log ("Success rate: " ^ percentage success calls ^ "%");
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   198
  log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time time_isa));
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   199
  log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_atp));
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   200
  log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_atp_fail));
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   201
  log ("Average time for sledgehammer calls (Isabelle): " ^
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   202
    str3 (avg_time time_isa calls));
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   203
  log ("Average time for successful sledgehammer calls (ATP): " ^
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   204
    str3 (avg_time time_atp success));
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   205
  log ("Average time for failed sledgehammer calls (ATP): " ^
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   206
    str3 (avg_time time_atp_fail (calls - success)))
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   207
  )
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   208
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   209
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   210
fun str_of_pos pos =
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   211
  let val str0 = string_of_int o the_default 0
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   212
  in str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos) end
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   213
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   214
fun log_metis_data log tag sh_calls sh_success metis_calls metis_success metis_proofs metis_time
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   215
    metis_timeout (lemmas, lems_sos, lems_max) metis_posns =
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   216
 (log ("Total number of " ^ tag ^ "metis calls: " ^ str metis_calls);
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   217
  log ("Number of successful " ^ tag ^ "metis calls: " ^ str metis_success ^
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   218
    " (proof: " ^ str metis_proofs ^ ")");
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   219
  log ("Number of " ^ tag ^ "metis timeouts: " ^ str metis_timeout);
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   220
  log ("Success rate: " ^ percentage metis_success sh_calls ^ "%");
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   221
  log ("Number of successful " ^ tag ^ "metis lemmas: " ^ str lemmas);
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   222
  log ("SOS of successful " ^ tag ^ "metis lemmas: " ^ str lems_sos);
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   223
  log ("Max number of successful " ^ tag ^ "metis lemmas: " ^ str lems_max);
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   224
  log ("Total time for successful " ^ tag ^ "metis calls: " ^ str3 (time metis_time));
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   225
  log ("Average time for successful metis calls: " ^
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   226
    str3 (avg_time metis_time metis_success));
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   227
  if tag=""
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   228
  then log ("Proved: " ^ space_implode " " (map str_of_pos metis_posns))
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   229
  else ()
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   230
 )
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   231
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   232
fun log_min_data log succs ab_ratios it_ratios =
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   233
  (log ("Number of successful minimizations: " ^ string_of_int succs);
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   234
   log ("After/before ratios: " ^ string_of_int ab_ratios);
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   235
   log ("Iterations ratios: " ^ string_of_int it_ratios)
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   236
  )
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   237
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   238
in
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   239
32567
de411627a985 explicitly export type abbreviations (as usual in SML97);
wenzelm
parents: 32564
diff changeset
   240
fun log_data id log (Data
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   241
   (ShData{calls=sh_calls, lemmas=sh_lemmas,  max_lems=sh_max_lems, success=sh_success,
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
   242
      time_isa=sh_time_isa,time_atp=sh_time_atp,time_atp_fail=sh_time_atp_fail},
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   243
    MeData{calls=metis_calls0, proofs=metis_proofs0,
32567
de411627a985 explicitly export type abbreviations (as usual in SML97);
wenzelm
parents: 32564
diff changeset
   244
      success=metis_success0, time=metis_time0, timeout=metis_timeout0,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   245
      lemmas=metis_lemmas0, posns=metis_posns0},
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   246
    MinData{succs=min_succs, ab_ratios=ab_ratios, it_ratios=it_ratios},
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   247
    MeData{calls=metis_calls, proofs=metis_proofs,
32567
de411627a985 explicitly export type abbreviations (as usual in SML97);
wenzelm
parents: 32564
diff changeset
   248
      success=metis_success, time=metis_time, timeout=metis_timeout,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   249
      lemmas=metis_lemmas, posns=metis_posns})) =
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   250
  if sh_calls > 0
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   251
  then
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   252
   (log ("\n\n\nReport #" ^ string_of_int id ^ ":\n");
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   253
    log_sh_data log sh_calls sh_success sh_lemmas sh_max_lems sh_time_isa sh_time_atp sh_time_atp_fail;
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   254
    log "";
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   255
    if metis_calls > 0 then log_metis_data log "" sh_calls sh_success metis_calls
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   256
      metis_success metis_proofs metis_time metis_timeout metis_lemmas metis_posns else ();
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   257
    log "";
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   258
    if metis_calls0 > 0
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   259
      then (log_min_data log min_succs ab_ratios it_ratios; log "";
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   260
            log_metis_data log "unminimized " sh_calls sh_success metis_calls0
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   261
              metis_success0 metis_proofs0 metis_time0 metis_timeout0 metis_lemmas0 metis_posns0)
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   262
      else ()
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   263
   )
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   264
  else ()
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   265
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   266
end
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   267
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   268
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   269
(* Warning: we implicitly assume single-threaded execution here! *)
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   270
val data = Unsynchronized.ref ([] : (int * data) list)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   271
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   272
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
   273
fun done id ({log, ...}: Mirabelle.done_args) =
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   274
  AList.lookup (op =) (!data) id
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   275
  |> Option.map (log_data id log)
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   276
  |> K ()
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   277
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   278
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
   279
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   280
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   281
fun get_atp thy args =
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   282
  AList.lookup (op =) args proverK
32937
34f66c9dd8a2 structure ATP_Manager: eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32936
diff changeset
   283
  |> the_default (hd (ATP_Manager.get_atps ()))  (* FIXME partial *)
32947
3c19b98a35cd ATP_Manager.get_prover: canonical argument order;
wenzelm
parents: 32941
diff changeset
   284
  |> (fn name => (name, the (ATP_Manager.get_prover thy name)))  (* FIXME partial *)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   285
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   286
local
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   287
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   288
datatype sh_result =
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   289
  SH_OK of int * int * string list |
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   290
  SH_FAIL of int * int |
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   291
  SH_ERROR
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   292
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32819
diff changeset
   293
fun run_sh prover hard_timeout timeout dir st =
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   294
  let
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32819
diff changeset
   295
    val (ctxt, goal) = Proof.get_goal st
32868
5f1805c6ef2a avoid exception Option: only apply "the" if needed
boehmes
parents: 32864
diff changeset
   296
    val ctxt' = if is_none dir then ctxt
32936
9491bec20595 modernized structure names;
wenzelm
parents: 32868
diff changeset
   297
      else Config.put ATP_Wrapper.destdir (the dir) ctxt
32948
e95a4be101a8 natural argument order for prover;
wenzelm
parents: 32947
diff changeset
   298
    val problem = ATP_Wrapper.problem_of_goal (! ATP_Manager.full_types) 1 (ctxt', goal);
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32819
diff changeset
   299
32574
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   300
    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
   301
      (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
   302
        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
   303
      | SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs))
32985
fbc642835ecb made SML/NJ happy;
wenzelm
parents: 32948
diff changeset
   304
    val ({success, message, theorem_names, runtime=time_atp, ...}: ATP_Wrapper.prover_result,
fbc642835ecb made SML/NJ happy;
wenzelm
parents: 32948
diff changeset
   305
        time_isa) = time_limit (Mirabelle.cpu_time (fn t => prover t problem)) timeout
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   306
  in
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32819
diff changeset
   307
    if success then (message, SH_OK (time_isa, time_atp, theorem_names))
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   308
    else (message, SH_FAIL(time_isa, time_atp))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   309
  end
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   310
  handle ResHolClause.TOO_TRIVIAL => ("trivial", SH_OK (0, 0, []))
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   311
       | ERROR msg => ("error: " ^ msg, SH_ERROR)
32574
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   312
       | TimeLimit.TimeOut => ("timeout", SH_ERROR)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   313
32454
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   314
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
   315
  let
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   316
    val lex = OuterKeyword.get_lexicons
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   317
    val get = maps (ProofContext.get_fact ctxt o fst)
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   318
  in
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   319
    Source.of_string name
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   320
    |> Symbol.source {do_recover=false}
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   321
    |> OuterLex.source {do_recover=SOME false} lex Position.start
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   322
    |> OuterLex.source_proper
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   323
    |> Source.source OuterLex.stopper (SpecParse.xthms1 >> get) NONE
a1a5589207ad Mirabelle: proper parsing of theorem names found by sledgehammer, respecting test intervals given along with file names
boehmes
parents: 32452
diff changeset
   324
    |> 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
   325
  end
32452
d84edd022efe apply metis with found theorems in case sledgehammer was successful
boehmes
parents: 32434
diff changeset
   326
32498
1132c7c13f36 Mirabelle: actions are responsible for handling exceptions,
boehmes
parents: 32496
diff changeset
   327
in
1132c7c13f36 Mirabelle: actions are responsible for handling exceptions,
boehmes
parents: 32496
diff changeset
   328
32567
de411627a985 explicitly export type abbreviations (as usual in SML97);
wenzelm
parents: 32564
diff changeset
   329
fun run_sledgehammer args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) =
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   330
  let
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   331
    val _ = change_data id inc_sh_calls
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32819
diff changeset
   332
    val (prover_name, prover) = get_atp (Proof.theory_of st) args
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   333
    val dir = AList.lookup (op =) args keepK
32541
cea1716eb106 timeout option for ATPs
boehmes
parents: 32536
diff changeset
   334
    val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30)
32574
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   335
    val hard_timeout = AList.lookup (op =) args prover_hard_timeoutK
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   336
      |> Option.map (fst o read_int o explode)
32864
a226f29d4bdc re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents: 32819
diff changeset
   337
    val (msg, result) = run_sh prover hard_timeout timeout dir st
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   338
  in
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   339
    case result of
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   340
      SH_OK (time_isa, time_atp, names) =>
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   341
        let fun get_thms name = (name, thms_of_name (Proof.context_of st) name)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   342
        in
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   343
          change_data id inc_sh_success;
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   344
          change_data id (inc_sh_lemmas (length names));
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   345
          change_data id (inc_sh_max_lems (length names));
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   346
          change_data id (inc_sh_time_isa time_isa);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   347
          change_data id (inc_sh_time_atp time_atp);
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   348
          named_thms := SOME (map get_thms names);
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   349
          log (sh_tag id ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   350
            string_of_int time_atp ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   351
        end
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   352
    | SH_FAIL (time_isa, time_atp) =>
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   353
        let
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   354
          val _ = change_data id (inc_sh_time_isa time_isa)
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   355
          val _ = change_data id (inc_sh_time_atp_fail time_atp)
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   356
        in log (sh_tag id ^ "failed: " ^ msg) end
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   357
    | SH_ERROR => log (sh_tag id ^ "failed: " ^ msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   358
  end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   359
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   360
end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   361
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   362
32567
de411627a985 explicitly export type abbreviations (as usual in SML97);
wenzelm
parents: 32564
diff changeset
   363
fun run_minimize args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) =
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   364
  let
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   365
    val n0 = length (these (!named_thms))
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   366
    val (prover_name, prover) = get_atp (Proof.theory_of st) args
32936
9491bec20595 modernized structure names;
wenzelm
parents: 32868
diff changeset
   367
    val minimize = ATP_Minimal.minimalize prover prover_name
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   368
    val timeout =
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   369
      AList.lookup (op =) args minimize_timeoutK
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   370
      |> Option.map (fst o read_int o explode)
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   371
      |> the_default 5
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   372
    val _ = log separator
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   373
  in
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   374
    case minimize timeout st (these (!named_thms)) of
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   375
      (SOME (named_thms',its), msg) =>
32609
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   376
        (change_data id inc_min_succs;
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   377
         change_data id (inc_min_ab_ratios ((100 * length named_thms') div n0));
2f3e7a92b522 modified minimization log
nipkow
parents: 32607
diff changeset
   378
         change_data id (inc_min_it_ratios ((100*its) div n0));
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   379
         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
   380
         then log (minimize_tag id ^ "already minimal")
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   381
         else (named_thms := SOME named_thms';
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   382
               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
   383
        )
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   384
    | (NONE, msg) => log (minimize_tag id ^ "failed: " ^ msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   385
  end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   386
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   387
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   388
fun run_metis (inc_metis_calls, inc_metis_success, inc_metis_proofs, inc_metis_time, inc_metis_timeout,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   389
    inc_metis_lemmas, inc_metis_posns) args name named_thms id
32567
de411627a985 explicitly export type abbreviations (as usual in SML97);
wenzelm
parents: 32564
diff changeset
   390
    ({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) =
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   391
  let
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   392
    fun metis thms ctxt = MetisTools.metis_tac ctxt thms
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   393
    fun apply_metis thms = Mirabelle.can_apply timeout (metis thms) st
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   394
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   395
    fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   396
      | with_time (true, t) = (change_data id inc_metis_success;
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
   397
          change_data id (inc_metis_lemmas (length named_thms));
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   398
          change_data id (inc_metis_time t);
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   399
          change_data id (inc_metis_posns pos);
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   400
          if name = "proof" then change_data id inc_metis_proofs else ();
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   401
          "succeeded (" ^ string_of_int t ^ ")")
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   402
    fun timed_metis thms = with_time (Mirabelle.cpu_time apply_metis thms)
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   403
      handle TimeLimit.TimeOut => (change_data id inc_metis_timeout; "timeout")
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   404
           | ERROR msg => "error: " ^ msg
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   405
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   406
    val _ = log separator
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   407
    val _ = change_data id inc_metis_calls
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   408
  in
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   409
    maps snd named_thms
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   410
    |> timed_metis
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   411
    |> log o prefix (metis_tag id) 
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   412
  end
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   413
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   414
fun sledgehammer_action args id (st as {log, pre, name, ...}: Mirabelle.run_args) =
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   415
  let val goal = Thm.major_prem_of(snd(snd(Proof.get_goal pre))) in
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   416
  if can Logic.dest_conjunction goal orelse can Logic.dest_equals goal
32607
e7fe01b74a92 skip &&& goals
nipkow
parents: 32598
diff changeset
   417
  then () else
32515
e7c0d3c0494a Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents: 32511
diff changeset
   418
  let
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   419
    val metis_fns = (inc_metis_calls, inc_metis_success, inc_metis_proofs, inc_metis_time,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   420
        inc_metis_timeout, inc_metis_lemmas, inc_metis_posns)
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   421
    val metis0_fns = (inc_metis_calls0, inc_metis_success0, inc_metis_proofs0, inc_metis_time0,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   422
        inc_metis_timeout0, inc_metis_lemmas0, inc_metis_posns0)
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32676
diff changeset
   423
    val named_thms = Unsynchronized.ref (NONE : (string * thm list) list option)
32612
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   424
    val minimize = AList.defined (op =) args minimizeK
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   425
  in 
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   426
    Mirabelle.catch sh_tag (run_sledgehammer args named_thms) id st;
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   427
    if is_some (!named_thms)
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   428
      then
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   429
       (if minimize
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   430
          then Mirabelle.catch metis_tag (run_metis metis0_fns args name (these (!named_thms))) id st
32612
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   431
          else ();
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   432
       if minimize andalso not(null(these(!named_thms)))
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   433
         then Mirabelle.catch minimize_tag (run_minimize args named_thms) id st
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   434
         else ();
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   435
       Mirabelle.catch metis_tag (run_metis metis_fns args name (these (!named_thms))) id st)
32612
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   436
    else ()
76dddd260d2f restructured code
nipkow
parents: 32609
diff changeset
   437
  end
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   438
  end
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   439
32511
43d2ac4aa2de added option full_typed for sledgehammer action
boehmes
parents: 32510
diff changeset
   440
fun invoke args =
32515
e7c0d3c0494a Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
boehmes
parents: 32511
diff changeset
   441
  let
32937
34f66c9dd8a2 structure ATP_Manager: eliminated slightly odd get/set operations in favour of Unsynchronized.ref;
wenzelm
parents: 32936
diff changeset
   442
    val _ = ATP_Manager.full_types := AList.defined (op =) args full_typesK
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   443
  in Mirabelle.register (init, sledgehammer_action args, done) end
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   444
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   445
end