src/HOL/Tools/Mirabelle/mirabelle_sledgehammer.ML
author wenzelm
Fri, 29 Oct 2021 11:59:02 +0200
changeset 74620 d622d1dce05c
parent 74472 9d304ef5c932
child 74892 3094dae03764
permissions -rw-r--r--
avoid persistence of static context: instantiation arguments should provide proper dynamic context (see also e2e2bc1f9570);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47847
7cddb6c8f93c updated headers;
wenzelm
parents: 47732
diff changeset
     1
(*  Title:      HOL/Mirabelle/Tools/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
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
     6
    Author:     Martin Desharnais, UniBw Munich
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"*)
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    23
val e_selection_heuristicK = "e_selection_heuristic" (*=STRING: E clause selection heuristic*)
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    24
val fact_filterK = "fact_filter" (*=STRING: fact filter*)
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    25
val force_sosK = "force_sos" (*=BOOL: use set-of-support (in Vampire)*)
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    26
val isar_proofsK = "isar_proofs" (*=SMART_BOOL: enable Isar proof generation*)
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
    27
val keepK = "keep" (*=BOOL: keep temporary files created by sledgehammer*)
72173
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    28
val lam_transK = "lam_trans" (*=STRING: lambda translation scheme*)
48293
914ca0827804 renamed Sledgehammer options
blanchet
parents: 48292
diff changeset
    29
val max_factsK = "max_facts" (*=NUM: max. relevant clauses to use*)
72173
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    30
val max_mono_itersK = "max_mono_iters" (*=NUM: max. iterations of monomorphiser*)
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    31
val max_new_mono_instancesK = "max_new_mono_instances" (*=NUM: max. new monomorphic instances*)
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
    32
val max_relevantK = "max_relevant" (*=NUM: max. relevant clauses to use*)
72173
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    33
val minimizeK = "minimize" (*=BOOL: instruct sledgehammer to run its minimizer*)
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
    34
val preplay_timeoutK = "preplay_timeout" (*=TIME: timeout for finding reconstructed proof*)
73849
4eac16052a94 tuned Mirabelle
desharna
parents: 73847
diff changeset
    35
val proof_methodK = "proof_method" (*=STRING: how to reconstruct proofs (e.g. using metis)*)
72173
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    36
val proverK = "prover" (*=STRING: name of the external prover to call*)
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    37
val prover_timeoutK = "prover_timeout" (*=TIME: timeout for invoked ATP (seconds of process time)*)
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    38
val sliceK = "slice" (*=BOOL: allow sledgehammer-level strategy-scheduling*)
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    39
val smt_proofsK = "smt_proofs" (*=BOOL: enable SMT proof generation*)
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    40
val strictK = "strict" (*=BOOL: run in strict mode*)
618a0ab13868 tuned Mirabelle comments
blanchet
parents: 69706
diff changeset
    41
val term_orderK = "term_order" (*=STRING: term order (in E)*)
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
    42
val type_encK = "type_enc" (*=STRING: type encoding scheme*)
24d8c9e9dae4 Mirabelle's 'usage' description relating to Sledgehammer now synchs with the ML code of the Mirabelle-Sledgehammer action
sultana
parents: 47477
diff changeset
    43
val uncurried_aliasesK = "uncurried_aliases" (*=SMART_BOOL: use fresh function names to alias curried applications*)
44448
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
    44
47481
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
    45
(*FIXME sensible to have Mirabelle-level Sledgehammer defaults?*)
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
    46
(*defaults used in this Mirabelle action*)
57751
f453bbc289fa modernized Mirabelle (a bit) and made it compile
blanchet
parents: 57154
diff changeset
    47
val preplay_timeout_default = "1"
47481
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
    48
val lam_trans_default = "smart"
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
    49
val uncurried_aliases_default = "smart"
50334
74d453d1b084 added "fact_filter" option to Mirabelle
blanchet
parents: 50267
diff changeset
    50
val fact_filter_default = "smart"
47481
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
    51
val type_enc_default = "smart"
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
    52
val strict_default = "false"
48293
914ca0827804 renamed Sledgehammer options
blanchet
parents: 48292
diff changeset
    53
val max_facts_default = "smart"
47481
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
    54
val slice_default = "true"
73292
f84a93f1de2f tuned Mirabelle to parse option check_trivial only once
desharna
parents: 73289
diff changeset
    55
val check_trivial_default = false
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
    56
val keep_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
    57
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
    58
(*If a key is present in args then augment a list with its pair*)
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
    59
(*This is used to avoid fixing default values at the Mirabelle level, and
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
    60
  instead use the default values of the tool (Sledgehammer in this case).*)
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
    61
fun available_parameter args key label list =
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
    62
  let
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
    63
    val value = AList.lookup (op =) args key
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
    64
  in if is_some value then (label, the value) :: list else list end
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
    65
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    66
datatype sh_data = ShData of {
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    67
  calls: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    68
  success: int,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    69
  nontriv_calls: int,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    70
  nontriv_success: int,
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
    71
  lemmas: int,
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    72
  max_lems: int,
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    73
  time_isa: int,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    74
  time_prover: int,
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    75
  time_prover_fail: int}
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    76
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    77
datatype re_data = ReData of {
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    78
  calls: int,
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    79
  success: int,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    80
  nontriv_calls: int,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    81
  nontriv_success: int,
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
    82
  proofs: int,
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    83
  time: int,
32550
d57c7a2d927c logging number of metis lemmas
nipkow
parents: 32549
diff changeset
    84
  timeout: int,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
    85
  lemmas: int * int * int,
39341
d2b981a0429a indicate triviality in the list of proved things
blanchet
parents: 39340
diff changeset
    86
  posns: (Position.T * bool) list
32550
d57c7a2d927c logging number of metis lemmas
nipkow
parents: 32549
diff changeset
    87
  }
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
    88
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
    89
fun make_sh_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    90
      (calls,success,nontriv_calls,nontriv_success,lemmas,max_lems,time_isa,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    91
       time_prover,time_prover_fail) =
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    92
  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
    93
         nontriv_success=nontriv_success, lemmas=lemmas, max_lems=max_lems,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    94
         time_isa=time_isa, time_prover=time_prover,
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
    95
         time_prover_fail=time_prover_fail}
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
    96
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    97
fun make_re_data (calls,success,nontriv_calls,nontriv_success,proofs,time,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
    98
                  timeout,lemmas,posns) =
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
    99
  ReData{calls=calls, success=success, nontriv_calls=nontriv_calls,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   100
         nontriv_success=nontriv_success, proofs=proofs, time=time,
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   101
         timeout=timeout, lemmas=lemmas, posns=posns}
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   102
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   103
val empty_sh_data = make_sh_data (0, 0, 0, 0, 0, 0, 0, 0, 0)
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   104
val empty_re_data = make_re_data (0, 0, 0, 0, 0, 0, 0, (0,0,0), [])
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   105
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   106
fun tuple_of_sh_data (ShData {calls, success, nontriv_calls, nontriv_success,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   107
                              lemmas, max_lems, time_isa,
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   108
  time_prover, time_prover_fail}) = (calls, success, nontriv_calls,
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   109
  nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   110
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   111
fun tuple_of_re_data (ReData {calls, success, nontriv_calls, nontriv_success,
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   112
  proofs, time, timeout, lemmas, posns}) = (calls, success, nontriv_calls,
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   113
  nontriv_success, proofs, time, timeout, lemmas, posns)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   114
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   115
datatype data = Data of {
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   116
  sh: sh_data,
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   117
  re_u: re_data (* proof method with unminimized set of lemmas *)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   118
  }
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   119
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   120
type change_data = (data -> data) -> unit
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   121
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   122
fun make_data (sh, re_u) = Data {sh=sh, re_u=re_u}
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   123
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   124
val empty_data = make_data (empty_sh_data, empty_re_data)
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   125
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   126
fun map_sh_data f (Data {sh, re_u}) =
57752
708347f9bfc6 removed Mirabelle minimization code
blanchet
parents: 57751
diff changeset
   127
  let val sh' = make_sh_data (f (tuple_of_sh_data sh))
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   128
  in make_data (sh', re_u) end
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   129
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   130
fun map_re_data f (Data {sh, re_u}) =
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   131
  let
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   132
    val f' = make_re_data o f o tuple_of_re_data
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   133
    val re_u' = f' re_u
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   134
  in make_data (sh, re_u') end
32990
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   135
717680b14041 added sums of squares for standard deviation
nipkow
parents: 32868
diff changeset
   136
fun inc_max (n:int) (s,sos,m) = (s+n, sos + n*n, Int.max(m,n));
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   137
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   138
val inc_sh_calls =  map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   139
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   140
    => (calls + 1, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   141
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   142
val inc_sh_success = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   143
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   144
    => (calls, success + 1, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail))
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   145
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   146
val inc_sh_nontriv_calls =  map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   147
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   148
    => (calls, success, nontriv_calls + 1, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail))
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   149
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   150
val inc_sh_nontriv_success = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   151
  (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   152
    => (calls, success, nontriv_calls, nontriv_success + 1, lemmas,max_lems, time_isa, time_prover, time_prover_fail))
32585
e788b33dd2b4 revised lemma counting
nipkow
parents: 32574
diff changeset
   153
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   154
fun inc_sh_lemmas n = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   155
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   156
    => (calls,success,nontriv_calls, nontriv_success, lemmas+n,max_lems,time_isa,time_prover,time_prover_fail))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   157
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   158
fun inc_sh_max_lems n = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   159
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   160
    => (calls,success,nontriv_calls, nontriv_success, lemmas,Int.max(max_lems,n),time_isa,time_prover,time_prover_fail))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   161
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   162
fun inc_sh_time_isa t = map_sh_data
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   163
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   164
    => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa + t,time_prover,time_prover_fail))
