src/HOL/Metis_Examples/Type_Encodings.thy
author wenzelm
Fri, 19 May 2017 18:10:19 +0200
changeset 65875 12c90c0c4b32
parent 63167 0909deb8059b
child 69989 bcba61d92558
permissions -rw-r--r--
suppress ANSI control sequences in Scala console;
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
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61076
diff changeset
     7
section \<open>
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
     8
Example that Exercises Metis's (and Hence Sledgehammer's) Type Encodings
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61076
diff changeset
     9
\<close>
43197
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
50705
0e943b33d907 use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
blanchet
parents: 48664
diff changeset
    15
declare [[metis_new_skolem]]
43197
c71657bbdbc0 tuned Metis examples
blanchet
parents: 43196
diff changeset
    16
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61076
diff changeset
    17
text \<open>Setup for testing Metis exhaustively\<close>
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    18
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    19
lemma fork: "P \<Longrightarrow> P \<Longrightarrow> P" by assumption
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    20
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61076
diff changeset
    21
ML \<open>
43626
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43212
diff changeset
    22
val type_encs =
43208
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    23
  ["erased",
43989
eb763b3ff9ed renamed "preds" encodings to "guards"
blanchet
parents: 43629
diff changeset
    24
   "poly_guards",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    25
   "poly_guards?",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    26
   "poly_guards??",
48146
14e317033809 reintroduced "t@" encoding, this time sound
blanchet
parents: 48098
diff changeset
    27
   "poly_guards@",
43208
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    28
   "poly_tags",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    29
   "poly_tags?",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    30
   "poly_tags??",
48146
14e317033809 reintroduced "t@" encoding, this time sound
blanchet
parents: 48098
diff changeset
    31
   "poly_tags@",
43208
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    32
   "poly_args",
48092
c84abbf3c5d8 added "args_query" encodings
blanchet
parents: 48091
diff changeset
    33
   "poly_args?",
44494
a77901b3774e rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents: 44412
diff changeset
    34
   "raw_mono_guards",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    35
   "raw_mono_guards?",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    36
   "raw_mono_guards??",
48146
14e317033809 reintroduced "t@" encoding, this time sound
blanchet
parents: 48098
diff changeset
    37
   "raw_mono_guards@",
44494
a77901b3774e rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents: 44412
diff changeset
    38
   "raw_mono_tags",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    39
   "raw_mono_tags?",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    40
   "raw_mono_tags??",
48146
14e317033809 reintroduced "t@" encoding, this time sound
blanchet
parents: 48098
diff changeset
    41
   "raw_mono_tags@",
44494
a77901b3774e rationalized option names -- mono becomes raw_mono and mangled becomes mono
blanchet
parents: 44412
diff changeset
    42
   "raw_mono_args",
48092
c84abbf3c5d8 added "args_query" encodings
blanchet
parents: 48091
diff changeset
    43
   "raw_mono_args?",
43989
eb763b3ff9ed renamed "preds" encodings to "guards"
blanchet
parents: 43629
diff changeset
    44
   "mono_guards",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    45
   "mono_guards?",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    46
   "mono_guards??",
43208
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    47
   "mono_tags",
acfc370df64b compile
blanchet
parents: 43205
diff changeset
    48
   "mono_tags?",
44768
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    49
   "mono_tags??",
a7bc1bdb8bb4 rationalize uniform encodings
blanchet
parents: 44651
diff changeset
    50
   "mono_args"]
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    51
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
    52
fun metis_exhaust_tac ctxt ths =
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    53
  let
43172
blanchet
parents: 43162
diff changeset
    54
    fun tac [] st = all_tac st
43626
a867ebb12209 renamed "type_sys" to "type_enc", which is more accurate
blanchet
parents: 43212
diff changeset
    55
      | tac (type_enc :: type_encs) st =
51929
5e8a0b8bb070 avoid PolyML.makestring, even in dead code;
wenzelm
parents: 50705
diff changeset
    56
        st (* |> tap (fn _ => tracing (@{make_string} type_enc)) *)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 58889
diff changeset
    57
           |> ((if null type_encs then all_tac else resolve_tac ctxt @{thms fork} 1)
46369
9ac0c64ad8e7 example tuning
blanchet
parents: 45951
diff changeset
    58
               THEN Metis_Tactic.metis_tac [type_enc]
9ac0c64ad8e7 example tuning
blanchet
parents: 45951
diff changeset
    59
                    ATP_Problem_Generate.combsN ctxt ths 1
43172
blanchet
parents: 43162
diff changeset
    60
               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
    61
               THEN tac type_encs)
46436
blanchet
parents: 46369
diff changeset
    62
  in tac type_encs end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61076
diff changeset
    63
\<close>
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    64
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61076
diff changeset
    65
method_setup metis_exhaust = \<open>
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    66
  Attrib.thms >>
46436
blanchet
parents: 46369
diff changeset
    67
    (fn ths => fn ctxt => SIMPLE_METHOD (metis_exhaust_tac ctxt ths))
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61076
diff changeset
    68
\<close> "exhaustively run Metis with all type encodings"
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    69
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61076
diff changeset
    70
text \<open>Miscellaneous tests\<close>
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    71
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    72
lemma "x = y \<Longrightarrow> y = x"
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
    73
by metis_exhaust
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    74
48098
dd611ab202a8 tweak Metis example to avoid glitch in proof reconstruction with a few guard-based, type-argument-less encodings
blanchet
parents: 48092
diff changeset
    75
lemma "[a] = [Suc 0] \<Longrightarrow> a = 1"
dd611ab202a8 tweak Metis example to avoid glitch in proof reconstruction with a few guard-based, type-argument-less encodings
blanchet
parents: 48092
diff changeset
    76
by (metis_exhaust last.simps One_nat_def)
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    77
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    78
lemma "map Suc [0] = [Suc 0]"
55465
0d31c0546286 merged 'List.map' and 'List.list.map'
blanchet
parents: 53989
diff changeset
    79
by (metis_exhaust list.map)
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    80
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    81
lemma "map Suc [1 + 1] = [Suc 2]"
55465
0d31c0546286 merged 'List.map' and 'List.list.map'
blanchet
parents: 53989
diff changeset
    82
by (metis_exhaust list.map nat_1_add_1)
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 "map Suc [2] = [Suc (1 + 1)]"
55465
0d31c0546286 merged 'List.map' and 'List.list.map'
blanchet
parents: 53989
diff changeset
    85
by (metis_exhaust list.map nat_1_add_1)
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
definition "null xs = (xs = [])"
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    88
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    89
lemma "P (null xs) \<Longrightarrow> null xs \<Longrightarrow> xs = []"
43212
050a03afe024 Metis code cleanup
blanchet
parents: 43208
diff changeset
    90
by (metis_exhaust null_def)
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    91
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60754
diff changeset
    92
lemma "(0::nat) + 0 = 0"
45502
6246bef495ff avoid theorem references like 'semiring_norm(111)'
huffman
parents: 44815
diff changeset
    93
by (metis_exhaust add_0_left)
43162
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    94
9a8acc5adfa3 added Metis examples to test the new type encodings
blanchet
parents:
diff changeset
    95
end