| author | blanchet | 
| Wed, 12 Feb 2014 08:35:56 +0100 | |
| changeset 55400 | 1e8dd9cd320b | 
| parent 53989 | 729700091556 | 
| child 58889 | 5b7a9633cfa8 | 
| 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 | |
| 53989 | 17 | sledgehammer_params [prover = spass, blocking, fact_filter = mepo, timeout = 30, | 
| 18 | preplay_timeout = 0, dont_minimize] | |
| 19 | ||
| 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 | 20 | 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 | 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 | 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 | 23 | 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 | 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 | 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 | 26 | "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 | 27 | |
| 
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 | 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 | 29 | sledgehammer [expect = some] (inc_def plus_1_not_0) | 
| 43212 | 30 | 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 | 31 | |
| 
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 | lemma "inc = (\<lambda>y. y + 1)" | 
| 47831 | 33 | sledgehammer [expect = some] | 
| 43212 | 34 | 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 | 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 | 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 | 37 | "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 | 38 | |
| 
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 | lemma "add_swap m n = n + m" | 
| 47892 | 40 | sledgehammer [expect = some] (add_swap_def) | 
| 43212 | 41 | 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 | 42 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 43 | 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 | 44 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 45 | 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 | 46 | (* 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 | 47 | 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 | 48 | 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 | 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 | 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 | 51 | 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 | 52 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 53 | 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 | 54 | 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 | 55 | |
| 45970 
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
 haftmann parents: 
44768diff
changeset | 56 | lemma "B \<le> C" | 
| 46733 
4a03b30e04cb
use SPASS instead of E for Metis examples -- it's seems faster for these small problems
 blanchet parents: 
46072diff
changeset | 57 | sledgehammer [expect = some] | 
| 43212 | 58 | 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 | 59 | |
| 
6b7ef9b724fd
added two mildly higher-order examples contributed by TN, removed references to obsoleted type systems, and moved things around
 blanchet parents: 
42722diff
changeset | 60 | 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 | 61 | |
| 42758 
865ce93ce025
handle equality proxy in a more backward-compatible way
 blanchet parents: 
42756diff
changeset | 62 | lemma "id (op =) x x" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 63 | sledgehammer [type_enc = erased, expect = none] (id_apply) | 
| 44768 | 64 | 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 | 65 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 66 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 67 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 68 | 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 | 69 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 70 | 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 | 71 | 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 | 72 | by (metis (full_types) id_apply) | 
| 42758 
865ce93ce025
handle equality proxy in a more backward-compatible way
 blanchet parents: 
42756diff
changeset | 73 | |
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 74 | lemma "id True" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 75 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 76 | 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 | 77 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 78 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 79 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 80 | 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 | 81 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 82 | 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 | 83 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 84 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 85 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 86 | lemma "\<not> id False" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 87 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 88 | 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 | 89 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 90 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 91 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 92 | 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 | 93 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 94 | 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 | 95 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 96 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 97 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 98 | lemma "x = id True \<or> x = id False" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 99 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 100 | 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 | 101 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 102 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 103 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 104 | 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 | 105 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 106 | 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 | 107 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 108 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 109 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 110 | 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 | 111 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 112 | 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 | 113 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 114 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 115 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 116 | 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 | 117 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 118 | 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 | 119 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 120 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 121 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 122 | 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 | 123 | sledgehammer [type_enc = erased, expect = none] () | 
| 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 124 | sledgehammer [type_enc = poly_args, expect = none] () | 
| 44768 | 125 | sledgehammer [type_enc = poly_tags??, expect = some] () | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 126 | sledgehammer [type_enc = poly_tags, expect = some] () | 
| 44768 | 127 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 128 | sledgehammer [type_enc = poly_guards, expect = some] () | 
| 44768 | 129 | sledgehammer [type_enc = mono_tags??, expect = some] () | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 130 | sledgehammer [type_enc = mono_tags, expect = some] () | 
| 44768 | 131 | sledgehammer [type_enc = mono_guards??, expect = some] () | 
| 44494 
a77901b3774e
rationalized option names -- mono becomes raw_mono and mangled becomes mono
 blanchet parents: 
43989diff
changeset | 132 | 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 | 133 | by (metis (full_types)) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 134 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 135 | lemma "id (\<not> a) \<Longrightarrow> \<not> id a" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 136 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 137 | 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 | 138 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 139 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 140 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 141 | 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 | 142 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 143 | 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 | 144 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 145 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 146 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 147 | lemma "id (\<not> \<not> a) \<Longrightarrow> id a" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 148 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 149 | 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 | 150 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 151 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 152 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 153 | 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 | 154 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 155 | 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 | 156 | 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 | 157 | by metis_exhaust | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 158 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 159 | 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 | 160 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 161 | 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 | 162 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 163 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 164 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 165 | 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 | 166 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 167 | 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 | 168 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 169 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 170 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 171 | lemma "id (a \<and> b) \<Longrightarrow> id a" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 172 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 173 | 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 | 174 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 175 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 176 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 177 | 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 | 178 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 179 | 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 | 180 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 181 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 182 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 183 | lemma "id (a \<and> b) \<Longrightarrow> id b" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 184 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 185 | 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 | 186 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 187 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 188 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 189 | 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 | 190 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 191 | 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 | 192 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 193 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 194 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 195 | 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 | 196 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 197 | 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 | 198 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 199 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 200 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 201 | 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 | 202 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 203 | 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 | 204 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 205 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 206 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 207 | lemma "id a \<Longrightarrow> id (a \<or> b)" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 208 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 209 | 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 | 210 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 211 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 212 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 213 | 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 | 214 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 215 | 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 | 216 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 217 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 218 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 219 | lemma "id b \<Longrightarrow> id (a \<or> b)" | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43228diff
changeset | 220 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 221 | 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 | 222 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 223 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 224 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 225 | 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 | 226 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 227 | 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 | 228 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 229 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 230 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 231 | 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 | 232 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 233 | 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 | 234 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 235 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 236 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 237 | 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 | 238 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 239 | 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 | 240 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 241 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 242 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 243 | 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 | 244 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 245 | 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 | 246 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 247 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 248 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 249 | 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 | 250 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 251 | 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 | 252 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 253 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 254 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 255 | 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 | 256 | sledgehammer [type_enc = erased, expect = some] (id_apply) | 
| 44768 | 257 | 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 | 258 | sledgehammer [type_enc = poly_tags, expect = some] (id_apply) | 
| 44768 | 259 | sledgehammer [type_enc = poly_guards??, expect = some] (id_apply) | 
| 43989 | 260 | sledgehammer [type_enc = poly_guards, expect = some] (id_apply) | 
| 44768 | 261 | 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 | 262 | sledgehammer [type_enc = mono_tags, expect = some] (id_apply) | 
| 44768 | 263 | 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 | 264 | sledgehammer [type_enc = mono_guards, expect = some] (id_apply) | 
| 43212 | 265 | by (metis_exhaust id_apply) | 
| 41141 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 266 | |
| 
ad923cdd4a5d
added example to exercise higher-order reasoning with Sledgehammer and Metis
 blanchet parents: diff
changeset | 267 | end |