32818
6c91e668b15e record max lemmas used
nipkow
parents: 32676
diff changeset
   165
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   166
fun inc_sh_time_prover t = map_sh_data
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   167
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   168
    => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover + t,time_prover_fail))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   169
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   170
fun inc_sh_time_prover_fail t = map_sh_data
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   171
  (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   172
    => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail + t))
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   173
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   174
val inc_proof_method_calls = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   175
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   176
    => (calls + 1, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))
32533
9735013cec72 tuned stats
nipkow
parents: 32526
diff changeset
   177
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   178
val inc_proof_method_success = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   179
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   180
    => (calls, success + 1, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   181
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   182
val inc_proof_method_nontriv_calls = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   183
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   184
    => (calls, success, nontriv_calls + 1, nontriv_success, proofs, time, timeout, lemmas,posns))
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   185
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   186
val inc_proof_method_nontriv_success = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   187
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   188
    => (calls, success, nontriv_calls, nontriv_success + 1, proofs, time, timeout, lemmas,posns))
32676
b1c85a117dec record how many "proof"s are solved by s/h
nipkow
parents: 32612
diff changeset
   189
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   190
val inc_proof_method_proofs = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   191
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   192
    => (calls, success, nontriv_calls, nontriv_success, proofs + 1, time, timeout, lemmas,posns))
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   193
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   194
fun inc_proof_method_time t = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   195
 (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   196
    => (calls, success, nontriv_calls, nontriv_success, proofs, time + t, timeout, lemmas,posns))
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   197
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   198
val inc_proof_method_timeout = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   199
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   200
    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout + 1, lemmas,posns))
32549
338ccfd37f67 minimization: comparing w/ and w/o.
nipkow
parents: 32541
diff changeset
   201
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   202
fun inc_proof_method_lemmas n = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   203
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   204
    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, inc_max n lemmas, posns))
