| author | wenzelm | 
| Thu, 05 Dec 2013 17:51:29 +0100 | |
| changeset 54669 | 1b153cb9699f | 
| parent 54307 | 903ab115e9fd | 
| child 54059 | 896b55752938 | 
| permissions | -rw-r--r-- | 
| 
36375
 
2482446a604c
move the minimizer to the Sledgehammer directory
 
blanchet 
parents: 
36373 
diff
changeset
 | 
1  | 
(* Title: HOL/Tools/Sledgehammer/sledgehammer_isar.ML  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
2  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
4  | 
Adds "sledgehammer" and related commands to Isabelle/Isar's outer syntax.  | 
| 
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
6  | 
|
| 35963 | 7  | 
signature SLEDGEHAMMER_ISAR =  | 
8  | 
sig  | 
|
| 
41087
 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 
blanchet 
parents: 
40941 
diff
changeset
 | 
9  | 
type params = Sledgehammer_Provers.params  | 
| 35963 | 10  | 
|
| 
40059
 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 
blanchet 
parents: 
39335 
diff
changeset
 | 
11  | 
val provers : string Unsynchronized.ref  | 
| 40069 | 12  | 
val default_params : Proof.context -> (string * string) list -> params  | 
| 35963 | 13  | 
end;  | 
14  | 
||
| 35971 | 15  | 
structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR =  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
16  | 
struct  | 
| 
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
17  | 
|
| 44784 | 18  | 
open ATP_Util  | 
| 38028 | 19  | 
open ATP_Systems  | 
| 46320 | 20  | 
open ATP_Problem_Generate  | 
21  | 
open ATP_Proof_Reconstruct  | 
|
| 35971 | 22  | 
open Sledgehammer_Util  | 
| 
48250
 
1065c307fafe
further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
 
blanchet 
parents: 
47962 
diff
changeset
 | 
23  | 
open Sledgehammer_Fact  | 
| 
41087
 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 
blanchet 
parents: 
40941 
diff
changeset
 | 
24  | 
open Sledgehammer_Provers  | 
| 38988 | 25  | 
open Sledgehammer_Minimize  | 
| 48381 | 26  | 
open Sledgehammer_MaSh  | 
| 
41087
 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 
blanchet 
parents: 
40941 
diff
changeset
 | 
27  | 
open Sledgehammer_Run  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
28  | 
|
| 48433 | 29  | 
(* val cvc3N = "cvc3" *)  | 
| 
48405
 
7682bc885e8a
use CVC3 and Yices by default if they are available and there are enough cores
 
blanchet 
parents: 
48400 
diff
changeset
 | 
30  | 
val yicesN = "yices"  | 
| 
 
7682bc885e8a
use CVC3 and Yices by default if they are available and there are enough cores
 
blanchet 
parents: 
48400 
diff
changeset
 | 
31  | 
val z3N = "z3"  | 
| 
 
7682bc885e8a
use CVC3 and Yices by default if they are available and there are enough cores
 
blanchet 
parents: 
48400 
diff
changeset
 | 
32  | 
|
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
33  | 
val runN = "run"  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
34  | 
val minN = "min"  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
35  | 
val messagesN = "messages"  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
36  | 
val supported_proversN = "supported_provers"  | 
| 
50719
 
58b0b44da54a
renamed "kill" subcommand to avoid clash with "kill" keyword (which confuses Proof General and results in strange syntax highlighting)
 
blanchet 
parents: 
50604 
diff
changeset
 | 
37  | 
val kill_allN = "kill_all"  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
38  | 
val running_proversN = "running_provers"  | 
| 
48319
 
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
 
blanchet 
parents: 
48314 
diff
changeset
 | 
39  | 
val running_learnersN = "running_learners"  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
40  | 
val refresh_tptpN = "refresh_tptp"  | 
| 48301 | 41  | 
|
| 39318 | 42  | 
val _ =  | 
| 52639 | 43  | 
ProofGeneral.preference_option ProofGeneral.category_tracing  | 
| 
52017
 
bc0238c1f73a
clarified preferences: "override" re-initialized on prover startup, and "default" sent to PG -- thus recover typical defaults like auto-quickcheck in PG 4.x;
 
wenzelm 
parents: 
52007 
diff
changeset
 | 
44  | 
NONE  | 
| 52639 | 45  | 
    @{option auto_sledgehammer}
 | 
| 
52017
 
bc0238c1f73a
clarified preferences: "override" re-initialized on prover startup, and "default" sent to PG -- thus recover typical defaults like auto-quickcheck in PG 4.x;
 
wenzelm 
parents: 
52007 
diff
changeset
 | 
46  | 
"auto-sledgehammer"  | 
| 
 
bc0238c1f73a
clarified preferences: "override" re-initialized on prover startup, and "default" sent to PG -- thus recover typical defaults like auto-quickcheck in PG 4.x;
 
wenzelm 
parents: 
52007 
diff
changeset
 | 
47  | 
"Run Sledgehammer automatically"  | 
| 39318 | 48  | 
|
| 
36394
 
1a48d18449d8
move "neg_clausify" method and "clausify" attribute to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36378 
diff
changeset
 | 
49  | 
(** Sledgehammer commands **)  | 
| 
 
1a48d18449d8
move "neg_clausify" method and "clausify" attribute to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36378 
diff
changeset
 | 
50  | 
|
| 48292 | 51  | 
fun add_fact_override ns : fact_override = {add = ns, del = [], only = false}
 | 
52  | 
fun del_fact_override ns : fact_override = {add = [], del = ns, only = false}
 | 
|
53  | 
fun only_fact_override ns : fact_override = {add = ns, del = [], only = true}
 | 
|
54  | 
fun merge_fact_override_pairwise (r1 : fact_override) (r2 : fact_override) =  | 
|
| 
35966
 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 
blanchet 
parents: 
35965 
diff
changeset
 | 
55  | 
  {add = #add r1 @ #add r2, del = #del r1 @ #del r2,
 | 
| 
36183
 
f723348b2231
Sledgehammer: the empty set of fact () should mean nothing, not unchanged
 
blanchet 
parents: 
36143 
diff
changeset
 | 
56  | 
only = #only r1 andalso #only r2}  | 
| 48292 | 57  | 
fun merge_fact_overrides rs =  | 
58  | 
fold merge_fact_override_pairwise rs (only_fact_override [])  | 
|
| 
35965
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
59  | 
|
| 
36371
 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36290 
diff
changeset
 | 
60  | 
(*** parameters ***)  | 
| 
 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36290 
diff
changeset
 | 
61  | 
|
| 
40059
 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 
blanchet 
parents: 
39335 
diff
changeset
 | 
62  | 
val provers = Unsynchronized.ref ""  | 
| 
36371
 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36290 
diff
changeset
 | 
63  | 
|
| 
 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36290 
diff
changeset
 | 
64  | 
val _ =  | 
| 
52007
 
0b1183012a3c
maintain ProofGeneral preferences within ProofGeneral module;
 
wenzelm 
parents: 
52006 
diff
changeset
 | 
65  | 
ProofGeneral.preference_string ProofGeneral.category_proof  | 
| 
52017
 
bc0238c1f73a
clarified preferences: "override" re-initialized on prover startup, and "default" sent to PG -- thus recover typical defaults like auto-quickcheck in PG 4.x;
 
wenzelm 
parents: 
52007 
diff
changeset
 | 
66  | 
NONE  | 
| 
52007
 
0b1183012a3c
maintain ProofGeneral preferences within ProofGeneral module;
 
wenzelm 
parents: 
52006 
diff
changeset
 | 
67  | 
provers  | 
| 
 
0b1183012a3c
maintain ProofGeneral preferences within ProofGeneral module;
 
wenzelm 
parents: 
52006 
diff
changeset
 | 
68  | 
"Sledgehammer: Provers"  | 
| 
 
0b1183012a3c
maintain ProofGeneral preferences within ProofGeneral module;
 
wenzelm 
parents: 
52006 
diff
changeset
 | 
69  | 
"Default automatic provers (separated by whitespace)"  | 
| 
36371
 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36290 
diff
changeset
 | 
70  | 
|
| 
 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36290 
diff
changeset
 | 
71  | 
val _ =  | 
| 
53057
 
e18a028b345c
prefer system option sledgehammer_timeout, with standard GUI in jEdit Plugin Options;
 
wenzelm 
parents: 
53055 
diff
changeset
 | 
72  | 
ProofGeneral.preference_option ProofGeneral.category_proof  | 
| 
52017
 
bc0238c1f73a
clarified preferences: "override" re-initialized on prover startup, and "default" sent to PG -- thus recover typical defaults like auto-quickcheck in PG 4.x;
 
wenzelm 
parents: 
52007 
diff
changeset
 | 
73  | 
NONE  | 
| 
53057
 
e18a028b345c
prefer system option sledgehammer_timeout, with standard GUI in jEdit Plugin Options;
 
wenzelm 
parents: 
53055 
diff
changeset
 | 
74  | 
    @{option sledgehammer_timeout}
 | 
| 
52007
 
0b1183012a3c
maintain ProofGeneral preferences within ProofGeneral module;
 
wenzelm 
parents: 
52006 
diff
changeset
 | 
75  | 
"Sledgehammer: Time Limit"  | 
| 
 
0b1183012a3c
maintain ProofGeneral preferences within ProofGeneral module;
 
wenzelm 
parents: 
52006 
diff
changeset
 | 
76  | 
"ATPs will be interrupted after this time (in seconds)"  | 
| 
36371
 
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36290 
diff
changeset
 | 
77  | 
|
| 35963 | 78  | 
type raw_param = string * string list  | 
79  | 
||
80  | 
val default_default_params =  | 
|
| 
41208
 
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
 
blanchet 
parents: 
41153 
diff
changeset
 | 
81  | 
  [("debug", "false"),
 | 
| 35963 | 82  | 
   ("verbose", "false"),
 | 
| 
36143
 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 
blanchet 
parents: 
36142 
diff
changeset
 | 
83  | 
   ("overlord", "false"),
 | 
| 53800 | 84  | 
   ("spy", "false"),
 | 
| 
41208
 
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
 
blanchet 
parents: 
41153 
diff
changeset
 | 
85  | 
   ("blocking", "false"),
 | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43572 
diff
changeset
 | 
86  | 
   ("type_enc", "smart"),
 | 
| 46301 | 87  | 
   ("strict", "false"),
 | 
| 45514 | 88  | 
   ("lam_trans", "smart"),
 | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46398 
diff
changeset
 | 
89  | 
   ("uncurried_aliases", "smart"),
 | 
| 48321 | 90  | 
   ("learn", "true"),
 | 
| 
48314
 
ee33ba3c0e05
added option to control which fact filter is used
 
blanchet 
parents: 
48308 
diff
changeset
 | 
91  | 
   ("fact_filter", "smart"),
 | 
| 
 
ee33ba3c0e05
added option to control which fact filter is used
 
blanchet 
parents: 
48308 
diff
changeset
 | 
92  | 
   ("max_facts", "smart"),
 | 
| 48293 | 93  | 
   ("fact_thresholds", "0.45 0.85"),
 | 
| 
47962
 
137883567114
lower the monomorphization thresholds for less scalable provers
 
blanchet 
parents: 
47642 
diff
changeset
 | 
94  | 
   ("max_mono_iters", "smart"),
 | 
| 
 
137883567114
lower the monomorphization thresholds for less scalable provers
 
blanchet 
parents: 
47642 
diff
changeset
 | 
95  | 
   ("max_new_mono_instances", "smart"),
 | 
| 
51190
 
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
 
blanchet 
parents: 
51189 
diff
changeset
 | 
96  | 
   ("isar_proofs", "smart"),
 | 
| 
51130
 
76d68444cd59
renamed sledgehammer_shrink to sledgehammer_compress
 
smolkas 
parents: 
50824 
diff
changeset
 | 
97  | 
   ("isar_compress", "10"),
 | 
| 53764 | 98  | 
   ("isar_try0", "true"),
 | 
| 45706 | 99  | 
   ("slice", "true"),
 | 
| 
45707
 
6bf7eec9b153
added "minimize" option for more control over automatic minimization
 
blanchet 
parents: 
45706 
diff
changeset
 | 
100  | 
   ("minimize", "smart"),
 | 
| 
53761
 
4d34e267fba9
use configuration mechanism for low-level tracing
 
blanchet 
parents: 
53759 
diff
changeset
 | 
101  | 
   ("preplay_timeout", "3")]
 | 
| 35963 | 102  | 
|
| 
36140
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
103  | 
val alias_params =  | 
| 51138 | 104  | 
  [("prover", ("provers", [])), (* undocumented *)
 | 
| 51189 | 105  | 
   ("dont_preplay", ("preplay_timeout", ["0"])),
 | 
| 52093 | 106  | 
   ("dont_compress_isar", ("isar_compress", ["0"])),
 | 
107  | 
   ("isar_proof", ("isar_proofs", [])) (* legacy *)]
 | 
|
| 
36140
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
108  | 
val negated_alias_params =  | 
| 
41208
 
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
 
blanchet 
parents: 
41153 
diff
changeset
 | 
109  | 
  [("no_debug", "debug"),
 | 
| 35963 | 110  | 
   ("quiet", "verbose"),
 | 
| 
36143
 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 
blanchet 
parents: 
36142 
diff
changeset
 | 
111  | 
   ("no_overlord", "overlord"),
 | 
| 53800 | 112  | 
   ("dont_spy", "spy"),
 | 
| 
41208
 
1b28c43a7074
make "debug" imply "blocking", since in blocking mode the exceptions flow through and are more instructive
 
blanchet 
parents: 
41153 
diff
changeset
 | 
113  | 
   ("non_blocking", "blocking"),
 | 
| 46301 | 114  | 
   ("non_strict", "strict"),
 | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46398 
diff
changeset
 | 
115  | 
   ("no_uncurried_aliases", "uncurried_aliases"),
 | 
| 48321 | 116  | 
   ("dont_learn", "learn"),
 | 
| 
49918
 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 
blanchet 
parents: 
49632 
diff
changeset
 | 
117  | 
   ("no_isar_proofs", "isar_proofs"),
 | 
| 
45707
 
6bf7eec9b153
added "minimize" option for more control over automatic minimization
 
blanchet 
parents: 
45706 
diff
changeset
 | 
118  | 
   ("dont_slice", "slice"),
 | 
| 51879 | 119  | 
   ("dont_minimize", "minimize"),
 | 
| 53764 | 120  | 
   ("dont_try0_isar", "isar_try0")]
 | 
| 35963 | 121  | 
|
| 53758 | 122  | 
val params_not_for_minimize =  | 
123  | 
["blocking", "fact_filter", "max_facts", "fact_thresholds", "slice",  | 
|
124  | 
"minimize"];  | 
|
| 
36281
 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 
blanchet 
parents: 
36235 
diff
changeset
 | 
125  | 
|
| 
43569
 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 
blanchet 
parents: 
43353 
diff
changeset
 | 
126  | 
val property_dependent_params = ["provers", "timeout"]  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
127  | 
|
| 35963 | 128  | 
fun is_known_raw_param s =  | 
129  | 
AList.defined (op =) default_default_params s orelse  | 
|
| 
36140
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
130  | 
AList.defined (op =) alias_params s orelse  | 
| 
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
131  | 
AList.defined (op =) negated_alias_params s orelse  | 
| 
38985
 
162bbbea4e4d
added "expect" feature of Nitpick to Sledgehammer, for regression testing
 
blanchet 
parents: 
38982 
diff
changeset
 | 
132  | 
member (op =) property_dependent_params s orelse s = "expect"  | 
| 35963 | 133  | 
|
| 
41258
 
73401632a80c
convenient syntax for setting provers -- useful for debugging, not for general consumption and hence not documented
 
blanchet 
parents: 
41208 
diff
changeset
 | 
134  | 
fun is_prover_list ctxt s =  | 
| 
 
73401632a80c
convenient syntax for setting provers -- useful for debugging, not for general consumption and hence not documented
 
blanchet 
parents: 
41208 
diff
changeset
 | 
135  | 
case space_explode " " s of  | 
| 
41727
 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 
blanchet 
parents: 
41472 
diff
changeset
 | 
136  | 
ss as _ :: _ => forall (is_prover_supported ctxt) ss  | 
| 
41258
 
73401632a80c
convenient syntax for setting provers -- useful for debugging, not for general consumption and hence not documented
 
blanchet 
parents: 
41208 
diff
changeset
 | 
137  | 
| _ => false  | 
| 
 
73401632a80c
convenient syntax for setting provers -- useful for debugging, not for general consumption and hence not documented
 
blanchet 
parents: 
41208 
diff
changeset
 | 
138  | 
|
| 
36140
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
139  | 
fun unalias_raw_param (name, value) =  | 
| 
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
140  | 
case AList.lookup (op =) alias_params name of  | 
| 47037 | 141  | 
SOME (name', []) => (name', value)  | 
142  | 
| SOME (param' as (name', _)) =>  | 
|
143  | 
if value <> ["false"] then  | 
|
144  | 
param'  | 
|
145  | 
else  | 
|
146  | 
      error ("Parameter " ^ quote name ^ " cannot be set to \"false\" \
 | 
|
147  | 
\(cf. " ^ quote name' ^ ").")  | 
|
| 
36140
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
148  | 
| NONE =>  | 
| 
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
149  | 
case AList.lookup (op =) negated_alias_params name of  | 
| 
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
150  | 
SOME name' => (name', case value of  | 
| 
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
151  | 
["false"] => ["true"]  | 
| 
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
152  | 
| ["true"] => ["false"]  | 
| 
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
153  | 
| [] => ["false"]  | 
| 
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
154  | 
| _ => value)  | 
| 
 
08b2a7ecb6c3
fixed handling of "sledgehammer_params" that get a default value from Isabelle menu;
 
blanchet 
parents: 
36064 
diff
changeset
 | 
155  | 
| NONE => (name, value)  | 
| 35963 | 156  | 
|
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52018 
diff
changeset
 | 
157  | 
val any_type_enc = type_enc_of_string Strict "erased"  | 
| 45514 | 158  | 
|
| 48397 | 159  | 
(* "provers =", "type_enc =", "lam_trans =", "fact_filter =", and "max_facts ="  | 
160  | 
can be omitted. For the last four, this is a secret feature. *)  | 
|
| 
44720
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
161  | 
fun normalize_raw_param ctxt =  | 
| 
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
162  | 
unalias_raw_param  | 
| 
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
163  | 
#> (fn (name, value) =>  | 
| 
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
164  | 
if is_known_raw_param name then  | 
| 
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
165  | 
(name, value)  | 
| 48533 | 166  | 
else if null value then  | 
167  | 
if is_prover_list ctxt name then  | 
|
168  | 
             ("provers", [name])
 | 
|
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52018 
diff
changeset
 | 
169  | 
else if can (type_enc_of_string Strict) name then  | 
| 48533 | 170  | 
             ("type_enc", [name])
 | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52018 
diff
changeset
 | 
171  | 
else if can (trans_lams_of_string ctxt any_type_enc) name then  | 
| 48533 | 172  | 
             ("lam_trans", [name])
 | 
173  | 
else if member (op =) fact_filters name then  | 
|
174  | 
             ("fact_filter", [name])
 | 
|
175  | 
else if is_some (Int.fromString name) then  | 
|
176  | 
             ("max_facts", [name])
 | 
|
177  | 
else  | 
|
178  | 
             error ("Unknown parameter: " ^ quote name ^ ".")
 | 
|
| 48400 | 179  | 
else  | 
180  | 
           error ("Unknown parameter: " ^ quote name ^ "."))
 | 
