src/HOL/Tools/Mirabelle/mirabelle_sledgehammer.ML
author blanchet
Fri, 25 Oct 2024 15:31:58 +0200
changeset 81254 d3c0734059ee
parent 78735 a0f85118488c
child 81362 f586fdabe670
permissions -rw-r--r--
variable instantiation in Sledgehammer and Metis
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
76183
8089593a364a proper file headers;
wenzelm
parents: 75372
diff changeset
     1
(*  Title:      HOL/Tools/Mirabelle/mirabelle_sledgehammer.ML
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
     2
    Author:     Jasmin Blanchette, TU Munich
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
     3
    Author:     Sascha Boehme, TU Munich
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
     4
    Author:     Tobias Nipkow, TU Munich
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
     5
    Author:     Makarius
74897
8b1ab558e3ee reused Sledgehammer code to parse parameters of sledgehammer action in Mirabelle
desharna
parents: 74892
diff changeset
     6
    Author:     Martin Desharnais, UniBw Munich, MPI-INF Saarbruecken
73691
2f9877db82a1 reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
wenzelm
parents: 73292
diff changeset
     7
2f9877db82a1 reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
wenzelm
parents: 73292
diff changeset
     8
Mirabelle action: "sledgehammer".
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
     9
*)
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
    10
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
    11
structure Mirabelle_Sledgehammer: MIRABELLE_ACTION =
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
    12
struct
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
    13
47480
24d8c9e9dae4 Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
sultana
parents: 47477
diff changeset
    14
(*To facilitate synching the description of Mirabelle Sledgehammer parameters
24d8c9e9dae4 Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
sultana
parents: 47477
diff changeset
    15
 (in ../lib/Tools/mirabelle) with the parameters actually used by this
24d8c9e9dae4 Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
sultana
parents: 47477
diff changeset
    16
 interface, the former extracts PARAMETER and DESCRIPTION from code below which
24d8c9e9dae4 Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
sultana
parents: 47477
diff changeset
    17
 has this pattern (provided it appears in a single line):
24d8c9e9dae4 Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
sultana
parents: 47477
diff changeset
    18
   val .*K = "PARAMETER" (*DESCRIPTION*)
24d8c9e9dae4 Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
sultana
parents: 47477
diff changeset
    19
*)
72342
4195e75a92ef [mirabelle] add initial documentation in Sledgehammer's doc
desharna
parents: 72173
diff changeset
    20
(* NOTE: Do not forget to update the Sledgehammer documentation to reflect changes here. *)
72173
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    21
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    22
val check_trivialK = "check_trivial" (*=BOOL: check if goals are "trivial"*)
75372
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
    23
val exhaustive_preplayK = "exhaustive_preplay" (*=BOOL: show exhaustive preplay data*)
74981
10df7a627ab6 split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents: 74967
diff changeset
    24
val keep_probsK = "keep_probs" (*=BOOL: keep temporary problem files created by sledgehammer*)
10df7a627ab6 split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents: 74967
diff changeset
    25
val keep_proofsK = "keep_proofs" (*=BOOL: keep temporary proof files created by ATPs*)
44448
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
    26
47481
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
    27
(*defaults used in this Mirabelle action*)
73292
f84a93f1de2f tuned Mirabelle to parse option check_trivial only once
desharna
parents: 73289
diff changeset
    28
val check_trivial_default = false
75372
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
    29
val exhaustive_preplay_default = false
74981
10df7a627ab6 split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents: 74967
diff changeset
    30
val keep_probs_default = false
10df7a627ab6 split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents: 74967
diff changeset
    31