32551
421323205efd position information is now passed to all actions;
nipkow
parents: 32550
diff changeset
   205
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   206
fun inc_proof_method_posns pos = map_re_data
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   207
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
57753
643e02500991 removed 'metisFT' support in Mirabelle
blanchet
parents: 57752
diff changeset
   208
    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas, pos::posns))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   209
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   210
val str0 = string_of_int o the_default 0
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   211
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   212
local
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   213
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   214
val str = string_of_int
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   215
val str3 = Real.fmt (StringCvt.FIX (SOME 3))
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   216
fun percentage a b = string_of_int (a * 100 div b)
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   217
fun ms t = Real.fromInt t / 1000.0
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   218
fun avg_time t n =
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   219
  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
   220
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   221
fun log_sh_data (ShData {calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa,
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   222
      time_prover, time_prover_fail}) =
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   223
  "\nTotal number of sledgehammer calls: " ^ str calls ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   224
  "\nNumber of successful sledgehammer calls: " ^ str success ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   225
  "\nNumber of sledgehammer lemmas: " ^ str lemmas ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   226
  "\nMax number of sledgehammer lemmas: " ^ str max_lems ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   227
  "\nSuccess rate: " ^ percentage success calls ^ "%" ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   228
  "\nTotal number of nontrivial sledgehammer calls: " ^ str nontriv_calls ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   229
  "\nNumber of successful nontrivial sledgehammer calls: " ^ str nontriv_success ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   230
  "\nTotal time for sledgehammer calls (Isabelle): " ^ str3 (ms time_isa) ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   231
  "\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
   232
  "\nTotal time for failed sledgehammer calls (ATP): " ^ str3 (ms time_prover_fail) ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   233
  "\nAverage time for sledgehammer calls (Isabelle): " ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   234
    str3 (avg_time time_isa calls) ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   235
  "\nAverage time for successful sledgehammer calls (ATP): " ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   236
    str3 (avg_time time_prover success) ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   237
  "\nAverage time for failed sledgehammer calls (ATP): " ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   238
    str3 (avg_time time_prover_fail (calls - success))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   239
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   240
fun log_re_data sh_calls (ReData {calls, success, nontriv_calls, nontriv_success, proofs, time,
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   241
      timeout, lemmas = (lemmas, lems_sos, lems_max), posns}) =
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   242
  let
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   243
    val proved =
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   244
      posns |> map (fn (pos, triv) =>
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   245
        str0 (Position.line_of pos) ^ ":" ^ str0 (Position.offset_of pos) ^
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   246
        (if triv then "[T]" else ""))
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   247
  in
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   248
    "\nTotal number of proof method calls: " ^ str calls ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   249
    "\nNumber of successful proof method calls: " ^ str success ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   250
      " (proof: " ^ str proofs ^ ")" ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   251
    "\nNumber of proof method timeouts: " ^ str timeout ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   252
    "\nSuccess rate: " ^ percentage success sh_calls ^ "%" ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   253
    "\nTotal number of nontrivial proof method calls: " ^ str nontriv_calls ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   254
    "\nNumber of successful nontrivial proof method calls: " ^ str nontriv_success ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   255
      " (proof: " ^ str proofs ^ ")" ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   256
    "\nNumber of successful proof method lemmas: " ^ str lemmas ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   257
    "\nSOS of successful proof method lemmas: " ^ str lems_sos ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   258
    "\nMax number of successful proof method lemmas: " ^ str lems_max ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   259
    "\nTotal time for successful proof method calls: " ^ str3 (ms time) ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   260
    "\nAverage time for successful proof method calls: " ^ str3 (avg_time time success) ^
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   261
    "\nProved: " ^ space_implode " " proved
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   262
  end
32571
d4bb776874b8 count number of iterations required for minimization (and fixed bug: minimization was always called)
nipkow
parents: 32567
diff changeset
   263
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   264
in
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   265
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   266
fun log_data (Data {sh, re_u}) =
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   267
  let
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   268
    val ShData {calls=sh_calls, ...} = sh
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   269
    val ReData {calls=re_calls, ...} = re_u
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   270
  in
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   271
    if sh_calls > 0 then
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   272
      let val text1 = log_sh_data sh in
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   273
        if re_calls > 0 then text1 ^ "\n" ^ log_re_data sh_calls re_u else text1
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   274
      end
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   275
    else
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   276
      ""