|
| 
44720
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
181  | 
|
| 46435 | 182  | 
(* Ensures that type encodings such as "mono_native?" and "poly_guards!!" are  | 
| 44785 | 183  | 
read correctly. *)  | 
| 44784 | 184  | 
val implode_param = strip_spaces_except_between_idents o space_implode " "  | 
| 
43009
 
f58bab6be6a9
reintroduced Waldmeister but limit the number of remote threads created
 
blanchet 
parents: 
43008 
diff
changeset
 | 
185  | 
|
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41258 
diff
changeset
 | 
186  | 
structure Data = Theory_Data  | 
| 
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41258 
diff
changeset
 | 
187  | 
(  | 
| 35963 | 188  | 
type T = raw_param list  | 
| 53800 | 189  | 
val empty =  | 
190  | 
default_default_params  | 
|
191  | 
    |> getenv "SLEDGEHAMMER_SPY" = "yes" ? AList.update (op =) ("spy", "true")
 | 
|
192  | 
|> map (apsnd single)  | 
|
| 35963 | 193  | 
val extend = I  | 
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41258 
diff
changeset
 | 
194  | 
fun merge data : T = AList.merge (op =) (K true) data  | 
| 
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
41258 
diff
changeset
 | 
195  | 
)  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
196  | 
|
| 
40059
 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 