val keep_proofs_default = false
46826
4c80c4034f1d added max_new_mono_instances, max_mono_iters, to Mirabelle-Sledgehammer; changed sh_minimize to avoid setting Mirabelle-level defaults;
sultana
parents: 46825
diff changeset
    32
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    33
datatype sh_data = ShData of {
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    34
  calls: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    35
  success: int,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    36
  nontriv_calls: int,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    37
  nontriv_success: int,
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
    38
  lemmas: int,
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    39
  max_lems: int,
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    40
  time_isa: int,
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    41
  time_prover: int}
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    42
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    43
fun make_sh_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    44
      (calls,success,nontriv_calls,nontriv_success,lemmas,max_lems,time_isa,
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    45
       time_prover) =
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    46
  ShData{calls=calls, success=success, nontriv_calls=nontriv_calls,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    47
         nontriv_success=nontriv_success, lemmas=lemmas, max_lems=max_lems,
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    48
         time_isa=time_isa, time_prover=time_prover}
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    49
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    50
val empty_sh_data = make_sh_data (0, 0, 0, 0, 0, 0, 0, 0)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    51
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    52
fun tuple_of_sh_data (ShData {calls, success, nontriv_calls, nontriv_success, lemmas, max_lems,
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    53
    time_isa, time_prover}) =
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    54
  (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    55
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
    56
fun map_sh_data f sh = make_sh_data (f (tuple_of_sh_data sh))
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
    57
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
    58
val inc_sh_calls = map_sh_data
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    59
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover) =>
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    60
    (calls + 1, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    61
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    62
val inc_sh_success = map_sh_data
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    63
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover) =>
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    64
    (calls, success + 1, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover))
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    65
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    66
val inc_sh_nontriv_calls =  map_sh_data
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    67
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover) =>
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    68
    (calls, success, nontriv_calls + 1, nontriv_success, lemmas, max_lems, time_isa, time_prover))
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    69
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    70
val inc_sh_nontriv_success = map_sh_data
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    71
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover) =>
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    72
    (calls, success, nontriv_calls, nontriv_success + 1, lemmas,max_lems, time_isa, time_prover))
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
    73
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    74
fun inc_sh_lemmas n = map_sh_data
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    75
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover) =>
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    76
    (calls, success, nontriv_calls, nontriv_success, lemmas+n, max_lems, time_isa, time_prover))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    77
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    78
fun inc_sh_max_lems n = map_sh_data
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    79
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover) =>
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    80
    (calls, success,nontriv_calls, nontriv_success, lemmas, Int.max (max_lems, n), time_isa,
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    81
     time_prover))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    82
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    83
fun inc_sh_time_isa t = map_sh_data
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    84
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover) =>
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    85
    (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa + t, time_prover))
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    86
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    87
fun inc_sh_time_prover t = map_sh_data
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    88
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover) =>
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
    89
    (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover + t))
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
    90
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
    91
val str0 = string_of_int o the_default 0
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
    92
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    93
local
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    94
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    95
val str = string_of_int
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    96
val str3 = Real.fmt (StringCvt.FIX (SOME 3))
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    97
fun percentage a b = string_of_int (a * 100 div b)
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
    98
fun ms t = Real.fromInt t / 1000.0
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    99
fun avg_time t n =
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   100
  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
   101
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   102
fun log_sh_data (ShData {calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa,
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
   103
      time_prover}) =
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   104
  "\nTotal number of sledgehammer calls: " ^ str calls ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   105
  "\nNumber of successful sledgehammer calls: " ^ str success ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   106
  "\nNumber of sledgehammer lemmas: " ^ str lemmas ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   107
  "\nMax number of sledgehammer lemmas: " ^ str max_lems ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   108
  "\nSuccess rate: " ^ percentage success calls ^ "%" ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   109
  "\nTotal number of nontrivial sledgehammer calls: " ^ str nontriv_calls ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   110
  "\nNumber of successful nontrivial sledgehammer calls: " ^ str nontriv_success ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   111
  "\nTotal time for sledgehammer calls (Isabelle): " ^ str3 (ms time_isa) ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   112
  "\nTotal time for successful sledgehammer calls (ATP): " ^ str3 (ms time_prover) ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   113
  "\nAverage time for sledgehammer calls (Isabelle): " ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   114
    str3 (avg_time time_isa calls) ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   115
  "\nAverage time for successful sledgehammer calls (ATP): " ^
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
   116
    str3 (avg_time time_prover success)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   117
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   118
in
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   119
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   120
fun log_data (sh as ShData {calls=sh_calls, ...}) =
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   121
  if sh_calls > 0 then
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   122
    log_sh_data sh
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   123
  else
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   124
    ""
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   125
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   126
end
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   127
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   128
local
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   129
75343
62f0fda48a39 compile mirabelle
blanchet
parents: 75038
diff changeset
   130
