src/HOLCF/Up.thy
changeset 35783 38538bfe9ca6
parent 35427 ad039d29e01c
child 35900 aa5dfb03eb1e
--- a/src/HOLCF/Up.thy	Sat Mar 13 21:07:20 2010 -0800
+++ b/src/HOLCF/Up.thy	Sat Mar 13 22:00:34 2010 -0800
@@ -237,13 +237,15 @@
 lemma up_below [simp]: "up\<cdot>x \<sqsubseteq> up\<cdot>y \<longleftrightarrow> x \<sqsubseteq> y"
 by (simp add: up_def cont_Iup)
 
-lemma upE [cases type: u]: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; \<And>x. p = up\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
+lemma upE [case_names bottom up, cases type: u]:
+  "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; \<And>x. p = up\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
 apply (cases p)
 apply (simp add: inst_up_pcpo)
 apply (simp add: up_def cont_Iup)
 done
 
-lemma up_induct [induct type: u]: "\<lbrakk>P \<bottom>; \<And>x. P (up\<cdot>x)\<rbrakk> \<Longrightarrow> P x"
+lemma up_induct [case_names bottom up, induct type: u]:
+  "\<lbrakk>P \<bottom>; \<And>x. P (up\<cdot>x)\<rbrakk> \<Longrightarrow> P x"
 by (cases x, simp_all)
 
 text {* lifting preserves chain-finiteness *}