src/HOL/Metis_Examples/Type_Encodings.thy
author blanchet
Thu, 25 Aug 2011 14:25:07 +0200
changeset 45349 a77901b3774e
parent 45271 c8b847625584
child 45462 b054ca3f07b5
permissions -rw-r--r--
rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet@44038
     1
(*  Title:      HOL/Metis_Examples/Type_Encodings.thy
blanchet@44003
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@44003
     3
blanchet@44038
     4
Example that exercises Metis's (and hence Sledgehammer's) type encodings.
blanchet@44003
     5
*)
blanchet@44003
     6
blanchet@44038
     7
header {*
blanchet@44038
     8
Example that Exercises Metis's (and Hence Sledgehammer's) Type Encodings
blanchet@44038
     9
*}
blanchet@44038
    10
blanchet@44038
    11
theory Type_Encodings
blanchet@45271
    12
imports Main
blanchet@44003
    13
begin
blanchet@44003
    14
blanchet@44038
    15
declare [[metis_new_skolemizer]]
blanchet@44038
    16
blanchet@44038
    17
sledgehammer_params [prover = e, blocking, timeout = 10, preplay_timeout = 0]
blanchet@44038
    18
blanchet@44038
    19
blanchet@44003
    20
text {* Setup for testing Metis exhaustively *}
blanchet@44003
    21
blanchet@44003
    22
lemma fork: "P \<Longrightarrow> P \<Longrightarrow> P" by assumption
blanchet@44003
    23
blanchet@44003
    24
ML {*
blanchet@44493
    25
val type_encs =
blanchet@44049
    26
  ["erased",
blanchet@44860
    27
   "poly_guards",
blanchet@45261
    28
   "poly_guards_uniform",
blanchet@44049
    29
   "poly_tags",
blanchet@45261
    30
   "poly_tags_uniform",
blanchet@44049
    31
   "poly_args",
blanchet@45349
    32
   "raw_mono_guards",
blanchet@45349
    33
   "raw_mono_guards_uniform",
blanchet@45349
    34
   "raw_mono_tags",
blanchet@45349
    35
   "raw_mono_tags_uniform",
blanchet@45349
    36
   "raw_mono_args",
blanchet@44860
    37
   "mono_guards",
blanchet@45261
    38
   "mono_guards_uniform",
blanchet@44049
    39
   "mono_tags",
blanchet@45261
    40
   "mono_tags_uniform",
blanchet@44049
    41
   "mono_args",
blanchet@44049
    42
   "simple",
blanchet@44860
    43
   "poly_guards?",
blanchet@45261
    44
   "poly_guards_uniform?",
blanchet@44496
    45
   "poly_tags?",
blanchet@45271
    46
   "poly_tags_uniform?",
blanchet@45349
    47
   "raw_mono_guards?",
blanchet@45349
    48
   "raw_mono_guards_uniform?",
blanchet@45349
    49
   "raw_mono_tags?",
blanchet@45349
    50
   "raw_mono_tags_uniform?",
blanchet@44860
    51
   "mono_guards?",
blanchet@45261
    52
   "mono_guards_uniform?",
blanchet@44049
    53
   "mono_tags?",
blanchet@45261
    54
   "mono_tags_uniform?",
blanchet@44049
    55
   "simple?",
blanchet@44860
    56
   "poly_guards!",
blanchet@45261
    57
   "poly_guards_uniform!",
blanchet@45271
    58
   "poly_tags!",
blanchet@45261
    59
   "poly_tags_uniform!",
blanchet@45349
    60
   "raw_mono_guards!",
blanchet@45349
    61
   "raw_mono_guards_uniform!",
blanchet@45349
    62
   "raw_mono_tags!",
blanchet@45349
    63
   "raw_mono_tags_uniform!",
blanchet@44860
    64
   "mono_guards!",
blanchet@45261
    65
   "mono_guards_uniform!",
blanchet@44049
    66
   "mono_tags!",
blanchet@45261
    67
   "mono_tags_uniform!",
blanchet@44049
    68
   "simple!"]
blanchet@44003
    69
blanchet@44053
    70
fun metis_exhaust_tac ctxt ths =
blanchet@44003
    71
  let
blanchet@44013
    72
    fun tac [] st = all_tac st
blanchet@44493
    73
      | tac (type_enc :: type_encs) st =
blanchet@44493
    74
        st (* |> tap (fn _ => tracing (PolyML.makestring type_enc)) *)
blanchet@44493
    75
           |> ((if null type_encs then all_tac else rtac @{thm fork} 1)
blanchet@44493
    76
               THEN Metis_Tactics.metis_tac [type_enc] ctxt ths 1
blanchet@44013
    77
               THEN COND (has_fewer_prems 2) all_tac no_tac
blanchet@44493
    78
               THEN tac type_encs)
blanchet@44003
    79
  in tac end
blanchet@44003
    80
*}
blanchet@44003
    81
blanchet@44053
    82
method_setup metis_exhaust = {*
blanchet@44003
    83
  Attrib.thms >>
blanchet@44493
    84
    (fn ths => fn ctxt => SIMPLE_METHOD (metis_exhaust_tac ctxt ths type_encs))
blanchet@44003
    85
*} "exhaustively run the new Metis with all type encodings"
blanchet@44003
    86
blanchet@44003
    87
blanchet@44038
    88
text {* Miscellaneous tests *}
blanchet@44003
    89
blanchet@44003
    90
lemma "x = y \<Longrightarrow> y = x"
blanchet@44053
    91
by metis_exhaust
blanchet@44003
    92
blanchet@44003
    93
lemma "[a] = [1 + 1] \<Longrightarrow> a = 1 + (1::int)"
blanchet@44053
    94
by (metis_exhaust last.simps)
blanchet@44003
    95
blanchet@44003
    96
lemma "map Suc [0] = [Suc 0]"
blanchet@44053
    97
by (metis_exhaust map.simps)
blanchet@44003
    98
blanchet@44003
    99
lemma "map Suc [1 + 1] = [Suc 2]"
blanchet@44053
   100
by (metis_exhaust map.simps nat_1_add_1)
blanchet@44003
   101
blanchet@44003
   102
lemma "map Suc [2] = [Suc (1 + 1)]"
blanchet@44053
   103
by (metis_exhaust map.simps nat_1_add_1)
blanchet@44003
   104
blanchet@44003
   105
definition "null xs = (xs = [])"
blanchet@44003
   106
blanchet@44003
   107
lemma "P (null xs) \<Longrightarrow> null xs \<Longrightarrow> xs = []"
blanchet@44053
   108
by (metis_exhaust null_def)
blanchet@44003
   109
blanchet@44003
   110
lemma "(0::nat) + 0 = 0"
blanchet@44053
   111
by (metis_exhaust arithmetic_simps(38))
blanchet@44003
   112
blanchet@44003
   113
end