| author | haftmann | 
| Sun, 12 Feb 2012 22:10:05 +0100 | |
| changeset 46545 | 69f45ffd5091 | 
| parent 46072 | 291c14a01b6d | 
| child 46733 | 4a03b30e04cb | 
| permissions | -rw-r--r-- | 
| 43197 | 1 | (* Title: HOL/Metis_Examples/Proxies.thy | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 2 | Author: Jasmin Blanchette, TU Muenchen | 
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 3 | |
| 43197 | 4 | Example that exercises Metis's and Sledgehammer's logical symbol proxies for | 
| 5 | rudimentary higher-order reasoning. | |
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 6 | *) | 
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 7 | |
| 43197 | 8 | header {*
 | 
| 9 | Example that Exercises Metis's and Sledgehammer's Logical Symbol Proxies for | |
| 10 | Rudimentary Higher-Order Reasoning. | |
| 11 | *} | |
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 12 | |
| 43197 | 13 | theory Proxies | 
| 14 | imports Type_Encodings | |
| 15 | begin | |
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 16 | |
| 42756 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 17 | text {* Extensionality and set constants *}
 | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 18 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 19 | lemma plus_1_not_0: "n + (1\<Colon>nat) \<noteq> 0" | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 20 | by simp | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 21 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 22 | definition inc :: "nat \<Rightarrow> nat" where | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 23 | "inc x = x + 1" | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 24 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 25 | lemma "inc \<noteq> (\<lambda>y. 0)" | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 26 | sledgehammer [expect = some] (inc_def plus_1_not_0) | 
| 43212 | 27 | by (metis_exhaust inc_def plus_1_not_0) | 
| 42756 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 28 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 29 | lemma "inc = (\<lambda>y. y + 1)" | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 30 | sledgehammer [expect = some] (inc_def) | 
| 43212 | 31 | by (metis_exhaust inc_def) | 
| 42756 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 32 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 33 | definition add_swap :: "nat \<Rightarrow> nat \<Rightarrow> nat" where | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 34 | "add_swap = (\<lambda>x y. y + x)" | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 35 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 36 | lemma "add_swap m n = n + m" | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 37 | sledgehammer [expect = some] (add_swap_def) | 
| 43212 | 38 | by (metis_exhaust add_swap_def) | 
| 42756 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 39 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 40 | definition "A = {xs\<Colon>'a list. True}"
 | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 41 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 42 | lemma "xs \<in> A" | 
| 46072 
291c14a01b6d
reintroduced Sledgehammer call taken out by 9bc924006136, with some hints to guide the naive relevance filter
 blanchet parents: 
46062diff
changeset | 43 | (* The "add:" argument is unfortunate. *) | 
| 
291c14a01b6d
reintroduced Sledgehammer call taken out by 9bc924006136, with some hints to guide the naive relevance filter
 blanchet parents: 
46062diff
changeset | 44 | sledgehammer [expect = some] (add: A_def mem_Collect_eq) | 
| 45970 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 haftmann parents: 
44768diff
changeset | 45 | by (metis_exhaust A_def mem_Collect_eq) | 
| 42756 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 46 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 47 | definition "B (y::int) \<equiv> y \<le> 0" | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 48 | definition "C (y::int) \<equiv> y \<le> 1" | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 49 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 50 | lemma int_le_0_imp_le_1: "x \<le> (0::int) \<Longrightarrow> x \<le> 1" | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 51 | by linarith | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 52 | |
| 45970 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 haftmann parents: 
44768diff
changeset | 53 | lemma "B \<le> C" | 
| 43629 | 54 | sledgehammer [type_enc = poly_args, max_relevant = 100, expect = some] | 
| 43212 | 55 | by (metis_exhaust B_def C_def int_le_0_imp_le_1 predicate1I) | 
| 42756 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 56 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 57 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 58 | text {* Proxies for logical constants *}
 | 
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 59 | |
| 42758 
865ce93ce025
handle equality proxy in a more backward-compatible way
 blanchet parents: 
42756diff
changeset | 60 | lemma "id (op =) x x" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 61 | sledgehammer [type_enc = erased, expect = none] (id_apply) | 
| 44768 | 62 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 63 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 64 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 65 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 66 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 67 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 68 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 69 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43228 
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
 blanchet parents: 
