add lemma cont2cont_if_bottom
authorhuffman
Sat, 27 Nov 2010 22:48:08 -0800
changeset 40794 d28d41ee4cef
parent 40777 4898bae6ef23
child 40795 c52cd8bc426d
add lemma cont2cont_if_bottom
src/HOL/HOLCF/Cfun.thy
--- a/src/HOL/HOLCF/Cfun.thy	Sat Nov 27 17:44:36 2010 -0800
+++ b/src/HOL/HOLCF/Cfun.thy	Sat Nov 27 22:48:08 2010 -0800
@@ -481,12 +481,19 @@
   seq :: "'a \<rightarrow> 'b \<rightarrow> 'b" where
   "seq = (\<Lambda> x. if x = \<bottom> then \<bottom> else ID)"
 
-lemma cont_seq: "cont (\<lambda>x. if x = \<bottom> then \<bottom> else y)"
-unfolding cont_def is_lub_def is_ub_def ball_simps
-by (simp add: lub_eq_bottom_iff)
+lemma cont2cont_if_bottom [cont2cont, simp]:
+  assumes f: "cont (\<lambda>x. f x)" and g: "cont (\<lambda>x. g x)"
+  shows "cont (\<lambda>x. if f x = \<bottom> then \<bottom> else g x)"
+proof (rule cont_apply [OF f])
+  show "\<And>x. cont (\<lambda>y. if y = \<bottom> then \<bottom> else g x)"
+    unfolding cont_def is_lub_def is_ub_def ball_simps
+    by (simp add: lub_eq_bottom_iff)
+  show "\<And>y. cont (\<lambda>x. if y = \<bottom> then \<bottom> else g x)"
+    by (simp add: g)
+qed
 
 lemma seq_conv_if: "seq\<cdot>x = (if x = \<bottom> then \<bottom> else ID)"
-unfolding seq_def by (simp add: cont_seq)
+unfolding seq_def by simp
 
 lemma seq1 [simp]: "seq\<cdot>\<bottom> = \<bottom>"
 by (simp add: seq_conv_if)