blanchet 
parents: 
39335 
diff
changeset
 | 
197  | 
(* The first ATP of the list is used by Auto Sledgehammer. Because of the low  | 
| 
48405
 
7682bc885e8a
use CVC3 and Yices by default if they are available and there are enough cores
 
blanchet 
parents: 
48400 
diff
changeset
 | 
198  | 
timeout, it makes sense to put E first. *)  | 
| 40069 | 199  | 
fun default_provers_param_value ctxt =  | 
| 
53759
 
a198ce71de11
took out Waldmeister from list of default provers -- it's usually just visual noise, and its integration in Sledgehammer leaves much to be desired
 
blanchet 
parents: 
53758 
diff
changeset
 | 
200  | 
[eN, spassN, vampireN, z3N, e_sineN, yicesN]  | 
| 
40941
 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 
blanchet 
parents: 
40931 
diff
changeset
 | 
201  | 
|> map_filter (remotify_prover_if_not_installed ctxt)  | 
| 
54307
 
903ab115e9fd
count remote threads as well when balancing CPU usage -- otherwise jEdit users and other users of the "blocking" mode may have to wait for 2 * timeout if they e.g. have 4 cores and 5 provers (the typical situation)
 
blanchet 
parents: 
53800 
diff
changeset
 | 
202  | 
|> take (Multithreading.max_threads_value ())  | 
| 42776 | 203  | 
|> implode_param  | 
| 
40059
 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 