fun run_sh params keep pos state =
78726
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   131
  \<^try>\<open>
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   132
    let
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   133
      fun set_file_name (SOME (dir, keep_probs, keep_proofs)) =
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   134
          let
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   135
            val filename = "prob_" ^
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   136
              StringCvt.padLeft #"0" 5 (str0 (Position.line_of pos)) ^ "_" ^
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   137
              StringCvt.padLeft #"0" 6 (str0 (Position.offset_of pos))
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   138
          in
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   139
            Config.put Sledgehammer_Prover_ATP.atp_problem_prefix (filename ^ "__")
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   140
            #> (keep_probs ? Config.put Sledgehammer_Prover_ATP.atp_problem_dest_dir dir)
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   141
            #> (keep_proofs ? Config.put Sledgehammer_Prover_ATP.atp_proof_dest_dir dir)
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   142
            #> Config.put SMT_Config.debug_files (dir ^ "/" ^ filename ^ "__" ^ serial_string ())
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   143
          end
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   144
        | set_file_name NONE = I
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   145
      val state' = state
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   146
        |> Proof.map_context (set_file_name keep)
74952
ae2185967e67 exported Sledgehammer.launch_prover and use it in Mirabelle
desharna
parents: 74951
diff changeset
   147
78726
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   148
      val ((_, (sledgehammer_outcome, msg)), cpu_time) = Mirabelle.cpu_time (fn () =>
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   149
        Sledgehammer.run_sledgehammer params Sledgehammer_Prover.Normal NONE 1
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   150
          Sledgehammer_Fact.no_fact_override state') ()
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   151
    in
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   152
      (sledgehammer_outcome, msg, cpu_time)
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   153
    end
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   154
    catch ERROR msg => (Sledgehammer.SH_Unknown, " error: " ^ msg, 0)
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   155
      | _ => (Sledgehammer.SH_Unknown, " error: unexpected error", 0)
810eca753919 tuned: prefer try-catch/finally over low-level 'handle';
wenzelm
parents: 76945
diff changeset
   156
  \<close>
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   157
32498
1132c7c13f36 Mirabelle: actions are responsible for handling exceptions,
boehmes
parents: 32496
diff changeset
   158
in
1132c7c13f36 Mirabelle: actions are responsible for handling exceptions,
boehmes
parents: 32496
diff changeset
   159
75343
62f0fda48a39 compile mirabelle
blanchet
parents: 75038
diff changeset
   160
fun run_sledgehammer (params as {provers, ...}) output_dir keep_probs keep_proofs
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   161
    exhaustive_preplay thy_index trivial pos st =
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   162
  let
39340
3998dc0bf82b indicate which goals are trivial
blanchet
parents: 39337
diff changeset
   163
    val triv_str = if trivial then "[T] " else ""
74981
10df7a627ab6 split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents: 74967
diff changeset
   164
    val keep =
10df7a627ab6 split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents: 74967
diff changeset
   165
      if keep_probs orelse keep_proofs then
78735
a0f85118488c moved variable bindings to tighter scope
desharna
parents: 78734
diff changeset
   166
        let
a0f85118488c moved variable bindings to tighter scope
desharna
parents: 78734
diff changeset
   167
          val thy_long_name = Context.theory_long_name (Proof.theory_of st)
a0f85118488c moved variable bindings to tighter scope
desharna
parents: 78734
diff changeset
   168
          val session_name = Long_Name.qualifier thy_long_name
a0f85118488c moved variable bindings to tighter scope
desharna
parents: 78734
diff changeset
   169
          val thy_name = Long_Name.base_name thy_long_name
a0f85118488c moved variable bindings to tighter scope
desharna
parents: 78734
diff changeset
   170
          val subdir = StringCvt.padLeft #"0" 4 (string_of_int thy_index) ^ "_" ^ thy_name
a0f85118488c moved variable bindings to tighter scope
desharna
parents: 78734
diff changeset
   171
        in
76945
fcd1df8f48fc added session to mirabelle output directory structure
desharna
parents: 76525
diff changeset
   172
          Path.append (Path.append output_dir (Path.basic session_name)) (Path.basic subdir)
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   173
          |> Isabelle_System.make_directory
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   174
          |> Path.implode
74981
10df7a627ab6 split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents: 74967
diff changeset
   175
          |> (fn dir => SOME (dir, keep_probs, keep_proofs))
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   176
        end
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   177
      else
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   178
        NONE
74897
8b1ab558e3ee reused Sledgehammer code to parse parameters of sledgehammer action in Mirabelle
desharna
parents: 74892
diff changeset
   179
    val prover_name = hd provers
75343
62f0fda48a39 compile mirabelle
blanchet
parents: 75038
diff changeset
   180
    val (sledgehamer_outcome, msg, cpu_time) = run_sh params keep pos st
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   181
    val (time_prover, change_data, exhaustive_preplay_msg) =
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
   182
      (case sledgehamer_outcome of
75372
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   183
        Sledgehammer.SH_Some ({used_facts, run_time, ...}, preplay_results) =>
38700
fcb0fe4c2f27 make Mirabelle happy
blanchet
parents: 38590
diff changeset
   184
        let
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
   185
          val num_used_facts = length used_facts
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
   186
          val time_prover = Time.toMilliseconds run_time
74996
1f4c39ffb116 tuned mirabelle_sledgehammer to have a single call to Synchronized.change per run
desharna
parents: 74981
diff changeset
   187
          val change_data =
1f4c39ffb116 tuned mirabelle_sledgehammer to have a single call to Synchronized.change per run
desharna
parents: 74981
diff changeset
   188
            inc_sh_success
1f4c39ffb116 tuned mirabelle_sledgehammer to have a single call to Synchronized.change per run
desharna
parents: 74981
diff changeset
   189
            #> not trivial ? inc_sh_nontriv_success
1f4c39ffb116 tuned mirabelle_sledgehammer to have a single call to Synchronized.change per run
desharna
parents: 74981
diff changeset
   190
            #> inc_sh_lemmas num_used_facts
1f4c39ffb116 tuned mirabelle_sledgehammer to have a single call to Synchronized.change per run
desharna
parents: 74981
diff changeset
   191
            #> inc_sh_max_lems num_used_facts
1f4c39ffb116 tuned mirabelle_sledgehammer to have a single call to Synchronized.change per run
desharna
parents: 74981
diff changeset
   192
            #> inc_sh_time_prover time_prover
75372
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   193
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   194
          val exhaustive_preplay_msg =
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   195
            if exhaustive_preplay then
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   196
              preplay_results
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   197
              |> map
76525
eb294dd8e266 compile
blanchet
parents: 76183
diff changeset
   198
                (fn (meth, (play_outcome, used_facts)) =>
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 78735
diff changeset
   199
                    "Preplay: " ^ Sledgehammer_Proof_Methods.string_of_proof_method
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 78735
diff changeset
   200
                      (map (ATP_Util.content_of_pretty o fst) used_facts) meth ^
75372
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   201
                    " (" ^ Sledgehammer_Proof_Methods.string_of_play_outcome play_outcome ^ ")")
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   202
              |> cat_lines
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   203
            else
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   204
              ""
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   205
        in
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   206
          (SOME time_prover, change_data, exhaustive_preplay_msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   207
        end
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   208
      | _ => (NONE, I, ""))
75001
d9a5824f68c8 tuned mirabelle_sledgehammer output
desharna
parents: 74997
diff changeset
   209
    val outcome_msg =
d9a5824f68c8 tuned mirabelle_sledgehammer output
desharna
parents: 74997
diff changeset
   210
      "(SH " ^ string_of_int cpu_time ^ "ms" ^
d9a5824f68c8 tuned mirabelle_sledgehammer output
desharna
parents: 74997
diff changeset
   211
      (case time_prover of NONE => "" | SOME ms => ", ATP " ^ string_of_int ms ^ "ms") ^
d9a5824f68c8 tuned mirabelle_sledgehammer output
desharna
parents: 74997
diff changeset
   212
      ") [" ^ prover_name ^ "]: "
74953
aade20a03edb tuned run_sledgehammer and called it directly from Mirabelle
desharna
parents: 74952
diff changeset
   213
  in
81254
d3c0734059ee variable instantiation in Sledgehammer and Metis
blanchet
parents: 78735
diff changeset
   214
    (sledgehamer_outcome, triv_str ^ outcome_msg ^ Protocol_Message.clean_output msg ^
75372
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   215
       (if exhaustive_preplay_msg = "" then "" else ("\n" ^ exhaustive_preplay_msg)),
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   216
     change_data #> inc_sh_time_isa cpu_time)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   217
  end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   218
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   219
end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   220
75004
8dc52ba4155b tuned trivial check in mirabelle_sledgehammer
desharna
parents: 75003
diff changeset
   221
val try0 = Try0.try0 (SOME (Time.fromSeconds 5)) ([], [], [], [])
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   222
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   223
fun make_action ({arguments, timeout, output_dir, ...} : Mirabelle.action_context) =
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   224
  let
74897
8b1ab558e3ee reused Sledgehammer code to parse parameters of sledgehammer action in Mirabelle
desharna
parents: 74892
diff changeset
   225
    (* Parse Mirabelle-specific parameters *)
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   226
    val check_trivial =
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   227
      Mirabelle.get_bool_argument arguments (check_trivialK, check_trivial_default)
74981
10df7a627ab6 split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents: 74967
diff changeset
   228
    val keep_probs = Mirabelle.get_bool_argument arguments (keep_probsK, keep_probs_default)
10df7a627ab6 split option "sledgehammer_atp_dest_dir" into "sledgehammer_atp_prob_dest_dir" and "sledgehammer_atp_proof_dest_dir"
desharna
parents: 74967
diff changeset
   229
    val keep_proofs = Mirabelle.get_bool_argument arguments (keep_proofsK, keep_proofs_default)
75372
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   230
    val exhaustive_preplay =
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   231
      Mirabelle.get_bool_argument arguments (exhaustive_preplayK, exhaustive_preplay_default)
74897
8b1ab558e3ee reused Sledgehammer code to parse parameters of sledgehammer action in Mirabelle
desharna
parents: 74892
diff changeset
   232
8b1ab558e3ee reused Sledgehammer code to parse parameters of sledgehammer action in Mirabelle
desharna
parents: 74892
diff changeset
   233
    val params = Sledgehammer_Commands.default_params \<^theory> arguments
44431
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   234
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   235
    val data = Synchronized.var "Mirabelle_Sledgehammer.data" empty_sh_data
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   236
74948
15ce207f69c8 added support for initialization messages to Mirabelle
desharna
parents: 74897
diff changeset
   237
    val init_msg = "Params for sledgehammer: " ^ Sledgehammer_Prover.string_of_params params
15ce207f69c8 added support for initialization messages to Mirabelle
desharna
parents: 74897
diff changeset
   238
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   239
    fun run ({theory_index, pos, pre, ...} : Mirabelle.command) =
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   240
      let val goal = Thm.major_prem_of (#goal (Proof.goal pre)) in
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   241
        if can Logic.dest_conjunction goal orelse can Logic.dest_equals goal then
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   242
          ""
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   243
        else
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   244
          let
75004
8dc52ba4155b tuned trivial check in mirabelle_sledgehammer
desharna
parents: 75003
diff changeset
   245
            val trivial = check_trivial andalso try0 pre handle Timeout.TIMEOUT _ => false
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   246
            val (outcome, log, change_data) =
75372
4c8d1ef258d3 added preplay results to sledgehammer_output
desharna
parents: 75343
diff changeset
   247
              run_sledgehammer params output_dir keep_probs keep_proofs exhaustive_preplay
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   248
                theory_index trivial pos pre
74996
1f4c39ffb116 tuned mirabelle_sledgehammer to have a single call to Synchronized.change per run
desharna
parents: 74981
diff changeset
   249
            val () = Synchronized.change data
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   250
              (change_data #> inc_sh_calls #> not trivial ? inc_sh_nontriv_calls)
74967
3f55c5feca58 prefixed all mirabelle_sledgehammer output lines with sledgehammer output
desharna
parents: 74953
diff changeset
   251
          in
78734
046e2ddff9ba removed proof reconstruction from Mirabelle; this is best handled directly in Sledgehammer
desharna
parents: 78733
diff changeset
   252
            log
74967
3f55c5feca58 prefixed all mirabelle_sledgehammer output lines with sledgehammer output
desharna
parents: 74953
diff changeset
   253
            |> Symbol.trim_blanks
3f55c5feca58 prefixed all mirabelle_sledgehammer output lines with sledgehammer output
desharna
parents: 74953
diff changeset
   254
            |> prefix_lines (Sledgehammer.short_string_of_sledgehammer_outcome outcome ^ " ")
3f55c5feca58 prefixed all mirabelle_sledgehammer output lines with sledgehammer output
desharna
parents: 74953
diff changeset
   255
          end
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   256
      end
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   257
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   258
    fun finalize () = log_data (Synchronized.value data)
75003
f21e7e6172a0 renamed run_action to run in Mirabelle.action record
desharna
parents: 75001
diff changeset
   259
  in (init_msg, {run = run, finalize = finalize}) end
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   260
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   261
val () = Mirabelle.register_action "sledgehammer" make_action
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   262
74951
0b6f795d3b78 proper filtering inf induction rules in Mirabelle
desharna
parents: 74950
diff changeset
   263
end