author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 82458 | c7bd567723b1 |
permissions | -rw-r--r-- |
82360 | 1 |
(* Title: HOL/Try0_HOL.thy |
2 |
Author: Jasmin Blanchette, LMU Muenchen |
|
3 |
Author: Martin Desharnais, LMU Muenchen |
|
4 |
Author: Fabian Huch, TU Muenchen |
|
5 |
*) |
|
6 |
theory Try0_HOL |
|
7 |
imports Try0 Presburger |
|
8 |
begin |
|
9 |
||
82363
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
10 |
ML \<open> |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
11 |
signature TRY0_HOL = |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
12 |
sig |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
13 |
val silence_methods : Proof.context -> Proof.context |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
14 |
end |
82360 | 15 |
|
82363
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
16 |
structure Try0_HOL : TRY0_HOL = struct |
82360 | 17 |
|
82363
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
18 |
(* Makes reconstructor tools as silent as possible. The "set_visible" calls suppresses "Unification |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
19 |
bound exceeded" warnings and the like. *) |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
20 |
fun silence_methods ctxt = |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
21 |
ctxt |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
22 |
|> Config.put Metis_Tactic.verbose false |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
23 |
|> Simplifier_Trace.disable |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
24 |
|> Context_Position.set_visible false |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
25 |
|> Config.put Unify.unify_trace false |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
26 |
|> Config.put Argo_Tactic.trace "none" |
82360 | 27 |
|
28 |
local |
|
29 |
||
82363
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
30 |
open Try0_Util |
82360 | 31 |
|
32 |
(* name * (run_if_auto_try * (all_goals * tags)) *) |
|
33 |
val raw_named_methods = |
|
82363
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
34 |
[("auto", (true, (true, full_attrs))), |
82360 | 35 |
("blast", (true, (false, clas_attrs))), |
36 |
("metis", (true, (false, metis_attrs))), |
|
37 |
("argo", (true, (false, no_attrs))), |
|
38 |
("linarith", (true, (false, no_attrs))), |
|
39 |
("presburger", (true, (false, no_attrs))), |
|
40 |
("algebra", (true, (false, no_attrs))), |
|
41 |
("fast", (false, (false, clas_attrs))), |
|
42 |
("fastforce", (false, (false, full_attrs))), |
|
43 |
("force", (false, (false, full_attrs))), |
|
44 |
("meson", (false, (false, metis_attrs))), |
|
45 |
("satx", (false, (false, no_attrs))), |
|
46 |
("order", (true, (false, no_attrs)))] |
|
47 |
||
48 |
in |
|
49 |
||
50 |
val () = raw_named_methods |
|
82361 | 51 |
|> List.app (fn (name, (run_if_auto_try, (all_goals, tags))) => |
82363
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
52 |
let |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
53 |
val meth : Try0.proof_method = |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
54 |
Try0_Util.apply_raw_named_method name all_goals tags silence_methods |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
55 |
in |
82360 | 56 |
Try0.register_proof_method name {run_if_auto_try = run_if_auto_try} meth |
57 |
handle Symtab.DUP _ => () |
|
58 |
end) |
|
59 |
||
60 |
end |
|
82363
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
61 |
|
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82361
diff
changeset
|
62 |
end |
82360 | 63 |
\<close> |
64 |
||
82396 | 65 |
declare [[try0_schedule = " |
82458
c7bd567723b1
removed iprover from try0 because its name is clashing with iProver in Sledgehammer
desharna
parents:
82397
diff
changeset
|
66 |
satx metis | |
82396 | 67 |
order presburger linarith algebra argo | |
68 |
simp auto blast fast fastforce force meson |
|
69 |
"]] |
|
70 |
||
82360 | 71 |
end |