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