author | huffman |
Thu, 17 Nov 2011 06:01:47 +0100 | |
changeset 45532 | 74b17a0881b3 |
parent 45523 | 741f308c0f36 |
child 45705 | a25ff4283352 |
permissions | -rw-r--r-- |
43197 | 1 |
(* Title: HOL/Metis_Examples/Type_Encodings.thy |
43162 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
3 |
||
43197 | 4 |
Example that exercises Metis's (and hence Sledgehammer's) type encodings. |
43162 | 5 |
*) |
6 |
||
43197 | 7 |
header {* |
8 |
Example that Exercises Metis's (and Hence Sledgehammer's) Type Encodings |
|
9 |
*} |
|
10 |
||
11 |
theory Type_Encodings |
|
44412
c8b847625584
include all encodings in tests, now that the incompleteness of some encodings has been addressed
blanchet
parents:
44408
diff
changeset
|
12 |
imports Main |
43162 | 13 |
begin |
14 |
||
43197 | 15 |
declare [[metis_new_skolemizer]] |
16 |
||
17 |
sledgehammer_params [prover = e, blocking, timeout = 10, preplay_timeout = 0] |
|
18 |
||
19 |
||
43162 | 20 |
text {* Setup for testing Metis exhaustively *} |
21 |
||
22 |
lemma fork: "P \<Longrightarrow> P \<Longrightarrow> P" by assumption |
|
23 |
||
24 |
ML {* |
|
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43212
diff
changeset
|
25 |
val type_encs = |
43208 | 26 |
["erased", |
43989 | 27 |
"poly_guards", |
44768 | 28 |
"poly_guards?", |
29 |
"poly_guards??", |
|
44813 | 30 |
"poly_guards@?", |
44768 | 31 |
"poly_guards!", |
32 |
"poly_guards!!", |
|
44813 | 33 |
"poly_guards@!", |
43208 | 34 |
"poly_tags", |
44768 | 35 |
"poly_tags?", |
36 |
"poly_tags??", |
|
44815 | 37 |
"poly_tags@?", |
44768 | 38 |
"poly_tags!", |
39 |
"poly_tags!!", |
|
44815 | 40 |
"poly_tags@!", |
43208 | 41 |
"poly_args", |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44412
diff
changeset
|
42 |
"raw_mono_guards", |
44768 | 43 |
"raw_mono_guards?", |
44 |
"raw_mono_guards??", |
|
44813 | 45 |
"raw_mono_guards@?", |
44768 | 46 |
"raw_mono_guards!", |
47 |
"raw_mono_guards!!", |
|
44813 | 48 |
"raw_mono_guards@!", |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44412
diff
changeset
|
49 |
"raw_mono_tags", |
44768 | 50 |
"raw_mono_tags?", |
51 |
"raw_mono_tags??", |
|
44815 | 52 |
"raw_mono_tags@?", |
44768 | 53 |
"raw_mono_tags!", |
54 |
"raw_mono_tags!!", |
|
44815 | 55 |
"raw_mono_tags@!", |
44494
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents:
44412
diff
changeset
|
56 |
"raw_mono_args", |
43989 | 57 |
"mono_guards", |
44768 | 58 |
"mono_guards?", |
59 |
"mono_guards??", |
|
60 |
"mono_guards!", |
|
61 |
"mono_guards!!", |
|
43208 | 62 |
"mono_tags", |
63 |
"mono_tags?", |
|
44768 | 64 |
"mono_tags??", |
43208 | 65 |
"mono_tags!", |
44768 | 66 |
"mono_tags!!", |
67 |
"mono_args"] |
|
43162 | 68 |
|
43212 | 69 |
fun metis_exhaust_tac ctxt ths = |
43162 | 70 |
let |
43172 | 71 |
fun tac [] st = all_tac st |
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43212
diff
changeset
|
72 |
| tac (type_enc :: type_encs) st = |
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43212
diff
changeset
|
73 |
st (* |> tap (fn _ => tracing (PolyML.makestring type_enc)) *) |
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43212
diff
changeset
|
74 |
|> ((if null type_encs then all_tac else rtac @{thm fork} 1) |
45523 | 75 |
THEN Metis_Tactic.metis_tac [type_enc] "combinators" ctxt ths 1 |
43172 | 76 |
THEN COND (has_fewer_prems 2) all_tac no_tac |
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43212
diff
changeset
|
77 |
THEN tac type_encs) |
43162 | 78 |
in tac end |
79 |
*} |
|
80 |
||
43212 | 81 |
method_setup metis_exhaust = {* |
43162 | 82 |
Attrib.thms >> |
43626
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents:
43212
diff
changeset
|
83 |
(fn ths => fn ctxt => SIMPLE_METHOD (metis_exhaust_tac ctxt ths type_encs)) |
43162 | 84 |
*} "exhaustively run the new Metis with all type encodings" |
85 |
||
86 |
||
43197 | 87 |
text {* Miscellaneous tests *} |
43162 | 88 |
|
89 |
lemma "x = y \<Longrightarrow> y = x" |
|
43212 | 90 |
by metis_exhaust |
43162 | 91 |
|
92 |
lemma "[a] = [1 + 1] \<Longrightarrow> a = 1 + (1::int)" |
|
43212 | 93 |
by (metis_exhaust last.simps) |
43162 | 94 |
|
95 |
lemma "map Suc [0] = [Suc 0]" |
|
43212 | 96 |
by (metis_exhaust map.simps) |
43162 | 97 |
|
98 |
lemma "map Suc [1 + 1] = [Suc 2]" |
|
43212 | 99 |
by (metis_exhaust map.simps nat_1_add_1) |
43162 | 100 |
|
101 |
lemma "map Suc [2] = [Suc (1 + 1)]" |
|
43212 | 102 |
by (metis_exhaust map.simps nat_1_add_1) |
43162 | 103 |
|
104 |
definition "null xs = (xs = [])" |
|
105 |
||
106 |
lemma "P (null xs) \<Longrightarrow> null xs \<Longrightarrow> xs = []" |
|
43212 | 107 |
by (metis_exhaust null_def) |
43162 | 108 |
|
109 |
lemma "(0::nat) + 0 = 0" |
|
45502
6246bef495ff
avoid theorem references like 'semiring_norm(111)'
huffman
parents:
44815
diff
changeset
|
110 |
by (metis_exhaust add_0_left) |
43162 | 111 |
|
112 |
end |