34035
08d34921b7dd also consider the fully-typed version of metis for Mirabelle measurements
boehmes
parents: 33316
diff changeset
   277
  end
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   278
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   279
end
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   280
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   281
fun get_prover_name thy args =
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   282
  let
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   283
    fun default_prover_name () =
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   284
      hd (#provers (Sledgehammer_Commands.default_params thy []))
63692
1bc4bc2c9fd1 killed final stops in Sledgehammer and friends
blanchet
parents: 62969
diff changeset
   285
      handle List.Empty => error "No ATP available"
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   286
  in
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   287
    (case AList.lookup (op =) args proverK of
50352
db8cae658807 added MaSh learning to Mirabelle
blanchet
parents: 50334
diff changeset
   288
      SOME name => name
db8cae658807 added MaSh learning to Mirabelle
blanchet
parents: 50334
diff changeset
   289
    | NONE => default_prover_name ())
db8cae658807 added MaSh learning to Mirabelle
blanchet
parents: 50334
diff changeset
   290
  end
db8cae658807 added MaSh learning to Mirabelle
blanchet
parents: 50334
diff changeset
   291
69706
6d6235b828fc get rid of visibility in MaSh -- it slows it down more than it helps
blanchet
parents: 67405
diff changeset
   292
fun get_prover ctxt name params goal =
50352
db8cae658807 added MaSh learning to Mirabelle
blanchet
parents: 50334
diff changeset
   293
  let
69706
6d6235b828fc get rid of visibility in MaSh -- it slows it down more than it helps
blanchet
parents: 67405
diff changeset
   294
    val learn = Sledgehammer_MaSh.mash_learn_proof ctxt params (Thm.prop_of goal)
50352
db8cae658807 added MaSh learning to Mirabelle
blanchet
parents: 50334
diff changeset
   295
  in
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55201
diff changeset
   296
    Sledgehammer_Prover_Minimize.get_minimizing_prover ctxt Sledgehammer_Prover.Normal learn name
33016
b73b74fe23c3 proper exceptions instead of unhandled partiality
boehmes
parents: 32991
diff changeset
   297
  end
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   298
46340
cac402c486b0 separate orthogonal components
blanchet
parents: 46320
diff changeset
   299
type stature = ATP_Problem_Generate.stature
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38700
diff changeset
   300
58473
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   301
fun is_good_line s =
53081
2a62d848a56a improved ad hoc success detection in Mirabelle -- if the metis call fails and the structured proof succeeds, remember only the success
blanchet
parents: 52555
diff changeset
   302
  (String.isSubstring " ms)" s orelse String.isSubstring " s)" s)
2a62d848a56a improved ad hoc success detection in Mirabelle -- if the metis call fails and the structured proof succeeds, remember only the success
blanchet
parents: 52555
diff changeset
   303
  andalso not (String.isSubstring "(> " s)
2a62d848a56a improved ad hoc success detection in Mirabelle -- if the metis call fails and the structured proof succeeds, remember only the success
blanchet
parents: 52555
diff changeset
   304
  andalso not (String.isSubstring ", > " s)
2a62d848a56a improved ad hoc success detection in Mirabelle -- if the metis call fails and the structured proof succeeds, remember only the success
blanchet
parents: 52555
diff changeset
   305
  andalso not (String.isSubstring "may fail" s)
2a62d848a56a improved ad hoc success detection in Mirabelle -- if the metis call fails and the structured proof succeeds, remember only the success
blanchet
parents: 52555
diff changeset
   306
51206
3fba6741ead2 improved hack
blanchet
parents: 51203
diff changeset
   307
(* Fragile hack *)
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   308
fun proof_method_from_msg args msg =
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   309
  (case AList.lookup (op =) args proof_methodK of
58473
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   310
    SOME name =>
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   311
    if name = "smart" then
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   312
      if exists is_good_line (split_lines msg) then
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   313
        "none"
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   314
      else
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   315
        "fail"
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   316
    else
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   317
      name
41357
ae76960d86a2 added "sledgehammer_tac" as possible reconstructor in Mirabelle
blanchet
parents: 41338
diff changeset
   318
  | NONE =>
58473
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   319
    if exists is_good_line (split_lines msg) then
51203
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51187
diff changeset
   320
      "none" (* trust the preplayed proof *)
4c6ae305462e trust preplayed proof in Mirabelle
blanchet
parents: 51187
diff changeset
   321
    else if String.isSubstring "metis (" msg then
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   322
      msg |> Substring.full
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   323
          |> Substring.position "metis ("
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   324
          |> snd |> Substring.position ")"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   325
          |> fst |> Substring.string
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   326
          |> suffix ")"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   327
    else if String.isSubstring "metis" msg then
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   328
      "metis"
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   329
    else
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   330
      "smt")
40667
b8579f24ce67 make Mirabelle take into consideration whether the SMT solver proof should be reconstructed with Metis or SMT
blanchet
parents: 40627
diff changeset
   331
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   332
local
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   333
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   334
datatype sh_result =
46340
cac402c486b0 separate orthogonal components
blanchet
parents: 46320
diff changeset
   335
  SH_OK of int * int * (string * stature) list |
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   336
  SH_FAIL of int * int |
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   337
  SH_ERROR
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   338
50352
db8cae658807 added MaSh learning to Mirabelle
blanchet
parents: 50334
diff changeset
   339
fun run_sh prover_name fact_filter type_enc strict max_facts slice
50334
74d453d1b084 added "fact_filter" option to Mirabelle
blanchet
parents: 50267
diff changeset
   340
      lam_trans uncurried_aliases e_selection_heuristic term_order force_sos
58473
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   341
      hard_timeout timeout preplay_timeout isar_proofsLST smt_proofsLST
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   342
      minimizeLST max_new_mono_instancesLST max_mono_itersLST dir pos st =
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   343
  let
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   344
    val thy = Proof.theory_of st
38998
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   345
    val {context = ctxt, facts = chained_ths, goal} = Proof.goal st
f11a861e0061 share the relevance filter among the provers
blanchet
parents: 38988
diff changeset
   346
    val i = 1
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   347
    fun set_file_name (SOME dir) =
55212
blanchet
parents: 55211
diff changeset
   348
        Config.put Sledgehammer_Prover_ATP.atp_dest_dir dir
blanchet
parents: 55211
diff changeset
   349
        #> Config.put Sledgehammer_Prover_ATP.atp_problem_prefix
74472
9d304ef5c932 added offset to Mirabelle's tptp output names
desharna
parents: 74078
diff changeset
   350
          ("prob_" ^
9d304ef5c932 added offset to Mirabelle's tptp output names
desharna
parents: 74078
diff changeset
   351
            StringCvt.padLeft #"0" 5 (str0 (Position.line_of pos)) ^ "_" ^
9d304ef5c932 added offset to Mirabelle's tptp output names
desharna
parents: 74078
diff changeset
   352
            StringCvt.padLeft #"0" 6 (str0 (Position.offset_of pos)) ^ "__")
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 58028
diff changeset
   353
        #> Config.put SMT_Config.debug_files
56811
b66639331db5 optional case enforcement
haftmann
parents: 56411
diff changeset
   354
          (dir ^ "/" ^ Name.desymbolize (SOME false) (ATP_Util.timestamp ()) ^ "_"
41338
ffd730fcf0ac avoid weird symbols in path
blanchet
parents: 41337
diff changeset
   355
          ^ serial_string ())
44090
6ce82b9e2896 add line number prefix to output file name
blanchet
parents: 44089
diff changeset
   356
      | set_file_name NONE = I
39321
23951979a362 make Mirabelle happy
blanchet
parents: 39005
diff changeset
   357
    val st' =
47030
7e80e14247fc internal renamings
blanchet
parents: 46826
diff changeset
   358
      st
7e80e14247fc internal renamings
blanchet
parents: 46826
diff changeset
   359
      |> Proof.map_context
7e80e14247fc internal renamings
blanchet
parents: 46826
diff changeset
   360
           (set_file_name dir
72400
abfeed05c323 tune filename
desharna
parents: 72342
diff changeset
   361
            #> (Option.map (Config.put Sledgehammer_ATP_Systems.e_selection_heuristic)
47032
73cdeed236c0 more weight attribute tuning
blanchet
parents: 47030
diff changeset
   362
                  e_selection_heuristic |> the_default I)
72400
abfeed05c323 tune filename
desharna
parents: 72342
diff changeset
   363
            #> (Option.map (Config.put Sledgehammer_ATP_Systems.term_order)
47049
72bd3311ecba added term_order option to Mirabelle
blanchet
parents: 47032
diff changeset
   364
                  term_order |> the_default I)
72400
abfeed05c323 tune filename
desharna
parents: 72342
diff changeset
   365
            #> (Option.map (Config.put Sledgehammer_ATP_Systems.force_sos)
47030
7e80e14247fc internal renamings
blanchet
parents: 46826
diff changeset
   366
                  force_sos |> the_default I))
57782
b5dc0562c7fb renamed 'sh_minimize' to 'minimize'; compile;
blanchet
parents: 57753
diff changeset
   367
    val params as {max_facts, minimize, preplay_timeout, ...} =
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   368
      Sledgehammer_Commands.default_params thy