blanchet 
parents: 
39335 
diff
changeset
 | 
204  | 
|
| 
44720
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
205  | 
fun set_default_raw_param param thy =  | 
| 
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
206  | 
let val ctxt = Proof_Context.init_global thy in  | 
| 
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
207  | 
thy |> Data.map (AList.update (op =) (normalize_raw_param ctxt param))  | 
| 
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
208  | 
end  | 
| 40069 | 209  | 
fun default_raw_params ctxt =  | 
| 42361 | 210  | 
let val thy = Proof_Context.theory_of ctxt in  | 
| 40069 | 211  | 
Data.get thy  | 
212  | 
|> fold (AList.default (op =))  | 
|
213  | 
            [("provers", [case !provers of
 | 
|
214  | 
"" => default_provers_param_value ctxt  | 
|
215  | 
| s => s]),  | 
|
| 
53057
 
e18a028b345c
prefer system option sledgehammer_timeout, with standard GUI in jEdit Plugin Options;
 
wenzelm 
parents: 
53055 
diff
changeset
 | 
216  | 
             ("timeout", let val timeout = Options.default_int @{option sledgehammer_timeout} in
 | 
| 40069 | 217  | 
[if timeout <= 0 then "none"  | 
| 
40341
 
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
 
blanchet 
parents: 
40181 
diff
changeset
 | 
218  | 
else string_of_int timeout]  | 
| 40069 | 219  | 
end)]  | 
220  | 
end  | 
|
| 35963 | 221  | 
|
| 43021 | 222  | 
fun extract_params mode default_params override_params =  | 
| 35963 | 223  | 
let  | 
224  | 
val raw_params = rev override_params @ rev default_params  | 
|
| 42776 | 225  | 
val lookup = Option.map implode_param o AList.lookup (op =) raw_params  | 
| 35963 | 226  | 
val lookup_string = the_default "" o lookup  | 
227  | 
fun general_lookup_bool option default_value name =  | 
|
228  | 
case lookup name of  | 
|
| 
35970
 
3d41a2a490f0
revert debugging output that shouldn't have been submitted in the first place
 
blanchet 
parents: 
35966 
diff
changeset
 | 
229  | 
SOME s => parse_bool_option option name s  | 
| 35963 | 230  | 
| NONE => default_value  | 
231  | 
val lookup_bool = the o general_lookup_bool false (SOME false)  | 
|
232  | 
fun lookup_time name =  | 
|
| 39318 | 233  | 
case lookup name of  | 
234  | 
SOME s => parse_time_option name s  | 
|
235  | 
| NONE => NONE  | 
|
| 
35966
 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 
blanchet 
parents: 
35965 
diff
changeset
 | 
236  | 
fun lookup_int name =  | 
| 35963 | 237  | 
case lookup name of  | 
| 
35966
 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 
blanchet 
parents: 
35965 
diff
changeset
 | 
238  | 
NONE => 0  | 
| 
 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 
blanchet 
parents: 
35965 
diff
changeset
 | 
239  | 
| SOME s => case Int.fromString s of  | 
| 
 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 
blanchet 
parents: 
35965 
diff
changeset
 | 
240  | 
SOME n => n  | 
| 
 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 
blanchet 
parents: 
35965 
diff
changeset
 | 
241  | 
                  | NONE => error ("Parameter " ^ quote name ^
 | 
| 
 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 
blanchet 
parents: 
35965 
diff
changeset
 | 
242  | 
" must be assigned an integer value.")  | 
| 
49918
 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 
blanchet 
parents: 
49632 
diff
changeset
 | 
243  | 
fun lookup_real name =  | 
| 
 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 
blanchet 
parents: 
49632 
diff
changeset
 | 
244  | 
case lookup name of  | 
| 
 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 
blanchet 
parents: 
49632 
diff
changeset
 | 
245  | 
NONE => 0.0  | 
| 
 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 
blanchet 
parents: 
49632 
diff
changeset
 | 
246  | 
| SOME s => case Real.fromString s of  | 
| 
 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 
blanchet 
parents: 
49632 
diff
changeset
 | 
247  | 
SOME n => n  | 
| 
 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 
blanchet 
parents: 
49632 
diff
changeset
 | 
248  | 
                  | NONE => error ("Parameter " ^ quote name ^
 | 
| 
 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 
blanchet 
parents: 
49632 
diff
changeset
 | 
249  | 
" must be assigned a real value.")  | 
| 
40343
 
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
 
blanchet 
parents: 
40341 
diff
changeset
 | 
250  | 
fun lookup_real_pair name =  | 
| 
38745
 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 
blanchet 
parents: 
38744 
diff
changeset
 | 
251  | 
case lookup name of  | 
| 
40343
 
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
 
blanchet 
parents: 
40341 
diff
changeset
 | 
252  | 
NONE => (0.0, 0.0)  | 
| 
 
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
 
blanchet 
parents: 
40341 
diff
changeset
 | 
253  | 
| SOME s => case s |> space_explode " " |> map Real.fromString of  | 
| 
 
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
 
blanchet 
parents: 
40341 
diff
changeset
 | 
254  | 
[SOME r1, SOME r2] => (r1, r2)  | 
| 
38745
 
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
 
blanchet 
parents: 
38744 
diff
changeset
 | 
