# HG changeset patch # User huffman # Date 1269364059 25200 # Node ID f135ebcc835c80ed35df34c4a343794ac842447e # Parent 86559356502de5f7004c8c6affd666b8418ab0db remove continuous let-binding function CLet; add cont2cont rule ordinary Let diff -r 86559356502d -r f135ebcc835c src/HOLCF/Cfun.thy --- a/src/HOLCF/Cfun.thy Tue Mar 23 09:39:21 2010 -0700 +++ b/src/HOLCF/Cfun.thy Tue Mar 23 10:07:39 2010 -0700 @@ -547,17 +547,26 @@ lemma strictify2 [simp]: "x \ \ \ strictify\f\x = f\x" by (simp add: strictify_conv_if) -subsection {* Continuous let-bindings *} +subsection {* Continuity of let-bindings *} -definition - CLet :: "'a \ ('a \ 'b) \ 'b" where - "CLet = (\ s f. f\s)" +lemma cont2cont_Let: + assumes f: "cont (\x. f x)" + assumes g1: "\y. cont (\x. g x y)" + assumes g2: "\x. cont (\y. g x y)" + shows "cont (\x. let y = f x in g x y)" +unfolding Let_def using f g2 g1 by (rule cont_apply) -syntax - "_CLet" :: "[letbinds, 'a] => 'a" ("(Let (_)/ in (_))" 10) - -translations - "_CLet (_binds b bs) e" == "_CLet b (_CLet bs e)" - "Let x = a in e" == "CONST CLet\a\(\ x. e)" +lemma cont2cont_Let' [cont2cont]: + assumes f: "cont (\x. f x)" + assumes g: "cont (\p. g (fst p) (snd p))" + shows "cont (\x. let y = f x in g x y)" +using f +proof (rule cont2cont_Let) + fix x show "cont (\y. g x y)" + using g by (rule cont_fst_snd_D2) +next + fix y show "cont (\x. g x y)" + using g by (rule cont_fst_snd_D1) +qed end