43213diff
changeset | 70 | by (metis (full_types) id_apply) | 
| 42758 
865ce93ce025
handle equality proxy in a more backward-compatible way
 blanchet parents: 
42756diff
changeset | 71 | |
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 72 | lemma "id True" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 73 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 74 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 75 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 76 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 77 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 78 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 79 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 80 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 81 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 82 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 83 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 84 | lemma "\<not> id False" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 85 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 86 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 87 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 88 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 89 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 90 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 91 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 92 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 93 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 94 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 95 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 96 | lemma "x = id True \<or> x = id False" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 97 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 98 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 99 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 100 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 101 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 102 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 103 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 104 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 105 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 106 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 107 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 108 | lemma "id x = id True \<or> id x = id False" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 109 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 110 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 111 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 112 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 113 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 114 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 115 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 116 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 117 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 118 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 119 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 120 | lemma "P True \<Longrightarrow> P False \<Longrightarrow> P x" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 121 | sledgehammer [type_enc = erased, expect = none] () | 
| 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 122 | sledgehammer [type_enc = poly_args, expect = none] () | 
| 44768 | 123 | sledgehammer [type_enc = poly_tags??, expect = some] () | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 124 | sledgehammer [type_enc = poly_tags, expect = some] () | 
| 44768 | 125 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 126 | sledgehammer [type_enc = poly_guards, expect = some] () | 
| 44768 | 127 | sledgehammer [type_enc = mono_tags??, expect = some] () | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 128 | sledgehammer [type_enc = mono_tags, expect = some] () | 
| 44768 | 129 | sledgehammer [type_enc = mono_guards??, expect = some] () | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 130 | sledgehammer [type_enc = mono_guards, expect = some] () | 
| 43228 
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
 blanchet parents: 
43213diff
changeset | 131 | by (metis (full_types)) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 132 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 133 | lemma "id (\<not> a) \<Longrightarrow> \<not> id a" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 134 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 135 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 136 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 137 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 138 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 139 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 140 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 141 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 142 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 143 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 144 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 145 | lemma "id (\<not> \<not> a) \<Longrightarrow> id a" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 146 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 147 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 148 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 149 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 150 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 151 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 152 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 153 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 154 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 45970 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 haftmann parents: 
44768diff
changeset | 155 | by metis_exhaust | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 156 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 157 | lemma "id (\<not> (id (\<not> a))) \<Longrightarrow> id a" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 158 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 159 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 160 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 161 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 162 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 163 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 164 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 165 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 166 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 167 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 168 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 169 | lemma "id (a \<and> b) \<Longrightarrow> id a" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 170 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 171 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 172 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 173 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 174 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 175 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 176 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 177 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 178 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 179 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 180 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 181 | lemma "id (a \<and> b) \<Longrightarrow> id b" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 182 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 183 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 184 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 185 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 186 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 187 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 188 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 189 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 190 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 191 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 192 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 193 | lemma "id a \<Longrightarrow> id b \<Longrightarrow> id (a \<and> b)" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 194 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 195 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 196 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 197 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 198 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 199 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 200 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 201 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 202 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 203 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 204 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 205 | lemma "id a \<Longrightarrow> id (a \<or> b)" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 206 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 207 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 208 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 209 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 210 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 211 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 212 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 213 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 214 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 215 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 216 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 217 | lemma "id b \<Longrightarrow> id (a \<or> b)" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 218 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 219 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 220 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 221 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 222 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 223 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 224 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 225 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 226 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 227 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 228 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 229 | lemma "id (\<not> a) \<Longrightarrow> id (\<not> b) \<Longrightarrow> id (\<not> (a \<or> b))" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 230 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 231 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 232 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 233 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 234 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 235 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 236 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 237 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 238 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 239 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 240 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 241 | lemma "id (\<not> a) \<Longrightarrow> id (a \<longrightarrow> b)" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 242 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 243 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 244 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 245 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 246 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 247 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 248 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 249 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 250 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 251 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 252 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 253 | lemma "id (a \<longrightarrow> b) \<longleftrightarrow> id (\<not> a \<or> b)" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 254 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 255 | sledgehammer [type_enc = poly_tags??, expect = some] (id_apply) | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 256 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 257 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 258 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 259 | sledgehammer [type_enc = mono_tags??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 260 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 261 | sledgehammer [type_enc = mono_guards??, expect = some] (id_apply) | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 262 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 263 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 264 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 265 | end |