# HG changeset patch # User huffman # Date 1290926888 28800 # Node ID d28d41ee4cef2b6974a8ccd7957d1dc3800313ef # Parent 4898bae6ef23aae7ffe7609f29ff04fa6d962b9a add lemma cont2cont_if_bottom diff -r 4898bae6ef23 -r d28d41ee4cef 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 \ 'b \ 'b" where "seq = (\ x. if x = \ then \ else ID)" -lemma cont_seq: "cont (\x. if x = \ then \ 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 (\x. f x)" and g: "cont (\x. g x)" + shows "cont (\x. if f x = \ then \ else g x)" +proof (rule cont_apply [OF f]) + show "\x. cont (\y. if y = \ then \ else g x)" + unfolding cont_def is_lub_def is_ub_def ball_simps + by (simp add: lub_eq_bottom_iff) + show "\y. cont (\x. if y = \ then \ else g x)" + by (simp add: g) +qed lemma seq_conv_if: "seq\x = (if x = \ then \ else ID)" -unfolding seq_def by (simp add: cont_seq) +unfolding seq_def by simp lemma seq1 [simp]: "seq\\ = \" by (simp add: seq_conv_if)