src/HOL/Metis_Examples/Type_Encodings.thy
author blanchet
Mon, 06 Feb 2012 23:01:02 +0100
changeset 46436 6f0f2b71fd69
parent 46369 9ac0c64ad8e7
child 46733 4a03b30e04cb
permissions -rw-r--r--
tuning
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
     1
(*  Title:      HOL/Metis_Examples/Type_Encodings.thy
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
     3
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
     4
Example that exercises Metis's (and hence Sledgehammer's) type encodings.
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
     5
*)
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
     6
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
     7
header {*
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
     8
Example that Exercises Metis's (and Hence Sledgehammer's) Type Encodings
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
     9
*}
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
    10
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
    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
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    13
begin
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    14
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
    15
declare [[metis_new_skolemizer]]
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
    16
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
    17
sledgehammer_params [prover = e, blocking, timeout = 10, preplay_timeout = 0]
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
    18
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    19
text {* Setup for testing Metis exhaustively *}
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    20
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    21
lemma fork: "P \<Longrightarrow> P \<Longrightarrow> P" by assumption
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    22
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    23
ML {*
43626
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43212
diff changeset
    24
val type_encs =
43208
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    25
  ["erased",
43989
eb763b3ff9ed renamed "preds" encodings to "guards"
blanchet
parents: 43629
diff changeset
    26
   "poly_guards",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    27
   "poly_guards?",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    28
   "poly_guards??",
44813
d7094cae7df4 added new guards encoding to test
blanchet
parents: 44768
diff changeset
    29
   "poly_guards@?",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    30
   "poly_guards!",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    31
   "poly_guards!!",
44813
d7094cae7df4 added new guards encoding to test
blanchet
parents: 44768
diff changeset
    32
   "poly_guards@!",
43208
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    33
   "poly_tags",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    34
   "poly_tags?",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    35
   "poly_tags??",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    36
   "poly_tags!",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    37
   "poly_tags!!",
43208
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    38
   "poly_args",
44494
a77901b3774e rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents: 44412
diff changeset
    39
   "raw_mono_guards",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    40
   "raw_mono_guards?",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    41
   "raw_mono_guards??",
44813
d7094cae7df4 added new guards encoding to test
blanchet
parents: 44768
diff changeset
    42
   "raw_mono_guards@?",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    43
   "raw_mono_guards!",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    44
   "raw_mono_guards!!",
44813
d7094cae7df4 added new guards encoding to test
blanchet
parents: 44768
diff changeset
    45
   "raw_mono_guards@!",
44494
a77901b3774e rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents: 44412
diff changeset
    46
   "raw_mono_tags",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    47
   "raw_mono_tags?",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    48
   "raw_mono_tags??",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    49
   "raw_mono_tags!",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    50
   "raw_mono_tags!!",
44494
a77901b3774e rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents: 44412
diff changeset
    51
   "raw_mono_args",
43989
eb763b3ff9ed renamed "preds" encodings to "guards"
blanchet
parents: 43629
diff changeset
    52
   "mono_guards",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    53
   "mono_guards?",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    54
   "mono_guards??",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    55
   "mono_guards!",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    56
   "mono_guards!!",
43208
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    57
   "mono_tags",
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    58
   "mono_tags?",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    59
   "mono_tags??",
43208
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    60
   "mono_tags!",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    61
   "mono_tags!!",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    62
   "mono_args"]
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    63
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
    64
fun metis_exhaust_tac ctxt ths =
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    65
  let
43172
blanchet
parents: 43162
diff changeset
    66
    fun tac [] st = all_tac st
43626
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43212
diff changeset
    67
      | tac (type_enc :: type_encs) st =
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43212
diff changeset
    68
        st (* |> tap (fn _ => tracing (PolyML.makestring type_enc)) *)
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43212
diff changeset
    69
           |> ((if null type_encs then all_tac else rtac @{thm fork} 1)
46369
9ac0c64ad8e7 example tuning
blanchet
parents: 45951
diff changeset
    70
               THEN Metis_Tactic.metis_tac [type_enc]
9ac0c64ad8e7 example tuning
blanchet
parents: 45951
diff changeset
    71
                    ATP_Problem_Generate.combsN ctxt ths 1
43172
blanchet
parents: 43162
diff changeset
    72
               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
    73
               THEN tac type_encs)
46436
blanchet
parents: 46369
diff changeset
    74
  in tac type_encs end
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    75
*}
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    76
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
    77
method_setup metis_exhaust = {*
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    78
  Attrib.thms >>
46436
blanchet
parents: 46369
diff changeset
    79
    (fn ths => fn ctxt => SIMPLE_METHOD (metis_exhaust_tac ctxt ths))
blanchet
parents: 46369
diff changeset
    80
*} "exhaustively run Metis with all type encodings"
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    81
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
    82
text {* Miscellaneous tests *}
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    83
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    84
lemma "x = y \<Longrightarrow> y = x"
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
    85
by metis_exhaust
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    86
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    87
lemma "[a] = [1 + 1] \<Longrightarrow> a = 1 + (1::int)"
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
    88
by (metis_exhaust last.simps)
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    89
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    90
lemma "map Suc [0] = [Suc 0]"
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
    91
by (metis_exhaust map.simps)
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    92
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    93
lemma "map Suc [1 + 1] = [Suc 2]"
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
    94
by (metis_exhaust map.simps nat_1_add_1)
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    95
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    96
lemma "map Suc [2] = [Suc (1 + 1)]"
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
    97
by (metis_exhaust map.simps nat_1_add_1)
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    98
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    99
definition "null xs = (xs = [])"
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
   100
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
   101
lemma "P (null xs) \<Longrightarrow> null xs \<Longrightarrow> xs = []"
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
   102
by (metis_exhaust null_def)
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
   103
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
   104
lemma "(0::nat) + 0 = 0"
45502
6246bef495ff avoid theorem references like 'semiring_norm(111)'
huffman
parents: 44815
diff changeset
   105
by (metis_exhaust add_0_left)
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
   106
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
   107
end