author | blanchet |
Mon, 23 Aug 2010 14:51:56 +0200 | |
changeset 38651 | 8aadda8e1338 |
parent 38650 | f22a564ac820 |
child 38652 | e063be321438 |
--- a/src/HOL/Predicate.thy Mon Aug 23 14:21:57 2010 +0200 +++ b/src/HOL/Predicate.thy Mon Aug 23 14:51:56 2010 +0200 @@ -92,7 +92,7 @@ lemma top2I [intro!]: "top x y" by (simp add: top_fun_eq top_bool_eq) -lemma bot1E [elim!]: "bot x \<Longrightarrow> P" +lemma bot1E [no_atp, elim!]: "bot x \<Longrightarrow> P" by (simp add: bot_fun_eq bot_bool_eq) lemma bot2E [elim!]: "bot x y \<Longrightarrow> P"