255  | 
                  | _ => error ("Parameter " ^ quote name ^
 | 
| 49632 | 256  | 
" must be assigned a pair of floating-point \  | 
| 
40343
 
4521d56aef63
use floating-point numbers for Sledgehammer's "thresholds" option rather than percentages;
 
blanchet 
parents: 
40341 
diff
changeset
 | 
257  | 
\values (e.g., \"0.6 0.95\")")  | 
| 
43064
 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 
blanchet 
parents: 
43057 
diff
changeset
 | 
258  | 
fun lookup_option lookup' name =  | 
| 
38589
 
b03f8fe043ec
added "max_relevant_per_iter" option to Sledgehammer
 
blanchet 
parents: 
38282 
diff
changeset
 | 
259  | 
case lookup name of  | 
| 
 
b03f8fe043ec
added "max_relevant_per_iter" option to Sledgehammer
 
blanchet 
parents: 
38282 
diff
changeset
 | 
260  | 
SOME "smart" => NONE  | 
| 
43064
 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 
blanchet 
parents: 
43057 
diff
changeset
 | 
261  | 
| _ => SOME (lookup' name)  | 
| 43021 | 262  | 
val debug = mode <> Auto_Try andalso lookup_bool "debug"  | 
263  | 
val verbose = debug orelse (mode <> Auto_Try andalso lookup_bool "verbose")  | 
|
| 
36143
 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 
blanchet 
parents: 
36142 
diff
changeset
 | 
264  | 
val overlord = lookup_bool "overlord"  | 
| 53800 | 265  | 
val spy = lookup_bool "spy"  | 
| 
42995
 
e23f61546cf0
always run Sledgehammer synchronously in the jEdit interface (until the multithreading support for Proof General is ported)
 
blanchet 
parents: 
42944 
diff
changeset
 | 
266  | 
val blocking =  | 
| 43021 | 267  | 
Isabelle_Process.is_active () orelse mode <> Normal orelse debug orelse  | 
| 
42995
 
e23f61546cf0
always run Sledgehammer synchronously in the jEdit interface (until the multithreading support for Proof General is ported)
 
blanchet 
parents: 
42944 
diff
changeset
 | 
268  | 
lookup_bool "blocking"  | 
| 
40059
 
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
 
blanchet 
parents: 
39335 
diff
changeset
 | 
269  | 
val provers = lookup_string "provers" |> space_explode " "  | 
| 43021 | 270  | 
|> mode = Auto_Try ? single o hd  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43572 
diff
changeset
 | 
271  | 
val type_enc =  | 
| 43021 | 272  | 
if mode = Auto_Try then  | 
| 
43569
 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 
blanchet 
parents: 
43353 
diff
changeset
 | 
273  | 
NONE  | 
| 
43626
 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 
blanchet 
parents: 
43572 
diff
changeset
 | 
274  | 
else case lookup_string "type_enc" of  | 
| 
43569
 
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
 
blanchet 
parents: 
43353 
diff
changeset
 | 
275  | 
"smart" => NONE  | 
| 
52031
 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 
blanchet 
parents: 
52018 
diff
changeset
 | 
276  | 
| s => (type_enc_of_string Strict s; SOME s)  | 
| 46301 | 277  | 
val strict = mode = Auto_Try orelse lookup_bool "strict"  | 
| 45520 | 278  | 
val lam_trans = lookup_option lookup_string "lam_trans"  | 
| 
46409
 
d4754183ccce
made option available to users (mostly for experiments)
 
blanchet 
parents: 
46398 
diff
changeset
 | 
279  | 
val uncurried_aliases = lookup_option lookup_bool "uncurried_aliases"  | 
| 48321 | 280  | 
val learn = lookup_bool "learn"  | 
| 
48314
 
ee33ba3c0e05
added option to control which fact filter is used
 
blanchet 
parents: 
48308 
diff
changeset
 | 
281  | 
val fact_filter = lookup_option lookup_string "fact_filter"  | 
| 
 
ee33ba3c0e05
added option to control which fact filter is used
 
blanchet 
parents: 
48308 
diff
changeset
 | 
282  | 
val max_facts = lookup_option lookup_int "max_facts"  | 
| 48293 | 283  | 
val fact_thresholds = lookup_real_pair "fact_thresholds"  | 
| 
47962
 
137883567114
lower the monomorphization thresholds for less scalable provers
 
blanchet 
parents: 
47642 
diff
changeset
 | 
284  | 
val max_mono_iters = lookup_option lookup_int "max_mono_iters"  | 
| 
 
137883567114
lower the monomorphization thresholds for less scalable provers
 
blanchet 
parents: 
47642 
diff
changeset
 | 
285  | 
val max_new_mono_instances =  | 
| 
 
137883567114
lower the monomorphization thresholds for less scalable provers
 
blanchet 
parents: 
47642 
diff
changeset
 | 
286  | 
lookup_option lookup_int "max_new_mono_instances"  | 
| 
51190
 
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
 
blanchet 
parents: 
51189 
diff
changeset
 | 
287  | 
val isar_proofs = lookup_option lookup_bool "isar_proofs"  | 
| 
51130
 
76d68444cd59
renamed sledgehammer_shrink to sledgehammer_compress
 
smolkas 
parents: 
50824 
diff
changeset
 | 
288  | 
val isar_compress = Real.max (1.0, lookup_real "isar_compress")  | 
| 
52632
 
23393c31c7fe
added sledgehammer parameters isar_try0 and isar_minimize and their negative aliases
 
smolkas 
parents: 
52556 
diff
changeset
 | 
289  | 
val isar_try0 = lookup_bool "isar_try0"  | 
| 45706 | 290  | 
val slice = mode <> Auto_Try andalso lookup_bool "slice"  | 
| 
45707
 
6bf7eec9b153
added "minimize" option for more control over automatic minimization
 
blanchet 
parents: 
45706 
diff
changeset
 | 
291  | 
val minimize =  | 
| 
 
6bf7eec9b153
added "minimize" option for more control over automatic minimization
 
blanchet 
parents: 
45706 
diff
changeset
 | 
292  | 
if mode = Auto_Try then NONE else lookup_option lookup_bool "minimize"  | 
| 50557 | 293  | 
val timeout = if mode = Auto_Try then NONE else lookup_time "timeout"  | 
| 
43015
 
21b6baec55b1
renamed "metis_timeout" to "preplay_timeout" and continued implementation
 
blanchet 
parents: 
43013 
diff
changeset
 | 
294  | 
val preplay_timeout =  | 
| 50557 | 295  | 
if mode = Auto_Try then SOME Time.zeroTime  | 
296  | 
else lookup_time "preplay_timeout"  | 
|
| 
38985
 
162bbbea4e4d
added "expect" feature of Nitpick to Sledgehammer, for regression testing
 
blanchet 
parents: 
38982 
diff
changeset
 | 
297  | 
val expect = lookup_string "expect"  | 
| 35963 | 298  | 
in  | 
| 53800 | 299  | 
    {debug = debug, verbose = verbose, overlord = overlord, spy = spy, blocking = blocking,
 | 
300  | 
provers = provers, type_enc = type_enc, strict = strict, lam_trans = lam_trans,  | 
|
301  | 
uncurried_aliases = uncurried_aliases, learn = learn, fact_filter = fact_filter,  | 
|
302  | 
max_facts = max_facts, fact_thresholds = fact_thresholds, max_mono_iters = max_mono_iters,  | 
|
| 
49918
 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 
blanchet 
parents: 
49632 
diff
changeset
 | 
303  | 
max_new_mono_instances = max_new_mono_instances, isar_proofs = isar_proofs,  | 
| 53800 | 304  | 
isar_compress = isar_compress, isar_try0 = isar_try0, slice = slice, minimize = minimize,  | 
305  | 
timeout = timeout, preplay_timeout = preplay_timeout, expect = expect}  | 
|
| 35963 | 306  | 
end  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
307  | 
|
| 
44720
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
308  | 
fun get_params mode = extract_params mode o default_raw_params  | 
| 43021 | 309  | 
fun default_params thy = get_params Normal thy o map (apsnd single)  | 
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
310  | 
|
| 
36373
 
66af0a49de39
move some sledgehammer stuff out of "atp_manager.ML"
 
blanchet 
parents: 
36371 
diff
changeset
 | 
311  | 
(* Sledgehammer the given subgoal *)  | 
| 
 
66af0a49de39
move some sledgehammer stuff out of "atp_manager.ML"
 
blanchet 
parents: 
36371 
diff
changeset
 | 
312  | 
|
| 
45519
 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 
blanchet 
parents: 
45514 
diff
changeset
 | 
313  | 
val default_minimize_prover = metisN  | 
| 43051 | 314  | 
|
| 
44720
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
315  | 
fun is_raw_param_relevant_for_minimize (name, _) =  | 
| 53758 | 316  | 
not (member (op =) params_not_for_minimize name)  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51879 
diff
changeset
 | 
317  | 
fun string_of_raw_param (key, values) =  | 
| 42776 | 318  | 
key ^ (case implode_param values of "" => "" | value => " = " ^ value)  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51879 
diff
changeset
 | 
319  | 
fun nice_string_of_raw_param (p as (key, ["false"])) =  | 
| 50748 | 320  | 
(case AList.find (op =) negated_alias_params key of  | 
321  | 
[neg] => neg  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51879 
diff
changeset
 | 
322  | 
| _ => string_of_raw_param p)  | 
| 
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51879 
diff
changeset
 | 
323  | 
| nice_string_of_raw_param p = string_of_raw_param p  | 
| 
36281
 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 
blanchet 
parents: 
36235 
diff
changeset
 | 
324  | 
|
| 45520 | 325  | 
fun minimize_command override_params i more_override_params prover_name  | 
326  | 
fact_names =  | 
|
| 43057 | 327  | 
let  | 
328  | 
val params =  | 
|
| 45520 | 329  | 
(override_params  | 
330  | 
|> filter_out (AList.defined (op =) more_override_params o fst)) @  | 
|
331  | 
more_override_params  | 
|
| 43057 | 332  | 
|> filter is_raw_param_relevant_for_minimize  | 
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51879 
diff
changeset
 | 
333  | 
|> map nice_string_of_raw_param  | 
| 43057 | 334  | 
|> (if prover_name = default_minimize_prover then I else cons prover_name)  | 
335  | 
|> space_implode ", "  | 
|
336  | 
in  | 
|
| 
48392
 
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
 
blanchet 
parents: 
48388 
diff
changeset
 | 
337  | 
sledgehammerN ^ " " ^ minN ^  | 
| 43057 | 338  | 
(if params = "" then "" else enclose " [" "]" params) ^  | 
339  | 
    " (" ^ space_implode " " fact_names ^ ")" ^
 | 
|
340  | 
(if i = 1 then "" else " " ^ string_of_int i)  | 
|
341  | 
end  | 
|
| 
36281
 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 
blanchet 
parents: 
36235 
diff
changeset
 | 
342  | 
|
| 50604 | 343  | 
val default_learn_prover_timeout = 2.0  | 
| 
48392
 
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
 
blanchet 
parents: 
48388 
diff
changeset
 | 
344  | 
|
| 48292 | 345  | 
fun hammer_away override_params subcommand opt_i fact_override state =  | 
| 35963 | 346  | 
let  | 
| 40069 | 347  | 
val ctxt = Proof.context_of state  | 
| 
44720
 
f3a8c19708c8
fixed handling of "sledgehammer_params", so that "sledgehammer_params [e]" is really the same as "sledgehammer_params [provers = e]"
 
blanchet 
parents: 
44651 
diff
changeset
 | 
348  | 
val override_params = override_params |> map (normalize_raw_param ctxt)  | 
| 50052 | 349  | 
val _ = Isabelle_System.mkdir (Path.explode (getenv "ISABELLE_TMP"))  | 
| 
35965
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
350  | 
in  | 
| 
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
351  | 
if subcommand = runN then  | 
| 
36281
 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 
blanchet 
parents: 
36235 
diff
changeset
 | 
352  | 
let val i = the_default 1 opt_i in  | 
| 
53048
 
0f76e620561f
more direct sledgehammer configuration via mode = Normal_Result and output_result;
 
wenzelm 
parents: 
52982 
diff
changeset
 | 
353  | 
run_sledgehammer (get_params Normal ctxt override_params) Normal NONE i  | 
| 48292 | 354  | 
fact_override (minimize_command override_params i)  | 
| 39318 | 355  | 
state  | 
356  | 
|> K ()  | 
|
| 
36281
 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 
blanchet 
parents: 
36235 
diff
changeset
 | 
357  | 
end  | 
| 
43008
 
bb212c2ad238
renamed "minimize" to "min" to make Sledgehammer output a little bit more concise
 
blanchet 
parents: 
42997 
diff
changeset
 | 
358  | 
else if subcommand = minN then  | 
| 
48384
 
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
 
blanchet 
parents: 
48383 
diff
changeset
 | 
359  | 
let  | 
| 
48395
 
85a7fb65507a
learning should honor the fact override and the chained facts
 
blanchet 
parents: 
48392 
diff
changeset
 | 
360  | 
val ctxt = ctxt |> Config.put instantiate_inducts false  | 
| 
48384
 
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
 
blanchet 
parents: 
48383 
diff
changeset
 | 
361  | 
val i = the_default 1 opt_i  | 
| 48400 | 362  | 
val params =  | 
| 
48384
 
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
 
blanchet 
parents: 
48383 
diff
changeset
 | 
363  | 
          get_params Minimize ctxt (("provers", [default_minimize_prover]) ::
 | 
| 
 
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
 
blanchet 
parents: 
48383 
diff
changeset
 | 
364  | 
override_params)  | 
| 
 
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
 
blanchet 
parents: 
48383 
diff
changeset
 | 
365  | 
val goal = prop_of (#goal (Proof.goal state))  | 
| 
 
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
 
blanchet 
parents: 
48383 
diff
changeset
 | 
366  | 
val facts = nearly_all_facts ctxt false fact_override Symtab.empty  | 
| 
 
83dc102041e6
learn on explicit "min" command but do the learning in a thread, since it may take a couple of seconds
 
blanchet 
parents: 
48383 
diff
changeset
 | 
367  | 
Termtab.empty [] [] goal  | 
| 
48399
 
4bacc8983b3d
learn from SMT proofs when they can be minimized by Metis
 
blanchet 
parents: 
48397 
diff
changeset
 | 
368  | 
fun learn prover = mash_learn_proof ctxt params prover goal facts  | 
| 
 
4bacc8983b3d
learn from SMT proofs when they can be minimized by Metis
 
blanchet 
parents: 
48397 
diff
changeset
 | 
369  | 
in run_minimize params learn i (#add fact_override) state end  | 
| 
35965
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
370  | 
else if subcommand = messagesN then  | 
| 
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
371  | 
messages opt_i  | 
| 
41727
 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 
blanchet 
parents: 
41472 
diff
changeset
 | 
372  | 
else if subcommand = supported_proversN then  | 
| 
 
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
 
blanchet 
parents: 
41472 
diff
changeset
 | 
373  | 
supported_provers ctxt  | 
| 
50719
 
58b0b44da54a
renamed "kill" subcommand to avoid clash with "kill" keyword (which confuses Proof General and results in strange syntax highlighting)
 
blanchet 
parents: 
50604 
diff
changeset
 | 
374  | 
else if subcommand = kill_allN then  | 
| 
53148
 
c898409d8630
fixed subtle bug with "take" + thread overlord through
 
blanchet 
parents: 
53142 
diff
changeset
 | 
375  | 
(kill_provers ();  | 
| 
 
c898409d8630
fixed subtle bug with "take" + thread overlord through
 
blanchet 
parents: 
53142 
diff
changeset
 | 
376  | 
kill_learners ctxt (get_params Normal ctxt override_params))  | 
| 
48319
 
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
 
blanchet 
parents: 
48314 
diff
changeset
 | 
377  | 
else if subcommand = running_proversN then  | 
| 
 
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
 
blanchet 
parents: 
48314 
diff
changeset
 | 
378  | 
running_provers ()  | 
| 48383 | 379  | 
else if subcommand = unlearnN then  | 
| 
53148
 
c898409d8630
fixed subtle bug with "take" + thread overlord through
 
blanchet 
parents: 
53142 
diff
changeset
 | 
380  | 
mash_unlearn ctxt (get_params Normal ctxt override_params)  | 
| 
50484
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
381  | 
else if subcommand = learn_isarN orelse subcommand = learn_proverN orelse  | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
382  | 
subcommand = relearn_isarN orelse subcommand = relearn_proverN then  | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
383  | 
(if subcommand = relearn_isarN orelse subcommand = relearn_proverN then  | 
| 
53148
 
c898409d8630
fixed subtle bug with "take" + thread overlord through
 
blanchet 
parents: 
53142 
diff
changeset
 | 
384  | 
mash_unlearn ctxt (get_params Normal ctxt override_params)  | 
| 
48392
 
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
 
blanchet 
parents: 
48388 
diff
changeset
 | 
385  | 
else  | 
| 
 
ca998fa08cd9
added "learn_from_atp" command to MaSh, for patient users
 
blanchet 
parents: 
48388 
diff
changeset
 | 
386  | 
();  | 
| 
50484
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
387  | 
mash_learn ctxt  | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
388  | 
(get_params Normal ctxt  | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
389  | 
                ([("timeout",
 | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
390  | 
[string_of_real default_learn_prover_timeout]),  | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
391  | 
                  ("slice", ["false"])] @
 | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
392  | 
override_params @  | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
393  | 
                 [("minimize", ["true"]),
 | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
394  | 
                  ("preplay_timeout", ["0"])]))
 | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
395  | 
fact_override (#facts (Proof.goal state))  | 
| 
 
8ec31bdb9d36
adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
 
blanchet 
parents: 
50052 
diff
changeset
 | 
396  | 
(subcommand = learn_proverN orelse subcommand = relearn_proverN))  | 
| 
48319
 
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
 
blanchet 
parents: 
48314 
diff
changeset
 | 
397  | 
else if subcommand = running_learnersN then  | 
| 
 
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
 
blanchet 
parents: 
48314 
diff
changeset
 | 
398  | 
running_learners ()  | 
| 
35965
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
399  | 
else if subcommand = refresh_tptpN then  | 
| 
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
400  | 
refresh_systems_on_tptp ()  | 
| 
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
401  | 
else  | 
| 
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
402  | 
      error ("Unknown subcommand: " ^ quote subcommand ^ ".")
 | 
| 
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
403  | 
end  | 
| 35963 | 404  | 
|
| 48292 | 405  | 
fun sledgehammer_trans (((subcommand, params), fact_override), opt_i) =  | 
| 
51557
 
4e4b56b7a3a5
more robust access Toplevel.proof_of -- prefer warning via Toplevel.unknown_proof over hard crash (notably for skipped proofs);
 
wenzelm 
parents: 
51200 
diff
changeset
 | 
406  | 
Toplevel.unknown_proof o  | 
| 48292 | 407  | 
Toplevel.keep (hammer_away params subcommand opt_i fact_override  | 
| 
35966
 
f8c738abaed8
honor the newly introduced Sledgehammer parameters and fixed the parsing;
 
blanchet 
parents: 
35965 
diff
changeset
 | 
408  | 
o Toplevel.proof_of)  | 
| 35963 | 409  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51879 
diff
changeset
 | 
410  | 
fun string_of_raw_param (name, value) = name ^ " = " ^ implode_param value  | 
| 35963 | 411  | 
|
| 
35965
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
412  | 
fun sledgehammer_params_trans params =  | 
| 35963 | 413  | 
Toplevel.theory  | 
| 
35965
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
414  | 
(fold set_default_raw_param params  | 
| 39318 | 415  | 
#> tap (fn thy =>  | 
| 42361 | 416  | 
let val ctxt = Proof_Context.init_global thy in  | 
| 40069 | 417  | 
                    writeln ("Default parameters for Sledgehammer:\n" ^
 | 
418  | 
(case default_raw_params ctxt |> rev of  | 
|
419  | 
[] => "none"  | 
|
420  | 
| params =>  | 
|
| 
51998
 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 
blanchet 
parents: 
51879 
diff
changeset
 | 
421  | 
params |> map string_of_raw_param  | 
| 
41258
 
73401632a80c
convenient syntax for setting provers -- useful for debugging, not for general consumption and hence not documented
 
blanchet 
parents: 
41208 
diff
changeset
 | 
422  | 
|> sort_strings |> cat_lines))  | 
| 40069 | 423  | 
end))  | 
| 35963 | 424  | 
|
| 46949 | 425  | 
val parse_query_bang = @{keyword "?"} || @{keyword "!"} || @{keyword "!!"}
 | 
| 42776 | 426  | 
val parse_key =  | 
| 44768 | 427  | 
Scan.repeat1 (Parse.typ_group || parse_query_bang) >> implode_param  | 
| 
42581
 
398fff2c6b8f
use ! for mildly unsound and !! for very unsound encodings
 
blanchet 
parents: 
42579 
diff
changeset
 | 
428  | 
val parse_value =  | 
| 44768 | 429  | 
Scan.repeat1 (Parse.xname || Parse.float_number || parse_query_bang)  | 
| 46949 | 430  | 
val parse_param = parse_key -- Scan.optional (@{keyword "="} |-- parse_value) []
 | 
| 
36960
 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 
wenzelm 
parents: 
36924 
diff
changeset
 | 
431  | 
val parse_params = Scan.optional (Args.bracks (Parse.list parse_param)) []  | 
| 
35965
 
0fce6db7babd
added a syntax for specifying facts to Sledgehammer;
 
blanchet 
parents: 
35963 
diff
changeset
 | 
432  | 
val parse_fact_refs =  | 
| 38996 | 433  | 
Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) Parse_Spec.xthm)  | 
| 48293 | 434  | 
val parse_fact_override_chunk =  | 
| 48292 | 435  | 
(Args.add |-- Args.colon |-- parse_fact_refs >> add_fact_override)  | 
436  | 
|| (Args.del |-- Args.colon |-- parse_fact_refs >> del_fact_override)  | 
|
437  | 
|| (parse_fact_refs >> only_fact_override)  | 
|
438  | 
val parse_fact_override =  | 
|
| 48293 | 439  | 
Scan.optional (Args.parens (Scan.repeat parse_fact_override_chunk  | 
| 48292 | 440  | 
>> merge_fact_overrides))  | 
441  | 
no_fact_override  | 
|
| 35963 | 442  | 
|
443  | 
val _ =  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
444  | 
  Outer_Syntax.improper_command @{command_spec "sledgehammer"}
 | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
445  | 
"search for first-order proof using automatic theorem provers"  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
446  | 
((Scan.optional Parse.short_ident runN -- parse_params  | 
| 48292 | 447  | 
-- parse_fact_override -- Scan.option Parse.nat) #>> sledgehammer_trans)  | 
| 35963 | 448  | 
val _ =  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
449  | 
  Outer_Syntax.command @{command_spec "sledgehammer_params"}
 | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
450  | 
"set and display the default parameters for Sledgehammer"  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46949 
diff
changeset
 | 
451  | 
(parse_params #>> sledgehammer_params_trans)  | 
| 
36394
 
1a48d18449d8
move "neg_clausify" method and "clausify" attribute to "sledgehammer_isar.ML"
 
blanchet 
parents: 
36378 
diff
changeset
 | 
452  | 
|
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
453  | 
fun try_sledgehammer auto state =  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
454  | 
let  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
455  | 
val ctxt = Proof.context_of state  | 
| 43021 | 456  | 
val mode = if auto then Auto_Try else Try  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
457  | 
val i = 1  | 
| 
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
458  | 
in  | 
| 
53048
 
0f76e620561f
more direct sledgehammer configuration via mode = Normal_Result and output_result;
 
wenzelm 
parents: 
52982 
diff
changeset
 | 
459  | 
run_sledgehammer (get_params mode ctxt []) mode NONE i no_fact_override  | 
| 
43020
 
abb5d1f907e4
added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
 
blanchet 
parents: 
43018 
diff
changeset
 | 
460  | 
(minimize_command [] i) state  | 
| 40931 | 461  | 
end  | 
| 39318 | 462  | 
|
| 
52641
 
c56b6fa636e8
hybrid "auto" tool setup, for TTY (within theory) and PIDE (global print function);
 
wenzelm 
parents: 
52639 
diff
changeset
 | 
463  | 
val _ = Try.tool_setup (sledgehammerN, (40, @{option auto_sledgehammer}, try_sledgehammer))
 | 
| 39318 | 464  | 
|
| 
52908
 
3461985dcbc3
dockable window for Sledgehammer, based on asynchronous/parallel query operation;
 
wenzelm 
parents: 
52653 
diff
changeset
 | 
465  | 
val _ =  | 
| 
52982
 
8e78bd316a53
clarified Query_Operation.register: avoid hard-wired parallel policy;
 
wenzelm 
parents: 
52908 
diff
changeset
 | 
466  | 
  Query_Operation.register sledgehammerN (fn {state = st, args, output_result} =>
 | 
| 
 
8e78bd316a53
clarified Query_Operation.register: avoid hard-wired parallel policy;
 
wenzelm 
parents: 
52908 
diff
changeset
 | 
467  | 
(case try Toplevel.proof_of st of  | 
| 
 
8e78bd316a53
clarified Query_Operation.register: avoid hard-wired parallel policy;
 
wenzelm 
parents: 
52908 
diff
changeset
 | 
468  | 
SOME state =>  | 
| 
 
8e78bd316a53
clarified Query_Operation.register: avoid hard-wired parallel policy;
 
wenzelm 
parents: 
52908 
diff
changeset
 | 
469  | 
let  | 
| 
53057
 
e18a028b345c
prefer system option sledgehammer_timeout, with standard GUI in jEdit Plugin Options;
 
wenzelm 
parents: 
53055 
diff
changeset
 | 
470  | 
val [provers_arg, isar_proofs_arg] = args;  | 
| 
52982
 
8e78bd316a53
clarified Query_Operation.register: avoid hard-wired parallel policy;
 
wenzelm 
parents: 
52908 
diff
changeset
 | 
471  | 
val ctxt = Proof.context_of state  | 
| 
52908
 
3461985dcbc3
dockable window for Sledgehammer, based on asynchronous/parallel query operation;
 
wenzelm 
parents: 
52653 
diff
changeset
 | 
472  | 
|
| 
52982
 
8e78bd316a53
clarified Query_Operation.register: avoid hard-wired parallel policy;
 
wenzelm 
parents: 
52908 
diff
changeset
 | 
473  | 
val override_params =  | 
| 
53053
 
6a3320758f0d
always enable "minimize" to simplify interaction model;
 
wenzelm 
parents: 
53052 
diff
changeset
 | 
474  | 
            ((if provers_arg = "" then [] else [("provers", space_explode " " provers_arg)]) @
 | 
| 
53057
 
e18a028b345c
prefer system option sledgehammer_timeout, with standard GUI in jEdit Plugin Options;
 
wenzelm 
parents: 
53055 
diff
changeset
 | 
475  | 
              [("isar_proofs", [isar_proofs_arg]),
 | 
| 
53053
 
6a3320758f0d
always enable "minimize" to simplify interaction model;
 
wenzelm 
parents: 
53052 
diff
changeset
 | 
476  | 
               ("blocking", ["true"]),
 | 
| 
 
6a3320758f0d
always enable "minimize" to simplify interaction model;
 
wenzelm 
parents: 
53052 
diff
changeset
 | 
477  | 
               ("minimize", ["true"]),
 | 
| 
 
6a3320758f0d
always enable "minimize" to simplify interaction model;
 
wenzelm 
parents: 
53052 
diff
changeset
 | 
478  | 
               ("debug", ["false"]),
 | 
| 
 
6a3320758f0d
always enable "minimize" to simplify interaction model;
 
wenzelm 
parents: 
53052 
diff
changeset
 | 
479  | 
               ("verbose", ["false"]),
 | 
| 
 
6a3320758f0d
always enable "minimize" to simplify interaction model;
 
wenzelm 
parents: 
53052 
diff
changeset
 | 
480  | 
               ("overlord", ["false"])])
 | 
| 
52982
 
8e78bd316a53
clarified Query_Operation.register: avoid hard-wired parallel policy;
 
wenzelm 
parents: 
52908 
diff
changeset
 | 
481  | 
|> map (normalize_raw_param ctxt)  | 
| 
53048
 
0f76e620561f
more direct sledgehammer configuration via mode = Normal_Result and output_result;
 
wenzelm 
parents: 
52982 
diff
changeset
 | 
482  | 
val _ =  | 
| 
53052
 
a0db255af8c5
sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
 
wenzelm 
parents: 
53049 
diff
changeset
 | 
483  | 
run_sledgehammer (get_params Normal ctxt override_params) Normal (SOME output_result) 1  | 
| 53049 | 484  | 
no_fact_override (minimize_command override_params 1) state  | 
| 
52982
 
8e78bd316a53
clarified Query_Operation.register: avoid hard-wired parallel policy;
 
wenzelm 
parents: 
52908 
diff
changeset
 | 
485  | 
in () end  | 
| 
53055
 
0fe8a9972eda
some protocol to determine provers according to ML;
 
wenzelm 
parents: 
53053 
diff
changeset
 | 
486  | 
| NONE => error "Unknown proof context"))  | 
| 
 
0fe8a9972eda
some protocol to determine provers according to ML;
 
wenzelm 
parents: 
53053 
diff
changeset
 | 
487  | 
|
| 
 
0fe8a9972eda
some protocol to determine provers according to ML;
 
wenzelm 
parents: 
53053 
diff
changeset
 | 
488  | 
val _ = Session.protocol_handler "isabelle.Sledgehammer_Params$Handler"  | 
| 
 
0fe8a9972eda
some protocol to determine provers according to ML;
 
wenzelm 
parents: 
53053 
diff
changeset
 | 
489  | 
|
| 
 
0fe8a9972eda
some protocol to determine provers according to ML;
 
wenzelm 
parents: 
53053 
diff
changeset
 | 
490  | 
val _ = Isabelle_Process.protocol_command "Sledgehammer.provers"  | 
| 
 
0fe8a9972eda
some protocol to determine provers according to ML;
 
wenzelm 
parents: 
53053 
diff
changeset
 | 
491  | 
(fn [] =>  | 
| 
 
0fe8a9972eda
some protocol to determine provers according to ML;
 
wenzelm 
parents: 
53053 
diff
changeset
 | 
492  | 
let  | 
| 
 
0fe8a9972eda
some protocol to determine provers according to ML;
 
wenzelm 
parents: 
53053 
diff
changeset
 | 
493  | 
      val this_ctxt = @{context}
 | 
| 
 
0fe8a9972eda
some protocol to determine provers according to ML;
 
wenzelm 
parents: 
53053 
diff
changeset
 | 
494  | 
val provers = space_implode " " (#provers (default_params this_ctxt []))  | 
| 
 
0fe8a9972eda
some protocol to determine provers according to ML;
 
wenzelm 
parents: 
53053 
diff
changeset
 | 
495  | 
in Output.protocol_message Markup.sledgehammer_provers provers end)  | 
| 
52908
 
3461985dcbc3
dockable window for Sledgehammer, based on asynchronous/parallel query operation;
 
wenzelm 
parents: 
52653 
diff
changeset
 | 
496  | 
|
| 
35866
 
513074557e06
move the Sledgehammer Isar commands together into one file;
 
blanchet 
parents:  
diff
changeset
 | 
497  | 
end;  |