58473
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   369
         ([(* ("verbose", "true"), *)
50334
74d453d1b084 added "fact_filter" option to Mirabelle
blanchet
parents: 50267
diff changeset
   370
           ("fact_filter", fact_filter),
43626
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43569
diff changeset
   371
           ("type_enc", type_enc),
46386
6b17c65d35c4 renamed Sledgehammer option
blanchet
parents: 46365
diff changeset
   372
           ("strict", strict),
47481
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
   373
           ("lam_trans", lam_trans |> the_default lam_trans_default),
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
   374
           ("uncurried_aliases", uncurried_aliases |> the_default uncurried_aliases_default),
48293
914ca0827804 renamed Sledgehammer options
blanchet
parents: 48292
diff changeset
   375
           ("max_facts", max_facts),
45706
418846ea4f99 renamed "slicing" to "slice"
blanchet
parents: 45520
diff changeset
   376
           ("slice", slice),
44448
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   377
           ("timeout", string_of_int timeout),
04bd6a9307c6 don't perform a triviality check if the goal is skipped anyway
blanchet
parents: 44447
diff changeset
   378
           ("preplay_timeout", preplay_timeout)]
56411
913dc982ef55 added option to Mirabelle
blanchet
parents: 55287
diff changeset
   379
          |> isar_proofsLST
58473
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   380
          |> smt_proofsLST
57782
b5dc0562c7fb renamed 'sh_minimize' to 'minimize'; compile;
blanchet
parents: 57753
diff changeset
   381
          |> minimizeLST (*don't confuse the two minimization flags*)
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
   382
          |> max_new_mono_instancesLST
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
   383
          |> max_mono_itersLST)
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   384
    val default_max_facts =
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   385
      Sledgehammer_Prover_Minimize.default_max_facts_of_prover ctxt prover_name
52196
2281f33e8da6 redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
blanchet
parents: 52125
diff changeset
   386
    val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal i ctxt
32574
719426c9e1eb added hard timeout for sledgehammer based on elapsed time (no need to trust ATP's timeout handling);
boehmes
parents: 32571
diff changeset
   387
    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
   388
      (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
   389
        NONE => I
62519
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 59582
diff changeset
   390
      | SOME secs => Timeout.apply (Time.fromSeconds secs))
42953
26111aafab12 detect inappropriate problems and crashes better in Waldmeister
blanchet
parents: 42952
diff changeset
   391
    fun failed failure =
51009
e8ff34a1fa9a thread through fact triple component from which used facts come, for accurate index output
blanchet
parents: 51008
diff changeset
   392
      ({outcome = SOME failure, used_facts = [], used_from = [],
57751
f453bbc289fa modernized Mirabelle (a bit) and made it compile
blanchet
parents: 57154
diff changeset
   393
        preferred_methss = (Sledgehammer_Proof_Methods.Auto_Method, []), run_time = Time.zeroTime,
f453bbc289fa modernized Mirabelle (a bit) and made it compile
blanchet
parents: 57154
diff changeset
   394
        message = K ""}, ~1)
f453bbc289fa modernized Mirabelle (a bit) and made it compile
blanchet
parents: 57154
diff changeset
   395
    val ({outcome, used_facts, preferred_methss, run_time, message, ...}
55201
1ee776da8da7 renamed ML file
blanchet
parents: 55198
diff changeset
   396
         : Sledgehammer_Prover.prover_result,
51009
e8ff34a1fa9a thread through fact triple component from which used facts come, for accurate index output
blanchet
parents: 51008
diff changeset
   397
         time_isa) = time_limit (Mirabelle.cpu_time (fn () =>
41275
3897af72c731 move relevance filter into hard timeout
blanchet
parents: 41274
diff changeset
   398
      let
55212
blanchet
parents: 55211
diff changeset
   399
        val ho_atp = Sledgehammer_Prover_ATP.is_ho_atp ctxt prover_name
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58921
diff changeset
   400
        val keywords = Thy_Header.get_keywords' ctxt
48299
5e5c6616f0fe centrally construct expensive data structures
blanchet
parents: 48293
diff changeset
   401
        val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
41275
3897af72c731 move relevance filter into hard timeout
blanchet
parents: 41274
diff changeset
   402
        val facts =
48288
255c6e1fd505 rationalize relevance filter, slowing moving code from Iter to MaSh
blanchet
parents: 48250
diff changeset
   403
          Sledgehammer_Fact.nearly_all_facts ctxt ho_atp
58921
ffdafc99f67b prefer explicit Keyword.keywords (cf. 82a71046dce8);
wenzelm
parents: 58903
diff changeset
   404
              Sledgehammer_Fact.no_fact_override keywords css_table chained_ths
48299
5e5c6616f0fe centrally construct expensive data structures
blanchet
parents: 48293
diff changeset
   405
              hyp_ts concl_t
51010
afd0213a3dab tuned data structure
blanchet
parents: 51009
diff changeset
   406
        val factss =
51007
4f694d52bf62 thread fact triple (MeSh, MePo, MaSh) to allow different filters in different slices
blanchet
parents: 51005
diff changeset
   407
          facts
48381
1b7d798460bb renamed ML structures
blanchet
parents: 48321
diff changeset
   408
          |> Sledgehammer_MaSh.relevant_facts ctxt params prover_name
48293
914ca0827804 renamed Sledgehammer options
blanchet
parents: 48292
diff changeset
   409
                 (the_default default_max_facts max_facts)
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48289
diff changeset
   410
                 Sledgehammer_Fact.no_fact_override hyp_ts concl_t
51010
afd0213a3dab tuned data structure
blanchet
parents: 51009
diff changeset
   411
          |> tap (fn factss =>
50868
4b30d461b3a6 more informative output
blanchet
parents: 50867
diff changeset
   412
                     "Line " ^ str0 (Position.line_of pos) ^ ": " ^
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55201
diff changeset
   413
                     Sledgehammer.string_of_factss factss
58843
521cea5fa777 discontinued obsolete Output.urgent_message;
wenzelm
parents: 58473
diff changeset
   414
                     |> writeln)
69706
6d6235b828fc get rid of visibility in MaSh -- it slows it down more than it helps
blanchet
parents: 67405
diff changeset
   415
        val prover = get_prover ctxt prover_name params goal
41275
3897af72c731 move relevance filter into hard timeout
blanchet
parents: 41274
diff changeset
   416
        val problem =
54141
f57f8e7a879f generate a comment storing the goal nickname in "learn_prover"
blanchet
parents: 54130
diff changeset
   417
          {comment = "", state = st', goal = goal, subgoal = i,
62738
fe827c6fa8c5 compile
blanchet
parents: 62519
diff changeset
   418
           subgoal_count = Sledgehammer_Util.subgoal_count st, factss = factss, found_proof = I}
57751
f453bbc289fa modernized Mirabelle (a bit) and made it compile
blanchet
parents: 57154
diff changeset
   419
      in prover params problem end)) ()
62519
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 59582
diff changeset
   420
      handle Timeout.TIMEOUT _ => failed ATP_Proof.TimedOut
42953
26111aafab12 detect inappropriate problems and crashes better in Waldmeister
blanchet
parents: 42952
diff changeset
   421
           | Fail "inappropriate" => failed ATP_Proof.Inappropriate
45371
blanchet
parents: 44768
diff changeset
   422
    val time_prover = run_time |> Time.toMilliseconds
57782
b5dc0562c7fb renamed 'sh_minimize' to 'minimize'; compile;
blanchet
parents: 57753
diff changeset
   423
    val msg = message (fn () => Sledgehammer.play_one_line_proof minimize preplay_timeout used_facts
b5dc0562c7fb renamed 'sh_minimize' to 'minimize'; compile;
blanchet
parents: 57753
diff changeset
   424
      st' i preferred_methss)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   425
  in
55211
5d027af93a08 moved ML code around
blanchet
parents: 55208
diff changeset
   426
    (case outcome of
43052
8d6a4978cc65 automatically minimize with Metis when this can be done within a few seconds
blanchet
parents: 43051
diff changeset
   427
      NONE => (msg, SH_OK (time_isa, time_prover, used_facts))
55211
5d027af93a08 moved ML code around
blanchet
parents: 55208
diff changeset
   428
    | SOME _ => (msg, SH_FAIL (time_isa, time_prover)))
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   429
  end
37994
b04307085a09 make TPTP generator accept full first-order formulas
blanchet
parents: 37631
diff changeset
   430
  handle ERROR msg => ("error: " ^ msg, SH_ERROR)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   431
32498
1132c7c13f36 Mirabelle: actions are responsible for handling exceptions,
boehmes
parents: 32496
diff changeset
   432
in
1132c7c13f36 Mirabelle: actions are responsible for handling exceptions,
boehmes
parents: 32496
diff changeset
   433
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   434
fun run_sledgehammer change_data thy_index trivial output_dir args proof_method named_thms pos st =
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   435
  let
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   436
    val thy = Proof.theory_of st
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   437
    val thy_name = Context.theory_name thy
39340
3998dc0bf82b indicate which goals are trivial
blanchet
parents: 39337
diff changeset
   438
    val triv_str = if trivial then "[T] " else ""
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   439
    val _ = change_data inc_sh_calls
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   440
    val _ = if trivial then () else change_data inc_sh_nontriv_calls
55205
8450622db0c5 refactor large ML file
blanchet
parents: 55202
diff changeset
   441
    val prover_name = get_prover_name thy args
50334
74d453d1b084 added "fact_filter" option to Mirabelle
blanchet
parents: 50267
diff changeset
   442
    val fact_filter = AList.lookup (op =) args fact_filterK |> the_default fact_filter_default
47481
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
   443
    val type_enc = AList.lookup (op =) args type_encK |> the_default type_enc_default
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
   444
    val strict = AList.lookup (op =) args strictK |> the_default strict_default
48293
914ca0827804 renamed Sledgehammer options
blanchet
parents: 48292
diff changeset
   445
    val max_facts =
55211
5d027af93a08 moved ML code around
blanchet
parents: 55208
diff changeset
   446
      (case AList.lookup (op =) args max_factsK of
48293
914ca0827804 renamed Sledgehammer options
blanchet
parents: 48292
diff changeset
   447
        SOME max => max
55211
5d027af93a08 moved ML code around
blanchet
parents: 55208
diff changeset
   448
      | NONE =>
5d027af93a08 moved ML code around
blanchet
parents: 55208
diff changeset
   449
        (case AList.lookup (op =) args max_relevantK of
5d027af93a08 moved ML code around
blanchet
parents: 55208
diff changeset
   450
          SOME max => max
5d027af93a08 moved ML code around
blanchet
parents: 55208
diff changeset
   451
        | NONE => max_facts_default))
47481
b5873d4ff1e2 gathered mirabelle_sledgehammer's hardcoded defaults
sultana
parents: 47480
diff changeset
   452
    val slice = AList.lookup (op =) args sliceK |> the_default slice_default
45514
973bb7846505 parse lambda translation option in Metis
blanchet
parents: 45512
diff changeset
   453
    val lam_trans = AList.lookup (op =) args lam_transK
46415
26153cbe97bf added option to Mirabelle/Sledgehammer
blanchet
parents: 46386
diff changeset
   454
    val uncurried_aliases = AList.lookup (op =) args uncurried_aliasesK
47032
73cdeed236c0 more weight attribute tuning
blanchet
parents: 47030
diff changeset
   455
    val e_selection_heuristic = AList.lookup (op =) args e_selection_heuristicK
47049
72bd3311ecba added term_order option to Mirabelle
blanchet
parents: 47032
diff changeset
   456
    val term_order = AList.lookup (op =) args term_orderK
44099
5e14f591515e support local HOATPs
blanchet
parents: 44090
diff changeset
   457
    val force_sos = AList.lookup (op =) args force_sosK
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67399
diff changeset
   458
      |> Option.map (curry (op <>) "false")
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   459
    val keep_dir =
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   460
      if Mirabelle.get_bool_argument args (keepK, keep_default) then
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   461
        let val subdir = StringCvt.padLeft #"0" 4 (string_of_int thy_index) ^ "_" ^ thy_name in
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   462
          Path.append output_dir (Path.basic subdir)
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   463
          |> Isabelle_System.make_directory
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   464
          |> Path.implode
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   465
          |> SOME
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   466
        end
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   467
      else
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   468
        NONE
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   469
    val timeout = Mirabelle.get_int_argument args (prover_timeoutK, 30)
41268
56b7e277fd7d higher hard timeout
blanchet
parents: 41266
diff changeset
   470
    (* always use a hard timeout, but give some slack so that the automatic
56b7e277fd7d higher hard timeout
blanchet
parents: 41266
diff changeset
   471
       minimizer has a chance to do its magic *)
46825
98300d5f9cc0 added sh_minimize and preplay_timeout options to Mirabelle's Sledgehammer action;
sultana
parents: 46641
diff changeset
   472
    val preplay_timeout = AList.lookup (op =) args preplay_timeoutK
98300d5f9cc0 added sh_minimize and preplay_timeout options to Mirabelle's Sledgehammer action;
sultana
parents: 46641
diff changeset
   473
      |> the_default preplay_timeout_default
56411
913dc982ef55 added option to Mirabelle
blanchet
parents: 55287
diff changeset
   474
    val isar_proofsLST = available_parameter args isar_proofsK "isar_proofs"
58473
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   475
    val smt_proofsLST = available_parameter args smt_proofsK "smt_proofs"
57782
b5dc0562c7fb renamed 'sh_minimize' to 'minimize'; compile;
blanchet
parents: 57753
diff changeset
   476
    val minimizeLST = available_parameter args minimizeK "minimize"
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
   477
    val max_new_mono_instancesLST =
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
   478
      available_parameter args max_new_mono_instancesK max_new_mono_instancesK
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
   479
    val max_mono_itersLST = available_parameter args max_mono_itersK max_mono_itersK
48702
e1752ccccc34 increase defensive timeout that should typically not kick in thanks to solid time limits in Sledgehammer (otherwise, Mirabelle-based evaluations might be distorted)
boehmes
parents: 48558
diff changeset
   480
    val hard_timeout = SOME (4 * timeout)
41155
85da8cbb4966 added support for "type_sys" option to Mirabelle
blanchet
parents: 41154
diff changeset
   481
    val (msg, result) =
50352
db8cae658807 added MaSh learning to Mirabelle
blanchet
parents: 50334
diff changeset
   482
      run_sh prover_name fact_filter type_enc strict max_facts slice lam_trans
db8cae658807 added MaSh learning to Mirabelle
blanchet
parents: 50334
diff changeset
   483
        uncurried_aliases e_selection_heuristic term_order force_sos
58473
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   484
        hard_timeout timeout preplay_timeout isar_proofsLST smt_proofsLST
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   485
        minimizeLST max_new_mono_instancesLST max_mono_itersLST keep_dir pos st
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   486
  in
55211
5d027af93a08 moved ML code around
blanchet
parents: 55208
diff changeset
   487
    (case result of
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   488
      SH_OK (time_isa, time_prover, names) =>
38700
fcb0fe4c2f27 make Mirabelle happy
blanchet
parents: 38590
diff changeset
   489
        let
46340
cac402c486b0 separate orthogonal components
blanchet
parents: 46320
diff changeset
   490
          fun get_thms (name, stature) =
50267
1da2e67242d6 moved thms_of_name to Sledgehammer_Util and removed copies, updated references
smolkas
parents: 49916
diff changeset
   491
            try (Sledgehammer_Util.thms_of_name (Proof.context_of st))
49916
412346127bfa fixed theorem lookup code in Isar proof reconstruction
blanchet
parents: 49914
diff changeset
   492
              name
47154
2c357e2b8436 more robustness in case a theorem cannot be retrieved (which typically happens with backtick facts)
blanchet
parents: 47060
diff changeset
   493
            |> Option.map (pair (name, stature))
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   494
        in
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   495
          change_data inc_sh_success;
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   496
          if trivial then () else change_data inc_sh_nontriv_success;
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   497
          change_data (inc_sh_lemmas (length names));
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   498
          change_data (inc_sh_max_lems (length names));
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   499
          change_data (inc_sh_time_isa time_isa);
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   500
          change_data (inc_sh_time_prover time_prover);
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   501
          proof_method := proof_method_from_msg args msg;
38826
f42f425edf24 drop chained facts
blanchet
parents: 38752
diff changeset
   502
          named_thms := SOME (map_filter get_thms names);
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   503
          triv_str ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   504
            string_of_int time_prover ^ ") [" ^ prover_name ^ "]:\n" ^ msg
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   505
        end
40062
cfaebaa8588f make Sledgehammer minimizer fully work with SMT
blanchet
parents: 40061
diff changeset
   506
    | SH_FAIL (time_isa, time_prover) =>
32536
ac56c62758d3 tuned stats
nipkow
parents: 32533
diff changeset
   507
        let
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   508
          val _ = change_data (inc_sh_time_isa time_isa)
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   509
          val _ = change_data (inc_sh_time_prover_fail time_prover)
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   510
        in triv_str ^ "failed: " ^ msg end
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   511
    | SH_ERROR => "failed: " ^ msg)
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   512
  end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   513
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   514
end
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   515
44542
3f5fd3635281 beef up sledgehammer_tac in Mirabelle some more
blanchet
parents: 44487
diff changeset
   516
fun override_params prover type_enc timeout =
3f5fd3635281 beef up sledgehammer_tac in Mirabelle some more
blanchet
parents: 44487
diff changeset
   517
  [("provers", prover),
48293
914ca0827804 renamed Sledgehammer options
blanchet
parents: 48292
diff changeset
   518
   ("max_facts", "0"),
44542
3f5fd3635281 beef up sledgehammer_tac in Mirabelle some more
blanchet
parents: 44487
diff changeset
   519
   ("type_enc", type_enc),
46386
6b17c65d35c4 renamed Sledgehammer option
blanchet
parents: 46365
diff changeset
   520
   ("strict", "true"),
45706
418846ea4f99 renamed "slicing" to "slice"
blanchet
parents: 45520
diff changeset
   521
   ("slice", "false"),
44461
5e19eecb0e1c specify timeout for "sledgehammer_tac"
blanchet
parents: 44449
diff changeset
   522
   ("timeout", timeout |> Time.toSeconds |> string_of_int)]
44430
c6f0490d432f beef up "sledgehammer_tac" reconstructor
blanchet
parents: 44424
diff changeset
   523
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   524
fun run_proof_method change_data trivial full name meth named_thms timeout pos st =
32525
ea322e847633 added signature ATP_MINIMAL,
boehmes
parents: 32523
diff changeset
   525
  let
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   526
    fun do_method named_thms ctxt =
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   527
      let
59083
88b0b1f28adc tuned signature;
wenzelm
parents: 58928
diff changeset
   528
        val ref_of_str = (* FIXME proper wrapper for parser combinators *)
88b0b1f28adc tuned signature;
wenzelm
parents: 58928
diff changeset
   529
          suffix ";" #> Token.explode (Thy_Header.get_keywords' ctxt) Position.none
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62826
diff changeset
   530
          #> Parse.thm #> fst
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   531
        val thms = named_thms |> maps snd
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   532
        val facts = named_thms |> map (ref_of_str o fst o fst)
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48289
diff changeset
   533
        val fact_override = {add = facts, del = [], only = true}
44566
bf8331161ad9 split timeout among ATPs in and add Metis to the mix as backup
blanchet
parents: 44542
diff changeset
   534
        fun my_timeout time_slice =
62826
eb94e570c1a4 prefer infix operations;
wenzelm
parents: 62738
diff changeset
   535
          timeout |> Time.toReal |> curry (op *) time_slice |> Time.fromReal
44566
bf8331161ad9 split timeout among ATPs in and add Metis to the mix as backup
blanchet
parents: 44542
diff changeset
   536
        fun sledge_tac time_slice prover type_enc =
44542
3f5fd3635281 beef up sledgehammer_tac in Mirabelle some more
blanchet
parents: 44487
diff changeset
   537
          Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt
57814
7a9aaddb3203 compile
blanchet
parents: 57782
diff changeset
   538
            (override_params prover type_enc (my_timeout time_slice)) fact_override []
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   539
      in
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   540
        if !meth = "sledgehammer_tac" then
57154
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 56811
diff changeset
   541
          sledge_tac 0.2 ATP_Proof.vampireN "mono_native"
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 56811
diff changeset
   542
          ORELSE' sledge_tac 0.2 ATP_Proof.eN "poly_guards??"
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 56811
diff changeset
   543
          ORELSE' sledge_tac 0.2 ATP_Proof.spassN "mono_native"
f0eff6393a32 basic setup for zipperposition prover
fleury
parents: 56811
diff changeset
   544
          ORELSE' sledge_tac 0.2 ATP_Proof.z3_tptpN "poly_tags??"
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 58028
diff changeset
   545
          ORELSE' SMT_Solver.smt_tac ctxt thms
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   546
        else if !meth = "smt" then
58061
3d060f43accb renamed new SMT module from 'SMT2' to 'SMT'
blanchet
parents: 58028
diff changeset
   547
          SMT_Solver.smt_tac ctxt thms
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   548
        else if full then
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45706
diff changeset
   549
          Metis_Tactic.metis_tac [ATP_Proof_Reconstruct.full_typesN]
54500
f625e0e79dd1 refactoring
blanchet
parents: 54141
diff changeset
   550
            ATP_Proof_Reconstruct.default_metis_lam_trans ctxt thms
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   551
        else if String.isPrefix "metis (" (!meth) then
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   552
          let
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   553
            val (type_encs, lam_trans) =
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   554
              !meth
59083
88b0b1f28adc tuned signature;
wenzelm
parents: 58928
diff changeset
   555
              |> Token.explode (Thy_Header.get_keywords' ctxt) Position.start
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   556
              |> filter Token.is_proper |> tl
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   557
              |> Metis_Tactic.parse_metis_options |> fst
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 45706
diff changeset
   558
              |>> the_default [ATP_Proof_Reconstruct.partial_typesN]
54500
f625e0e79dd1 refactoring
blanchet
parents: 54141
diff changeset
   559
              ||> the_default ATP_Proof_Reconstruct.default_metis_lam_trans
45519
cd6e78cb6ee8 make metis reconstruction handling more flexible
blanchet
parents: 45514
diff changeset
   560
          in Metis_Tactic.metis_tac type_encs lam_trans ctxt thms end
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   561
        else if !meth = "metis" then
55211
5d027af93a08 moved ML code around
blanchet
parents: 55208
diff changeset
   562
          Metis_Tactic.metis_tac [] ATP_Proof_Reconstruct.default_metis_lam_trans ctxt thms
58473
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   563
        else if !meth = "none" then
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   564
          K all_tac
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   565
        else if !meth = "fail" then
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   566
          K no_tac
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   567
        else
58473
d919cde25691 added options to Mirabelle
blanchet
parents: 58061
diff changeset
   568
          (warning ("Unknown method " ^ quote (!meth)); K no_tac)
44462
d9a657c44380 more reliable "sledgehammer\_tac" reconstruction, by avoiding "insert_tac"
blanchet
parents: 44461
diff changeset
   569
      end
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   570
    fun apply_method named_thms =
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   571
      Mirabelle.can_apply timeout (do_method named_thms) st
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   572
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   573
    fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   574
      | with_time (true, t) =
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   575
          (change_data inc_proof_method_success;
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   576
           if trivial then () else change_data inc_proof_method_nontriv_success;
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   577
           change_data (inc_proof_method_lemmas (length named_thms));
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   578
           change_data (inc_proof_method_time t);
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   579
           change_data (inc_proof_method_posns (pos, trivial));
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   580
           if name = "proof" then change_data inc_proof_method_proofs else ();
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   581
           "succeeded (" ^ string_of_int t ^ ")")
55285
e88ad20035f4 merged 'reconstructors' and 'proof methods'
blanchet
parents: 55212
diff changeset
   582
    fun timed_method named_thms =
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   583
      with_time (Mirabelle.cpu_time apply_method named_thms)
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   584
        handle Timeout.TIMEOUT _ => (change_data inc_proof_method_timeout; "timeout")
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   585
          | ERROR msg => ("error: " ^ msg)
32521
f20cc66b2c74 added initialization and cleanup of actions,
boehmes
parents: 32518
diff changeset
   586
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   587
    val _ = change_data inc_proof_method_calls
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   588
    val _ = if trivial then () else change_data inc_proof_method_nontriv_calls
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   589
  in timed_method named_thms end
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   590
41276
285aea0c153c two layers of timeouts seem to be less reliable than a single layer
blanchet
parents: 41275
diff changeset
   591
val try_timeout = seconds 5.0
39337
ffa577c0bbfa keep track of trivial vs. nontrivial calls using "try" for 30 seconds
blanchet
parents: 39321
diff changeset
   592
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   593
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
   594
  let
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   595
    val check_trivial =
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   596
      Mirabelle.get_bool_argument arguments (check_trivialK, check_trivial_default)
44431
3e0ce0ae1022 added "max_calls" option to get a fixed number of Sledgehammer calls per theory
blanchet
parents: 44430
diff changeset
   597
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   598
    val data = Synchronized.var "Mirabelle_Sledgehammer.data" empty_data
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   599
    val change_data = Synchronized.change data
32385
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   600
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   601
    fun run_action ({theory_index, name, pos, pre, ...} : Mirabelle.command) =
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   602
      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
   603
        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
   604
          ""
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   605
        else
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   606
          let
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   607
            val meth = Unsynchronized.ref ""
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   608
            val named_thms =
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   609
              Unsynchronized.ref (NONE : ((string * stature) * thm list) list option)
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   610
            val trivial =
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   611
              check_trivial andalso Try0.try0 (SOME try_timeout) ([], [], [], []) pre
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   612
              handle Timeout.TIMEOUT _ => false
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   613
            val log1 =
74078
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   614
              run_sledgehammer change_data theory_index trivial output_dir arguments meth named_thms
a2cbe81e1e32 changed Mirabelle_Sledgehammer keep option from path to boolean
desharna
parents: 73849
diff changeset
   615
                pos pre
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   616
            val log2 =
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   617
              (case !named_thms of
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   618
                SOME thms =>
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   619
                !meth ^ " (sledgehammer): " ^ run_proof_method change_data trivial false name meth
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   620
                  thms timeout pos pre
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   621
              | NONE => "")
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   622
          in log1 ^ "\n" ^ log2 end
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   623
      end
73697
0e7a5c7a14c8 reactive "sledgehammer";
wenzelm
parents: 73691
diff changeset
   624
73847
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   625
    fun finalize () = log_data (Synchronized.value data)
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   626
  in {run_action = run_action, finalize = finalize} end
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   627
58f6b41efe88 refactored Mirabelle to produce output in real time
desharna
parents: 73797
diff changeset
   628
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
   629
594890623c46 split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
boehmes
parents:
diff changeset
   630
end