# HG changeset patch # User huffman # Date 1117833835 -7200 # Node ID 36ee7f6af79fef7c6b3ea0f2532f77bf6bbdd473 # Parent cfe047ad6384e1ee3333b72edc40ad234812a92d removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext diff -r cfe047ad6384 -r 36ee7f6af79f src/HOLCF/Cfun.ML --- a/src/HOLCF/Cfun.ML Fri Jun 03 23:16:35 2005 +0200 +++ b/src/HOLCF/Cfun.ML Fri Jun 03 23:23:55 2005 +0200 @@ -2,21 +2,12 @@ (* legacy ML bindings *) val less_cfun_def = thm "less_cfun_def"; -val Rep_Cfun = thm "Rep_Cfun"; -val Rep_Cfun_inverse = thm "Rep_Cfun_inverse"; -val Abs_Cfun_inverse = thm "Abs_Cfun_inverse"; val cfun_cong = thm "cfun_cong"; val cfun_fun_cong = thm "cfun_fun_cong"; val cfun_arg_cong = thm "cfun_arg_cong"; -val Abs_Cfun_inverse2 = thm "Abs_Cfun_inverse2"; -val Cfunapp2 = thm "Cfunapp2"; +val Abs_CFun_inverse2 = thm "Abs_CFun_inverse2"; val beta_cfun = thm "beta_cfun"; val eta_cfun = thm "eta_cfun"; -val inst_cfun_po = thm "inst_cfun_po"; -val less_cfun = thm "less_cfun"; -val minimal_cfun = thm "minimal_cfun"; -val UU_cfun_def = thm "UU_cfun_def"; -val least_cfun = thm "least_cfun"; val cont_Rep_CFun2 = thm "cont_Rep_CFun2"; val monofun_Rep_CFun2 = thm "monofun_Rep_CFun2"; val contlub_Rep_CFun2 = thm "contlub_Rep_CFun2"; @@ -31,13 +22,13 @@ val ch2ch_Rep_CFunR = thm "ch2ch_Rep_CFunR"; val ch2ch_Rep_CFunL = thm "ch2ch_Rep_CFunL"; val lub_cfun_mono = thm "lub_cfun_mono"; -val ex_lubcfun = thm "ex_lubcfun"; -val cont_lubcfun = thm "cont_lubcfun"; +val ex_lub_cfun = thm "ex_lub_cfun"; +val cont_lub_cfun = thm "cont_lub_cfun"; val lub_cfun = thm "lub_cfun"; val thelub_cfun = thm "thelub_cfun"; val ext_cfun = thm "ext_cfun"; val semi_monofun_Abs_CFun = thm "semi_monofun_Abs_CFun"; -val less_cfun2 = thm "less_cfun2"; +val less_cfun_ext = thm "less_cfun_ext"; val Istrictify_def = thm "Istrictify_def"; val strictify_def = thm "strictify_def"; val ID_def = thm "ID_def"; @@ -52,20 +43,17 @@ val cont2cont_Rep_CFun = thm "cont2cont_Rep_CFun"; val cont2mono_LAM = thm "cont2mono_LAM"; val cont2cont_LAM = thm "cont2cont_LAM"; -val cont2cont_eta = thm "cont2cont_eta"; val cont_lemmas1 = [cont_const, cont_id, cont_Rep_CFun2, cont2cont_Rep_CFun, cont2cont_LAM]; -val strict_Rep_CFun1 = thm "strict_Rep_CFun1"; val Istrictify1 = thm "Istrictify1"; val Istrictify2 = thm "Istrictify2"; -val monofun_Istrictify1 = thm "monofun_Istrictify1"; val monofun_Istrictify2 = thm "monofun_Istrictify2"; -val contlub_Istrictify1 = thm "contlub_Istrictify1"; val contlub_Istrictify2 = thm "contlub_Istrictify2"; val cont_Istrictify1 = thm "cont_Istrictify1"; val cont_Istrictify2 = thm "cont_Istrictify2"; val strictify1 = thm "strictify1"; val strictify2 = thm "strictify2"; +val Rep_CFun_strict1 = thm "Rep_CFun_strict1"; val chfin_Rep_CFunR = thm "chfin_Rep_CFunR"; val retraction_strict = thm "retraction_strict"; val injection_eq = thm "injection_eq"; diff -r cfe047ad6384 -r 36ee7f6af79f src/HOLCF/Cfun.thy --- a/src/HOLCF/Cfun.thy Fri Jun 03 23:16:35 2005 +0200 +++ b/src/HOLCF/Cfun.thy Fri Jun 03 23:23:55 2005 +0200 @@ -16,69 +16,29 @@ subsection {* Definition of continuous function type *} typedef (CFun) ('a, 'b) "->" (infixr 0) = "{f::'a => 'b. cont f}" -by (rule exI, rule CfunI) +by (rule exI, fast intro: cont_const) syntax - Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("_$_" [999,1000] 999) - (* application *) - Abs_CFun :: "('a => 'b) => ('a -> 'b)" (binder "LAM " 10) - (* abstraction *) - less_cfun :: "[('a -> 'b),('a -> 'b)]=>bool" + Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("_$_" [999,1000] 999) + (* application *) + Abs_CFun :: "('a => 'b) => ('a -> 'b)" (binder "LAM " 10) + (* abstraction *) syntax (xsymbols) - "->" :: "[type, type] => type" ("(_ \/ _)" [1,0]0) - "LAM " :: "[idts, 'a => 'b] => ('a -> 'b)" + "->" :: "[type, type] => type" ("(_ \/ _)" [1,0]0) + "LAM " :: "[idts, 'a => 'b] => ('a -> 'b)" ("(3\_./ _)" [0, 10] 10) - Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\_)" [999,1000] 999) + Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\_)" [999,1000] 999) syntax (HTML output) - Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\_)" [999,1000] 999) - -text {* - Derive old type definition rules for @{term Abs_CFun} \& @{term Rep_CFun}. - @{term Rep_CFun} and @{term Abs_CFun} should be replaced by - @{term Rep_Cfun} and @{term Abs_Cfun} in future. -*} - -lemma Rep_Cfun: "Rep_CFun fo : CFun" -by (rule Rep_CFun) - -lemma Rep_Cfun_inverse: "Abs_CFun (Rep_CFun fo) = fo" -by (rule Rep_CFun_inverse) - -lemma Abs_Cfun_inverse: "f:CFun==>Rep_CFun(Abs_CFun f)=f" -by (erule Abs_CFun_inverse) - -text {* Additional lemma about the isomorphism between - @{typ "'a -> 'b"} and @{term Cfun} *} - -lemma Abs_Cfun_inverse2: "cont f ==> Rep_CFun (Abs_CFun f) = f" -apply (rule Abs_Cfun_inverse) -apply (unfold CFun_def) -apply (erule mem_Collect_eq [THEN ssubst]) -done - -text {* Simplification of application *} - -lemma Cfunapp2: "cont f ==> (Abs_CFun f)$x = f x" -by (erule Abs_Cfun_inverse2 [THEN fun_cong]) - -text {* Beta - equality for continuous functions *} - -lemma beta_cfun: "cont(c1) ==> (LAM x .c1 x)$u = c1 u" -by (rule Cfunapp2) - -text {* Eta - equality for continuous functions *} - -lemma eta_cfun: "(LAM x. f$x) = f" -by (rule Rep_CFun_inverse) + Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\_)" [999,1000] 999) subsection {* Class instances *} instance "->" :: (cpo, cpo) sq_ord .. defs (overloaded) - less_cfun_def: "(op <<) == (% fo1 fo2. Rep_CFun fo1 << Rep_CFun fo2 )" + less_cfun_def: "(op \) \ (\f g. Rep_CFun f \ Rep_CFun g)" lemma adm_CFun: "adm (\f. f \ CFun)" by (simp add: CFun_def, rule admI, rule cont_lub_fun) @@ -101,275 +61,190 @@ lemmas cont_Abs_CFun = typedef_cont_Abs [OF type_definition_CFun less_cfun_def adm_CFun] -lemmas strict_Rep_CFun = - typedef_strict_Rep [OF type_definition_CFun less_cfun_def UU_CFun] +lemmas Rep_CFun_strict = + typedef_Rep_strict [OF type_definition_CFun less_cfun_def UU_CFun] + +lemmas Abs_CFun_strict = + typedef_Abs_strict [OF type_definition_CFun less_cfun_def UU_CFun] -lemmas strict_Abs_CFun = - typedef_strict_Abs [OF type_definition_CFun less_cfun_def UU_CFun] +text {* Additional lemma about the isomorphism between + @{typ "'a -> 'b"} and @{term CFun} *} + +lemma Abs_CFun_inverse2: "cont f \ Rep_CFun (Abs_CFun f) = f" +by (simp add: Abs_CFun_inverse CFun_def) -text {* for compatibility with old HOLCF-Version *} -lemma inst_cfun_po: "(op <<)=(%f1 f2. Rep_CFun f1 << Rep_CFun f2)" -apply (fold less_cfun_def) -apply (rule refl) -done +text {* Beta-equality for continuous functions *} + +lemma beta_cfun [simp]: "cont f \ (\ x. f x)\u = f u" +by (simp add: Abs_CFun_inverse2) + +text {* Eta-equality for continuous functions *} + +lemma eta_cfun: "(\ x. f\x) = f" +by (rule Rep_CFun_inverse) + +text {* Extensionality for continuous functions *} + +lemma ext_cfun: "(\x. f\x = g\x) \ f = g" +by (simp add: Rep_CFun_inject [symmetric] ext) text {* lemmas about application of continuous functions *} -lemma cfun_cong: "[| f=g; x=y |] ==> f$x = g$y" +lemma cfun_cong: "\f = g; x = y\ \ f\x = g\y" by simp -lemma cfun_fun_cong: "f=g ==> f$x = g$x" +lemma cfun_fun_cong: "f = g \ f\x = g\x" by simp -lemma cfun_arg_cong: "x=y ==> f$x = f$y" +lemma cfun_arg_cong: "x = y \ f\x = f\y" by simp -text {* access to @{term less_cfun} in class po *} +subsection {* Continuity of application *} -lemma less_cfun: "( f1 << f2 ) = (Rep_CFun(f1) << Rep_CFun(f2))" -by (simp add: inst_cfun_po) - -subsection {* Type @{typ "'a -> 'b"} is pointed *} +lemma cont_Rep_CFun1: "cont (\f. f\x)" +by (rule cont_Rep_CFun [THEN cont2cont_CF1L]) -lemma minimal_cfun: "Abs_CFun(% x. UU) << f" -apply (subst less_cfun) -apply (subst Abs_Cfun_inverse2) -apply (rule cont_const) -apply (rule minimal_fun) -done - -lemmas UU_cfun_def = minimal_cfun [THEN minimal2UU, symmetric, standard] - -lemma least_cfun: "? x::'a->'b::pcpo.!y. x<x. f\x)" +apply (rule_tac P = "cont" in CollectD) +apply (fold CFun_def) +apply (rule Rep_CFun) done -subsection {* Monotonicity of application *} +lemmas monofun_Rep_CFun = cont_Rep_CFun [THEN cont2mono] +lemmas contlub_Rep_CFun = cont_Rep_CFun [THEN cont2contlub] -text {* - @{term Rep_CFun} yields continuous functions in @{typ "'a => 'b"}. - This is continuity of @{term Rep_CFun} in its 'second' argument: - @{prop "cont_Rep_CFun2 ==> monofun_Rep_CFun2 & contlub_Rep_CFun2"} -*} +lemmas monofun_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2mono, standard] +lemmas contlub_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2contlub, standard] +lemmas monofun_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2mono, standard] +lemmas contlub_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2contlub, standard] + +text {* contlub, cont properties of @{term Rep_CFun} in each argument *} + +lemma contlub_cfun_arg: "chain Y \ f\(lub (range Y)) = (\i. f\(Y i))" +by (rule contlub_Rep_CFun2 [THEN contlubE]) -lemma cont_Rep_CFun2: "cont (Rep_CFun fo)" -apply (rule_tac P = "cont" in CollectD) -apply (fold CFun_def) -apply (rule Rep_Cfun) -done +lemma cont_cfun_arg: "chain Y \ range (\i. f\(Y i)) <<| f\(lub (range Y))" +by (rule cont_Rep_CFun2 [THEN contE]) + +lemma contlub_cfun_fun: "chain F \ lub (range F)\x = (\i. F i\x)" +by (rule contlub_Rep_CFun1 [THEN contlubE]) -lemmas monofun_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2mono, standard] - -- {* @{thm monofun_Rep_CFun2} *} (* monofun(Rep_CFun(?fo)) *) +lemma cont_cfun_fun: "chain F \ range (\i. F i\x) <<| lub (range F)\x" +by (rule cont_Rep_CFun1 [THEN contE]) -lemmas contlub_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2contlub, standard] - -- {* @{thm contlub_Rep_CFun2} *} (* contlub(Rep_CFun(?fo)) *) +text {* Extensionality wrt. @{term "op <<"} in @{typ "'a -> 'b"} *} -text {* expanded thms @{thm [source] cont_Rep_CFun2}, @{thm [source] contlub_Rep_CFun2} look nice with mixfix syntax *} +lemma less_cfun_ext: "(\x. f\x \ g\x) \ f \ g" +by (simp add: less_cfun_def less_fun_def) -lemmas cont_cfun_arg = cont_Rep_CFun2 [THEN contE, THEN spec, THEN mp] - -- {* @{thm cont_cfun_arg} *} (* chain(x1) ==> range (%i. fo3$(x1 i)) <<| fo3$(lub (range ?x1)) *) - -lemmas contlub_cfun_arg = contlub_Rep_CFun2 [THEN contlubE, THEN spec, THEN mp] - -- {* @{thm contlub_cfun_arg} *} (* chain(?x1) ==> ?fo4$(lub (range ?x1)) = lub (range (%i. ?fo4$(?x1 i))) *) +text {* monotonicity of application *} + +lemma monofun_cfun_fun: "f \ g \ f\x \ g\x" +by (simp add: less_cfun_def less_fun_def) -text {* @{term Rep_CFun} is monotone in its 'first' argument *} +lemma monofun_cfun_arg: "x \ y \ f\x \ f\y" +by (rule monofun_Rep_CFun2 [THEN monofunE]) -lemma monofun_Rep_CFun1: "monofun(Rep_CFun)" -by (rule cont_Rep_CFun [THEN cont2mono]) +lemma monofun_cfun: "\f \ g; x \ y\ \ f\x \ g\y" +by (rule trans_less [OF monofun_cfun_fun monofun_cfun_arg]) -text {* monotonicity of application @{term Rep_CFun} in mixfix syntax @{text "[_]_"} *} +text {* ch2ch - rules for the type @{typ "'a -> 'b"} *} -lemma monofun_cfun_fun: "f1 << f2 ==> f1$x << f2$x" -apply (rule_tac x = "x" in spec) -apply (rule less_fun [THEN subst]) -apply (erule monofun_Rep_CFun1 [THEN monofunE [rule_format]]) -done +lemma chain_monofun: "chain Y \ chain (\i. f\(Y i))" +by (erule monofun_Rep_CFun2 [THEN ch2ch_monofun]) + +lemma ch2ch_Rep_CFunR: "chain Y \ chain (\i. f\(Y i))" +by (rule monofun_Rep_CFun2 [THEN ch2ch_monofun]) -lemmas monofun_cfun_arg = monofun_Rep_CFun2 [THEN monofunE [rule_format], standard] - -- {* @{thm monofun_cfun_arg} *} (* ?x2 << ?x1 ==> ?fo5$?x2 << ?fo5$?x1 *) +lemma ch2ch_Rep_CFunL: "chain F \ chain (\i. (F i)\x)" +by (rule monofun_Rep_CFun1 [THEN ch2ch_monofun]) -lemma chain_monofun: "chain Y ==> chain (%i. f\(Y i))" +lemma ch2ch_Rep_CFun: "\chain F; chain Y\ \ chain (\i. (F i)\(Y i))" apply (rule chainI) -apply (rule monofun_cfun_arg) +apply (rule monofun_cfun) +apply (erule chainE) apply (erule chainE) done -text {* monotonicity of @{term Rep_CFun} in both arguments in mixfix syntax @{text "[_]_"} *} +text {* contlub, cont properties of @{term Rep_CFun} in both arguments *} -lemma monofun_cfun: "[|f1< f1$x1 << f2$x2" -apply (rule trans_less) -apply (erule monofun_cfun_arg) -apply (erule monofun_cfun_fun) +lemma contlub_cfun: + "\chain F; chain Y\ \ (\i. F i)\(\i. Y i) = (\i. F i\(Y i))" +apply (simp only: contlub_cfun_fun) +apply (simp only: contlub_cfun_arg) +apply (rule diag_lub) +apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun]) +apply (erule monofun_Rep_CFun2 [THEN ch2ch_monofun]) done -lemma strictI: "f$x = UU ==> f$UU = UU" -apply (rule eq_UU_iff [THEN iffD2]) +lemma cont_cfun: + "\chain F; chain Y\ \ range (\i. F i\(Y i)) <<| (\i. F i)\(\i. Y i)" +apply (rule thelubE) +apply (simp only: ch2ch_Rep_CFun) +apply (simp only: contlub_cfun) +done + +text {* strictness *} + +lemma strictI: "f\x = \ \ f\\ = \" +apply (rule UU_I) apply (erule subst) apply (rule minimal [THEN monofun_cfun_arg]) done -subsection {* Type @{typ "'a -> 'b"} is a cpo *} - -text {* ch2ch - rules for the type @{typ "'a -> 'b"} use MF2 lemmas from Cont.thy *} - -lemma ch2ch_Rep_CFunR: "chain(Y) ==> chain(%i. f$(Y i))" -by (erule monofun_Rep_CFun2 [THEN ch2ch_MF2R]) +text {* the lub of a chain of continous functions is monotone *} -lemmas ch2ch_Rep_CFunL = monofun_Rep_CFun1 [THEN ch2ch_MF2L, standard] - -- {* @{thm ch2ch_Rep_CFunL} *} (* chain(?F) ==> chain (%i. ?F i$?x) *) - -text {* the lub of a chain of continous functions is monotone: uses MF2 lemmas from Cont.thy *} - -lemma lub_cfun_mono: "chain(F) ==> monofun(% x. lub(range(% j.(F j)$x)))" -apply (rule lub_MF2_mono) -apply (rule monofun_Rep_CFun1) -apply (rule monofun_Rep_CFun2 [THEN allI]) -apply assumption +lemma lub_cfun_mono: "chain F \ monofun (\x. \i. F i\x)" +apply (drule ch2ch_monofun [OF monofun_Rep_CFun]) +apply (simp add: thelub_fun [symmetric]) +apply (erule monofun_lub_fun) +apply (simp add: monofun_Rep_CFun2) done text {* a lemma about the exchange of lubs for type @{typ "'a -> 'b"}: uses MF2 lemmas from Cont.thy *} -lemma ex_lubcfun: "[| chain(F); chain(Y) |] ==> +lemma ex_lub_cfun: "[| chain(F); chain(Y) |] ==> lub(range(%j. lub(range(%i. F(j)$(Y i))))) = lub(range(%i. lub(range(%j. F(j)$(Y i)))))" -apply (rule ex_lubMF2) -apply (rule monofun_Rep_CFun1) -apply (rule monofun_Rep_CFun2 [THEN allI]) -apply assumption -apply assumption -done +by (simp add: diag_lub ch2ch_Rep_CFunL ch2ch_Rep_CFunR) text {* the lub of a chain of cont. functions is continuous *} -lemma cont_lubcfun: "chain(F) ==> cont(% x. lub(range(% j. F(j)$x)))" -apply (rule monocontlub2cont) -apply (erule lub_cfun_mono) -apply (rule contlubI [rule_format]) -apply (subst contlub_cfun_arg [THEN ext]) -apply assumption -apply (erule ex_lubcfun) -apply assumption +lemma cont_lub_cfun: "chain F \ cont (\x. \i. F i\x)" +apply (rule cont2cont_lub) +apply (erule monofun_Rep_CFun [THEN ch2ch_monofun]) +apply (rule cont_Rep_CFun2) done text {* type @{typ "'a -> 'b"} is chain complete *} -lemma lub_cfun: "chain(CCF) ==> range(CCF) <<| (LAM x. lub(range(% i. CCF(i)$x)))" -apply (rule is_lubI) -apply (rule ub_rangeI) -apply (subst less_cfun) -apply (subst Abs_Cfun_inverse2) -apply (erule cont_lubcfun) -apply (rule lub_fun [THEN is_lubD1, THEN ub_rangeD]) -apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun]) -apply (subst less_cfun) -apply (subst Abs_Cfun_inverse2) -apply (erule cont_lubcfun) -apply (rule lub_fun [THEN is_lub_lub]) -apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun]) -apply (erule monofun_Rep_CFun1 [THEN ub2ub_monofun]) +lemma lub_cfun: "chain F \ range F <<| (LAM x. LUB i. F i$x)" +apply (subst thelub_fun [symmetric]) +apply (erule monofun_Rep_CFun [THEN ch2ch_monofun]) +apply (erule typedef_is_lub [OF type_definition_CFun less_cfun_def adm_CFun]) done lemmas thelub_cfun = lub_cfun [THEN thelubI, standard] - -- {* @{thm thelub_cfun} *} (* -chain(?CCF1) ==> lub (range ?CCF1) = (LAM x. lub (range (%i. ?CCF1 i$x))) -*) + -- {* @{thm thelub_cfun} *} (* chain F \ lub (range F) = (\x. \i. F i\x) *) subsection {* Miscellaneous *} -text {* Extensionality in @{typ "'a -> 'b"} *} - -lemma ext_cfun: "(!!x. f$x = g$x) ==> f = g" -apply (rule Rep_CFun_inject [THEN iffD1]) -apply (rule ext) -apply simp -done - text {* Monotonicity of @{term Abs_CFun} *} lemma semi_monofun_Abs_CFun: "[| cont(f); cont(g); f< Abs_CFun(f)< 'b"} *} - -lemma less_cfun2: "(!!x. f$x << g$x) ==> f << g" -apply (rule_tac t = "f" in Rep_Cfun_inverse [THEN subst]) -apply (rule_tac t = "g" in Rep_Cfun_inverse [THEN subst]) -apply (rule semi_monofun_Abs_CFun) -apply (rule cont_Rep_CFun2) -apply (rule cont_Rep_CFun2) -apply (rule less_fun [THEN iffD2]) -apply simp -done +by (simp add: less_cfun_def Abs_CFun_inverse2) text {* for compatibility with old HOLCF-Version *} -lemma inst_cfun_pcpo: "UU = Abs_CFun(%x. UU)" -apply (simp add: UU_def UU_cfun_def) -done +lemma inst_cfun_pcpo: "\ = (\ x. \)" +by (simp add: inst_fun_pcpo [symmetric] Abs_CFun_strict) subsection {* Continuity of application *} -text {* the contlub property for @{term Rep_CFun} its 'first' argument *} - -lemma contlub_Rep_CFun1: "contlub(Rep_CFun)" -by (rule cont_Rep_CFun [THEN cont2contlub]) - -text {* the cont property for @{term Rep_CFun} in its first argument *} - -lemma cont_Rep_CFun1: "cont(Rep_CFun)" -by (rule cont_Rep_CFun) - -text {* contlub, cont properties of @{term Rep_CFun} in its first argument in mixfix @{text "_[_]"} *} - -lemma contlub_cfun_fun: -"chain(FY) ==> - lub(range FY)$x = lub(range (%i. FY(i)$x))" -apply (rule trans) -apply (erule contlub_Rep_CFun1 [THEN contlubE, THEN spec, THEN mp, THEN fun_cong]) -apply (subst thelub_fun) -apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun]) -apply (rule refl) -done - -lemma cont_cfun_fun: -"chain(FY) ==> - range(%i. FY(i)$x) <<| lub(range FY)$x" -apply (rule thelubE) -apply (erule ch2ch_Rep_CFunL) -apply (erule contlub_cfun_fun [symmetric]) -done - -text {* contlub, cont properties of @{term Rep_CFun} in both argument in mixfix @{text "_[_]"} *} - -lemma contlub_cfun: -"[|chain(FY);chain(TY)|] ==> - (lub(range FY))$(lub(range TY)) = lub(range(%i. FY(i)$(TY i)))" -apply (rule contlub_CF2) -apply (rule cont_Rep_CFun1) -apply (rule allI) -apply (rule cont_Rep_CFun2) -apply assumption -apply assumption -done - -lemma cont_cfun: -"[|chain(FY);chain(TY)|] ==> - range(%i.(FY i)$(TY i)) <<| (lub (range FY))$(lub(range TY))" -apply (rule thelubE) -apply (rule monofun_Rep_CFun1 [THEN ch2ch_MF2LR]) -apply (rule allI) -apply (rule monofun_Rep_CFun2) -apply assumption -apply assumption -apply (erule contlub_cfun [symmetric]) -apply assumption -done - text {* cont2cont lemma for @{term Rep_CFun} *} -lemma cont2cont_Rep_CFun: "[|cont(%x. ft x);cont(%x. tt x)|] ==> cont(%x. (ft x)$(tt x))" -apply (best intro: cont2cont_app2 cont_const cont_Rep_CFun1 cont_Rep_CFun2) -done +lemma cont2cont_Rep_CFun: + "\cont f; cont t\ \ cont (\x. (f x)\(t x))" +by (best intro: cont2cont_app2 cont_const cont_Rep_CFun cont_Rep_CFun2) text {* cont2mono Lemma for @{term "%x. LAM y. c1(x)(y)"} *} @@ -377,10 +252,10 @@ assumes p1: "!!x. cont(c1 x)" assumes p2: "!!y. monofun(%x. c1 x y)" shows "monofun(%x. LAM y. c1 x y)" -apply (rule monofunI [rule_format]) -apply (rule less_cfun2) -apply (simp add: beta_cfun p1) -apply (erule p2 [THEN monofunE [rule_format]]) +apply (rule monofunI) +apply (rule less_cfun_ext) +apply (simp add: p1) +apply (erule p2 [THEN monofunE]) done text {* cont2cont Lemma for @{term "%x. LAM y. c1 x y"} *} @@ -394,11 +269,6 @@ apply (simp add: p2 cont2cont_CF1L_rev) done -text {* cont2cont Lemma for @{term "%x. LAM y. c1 x$y"} *} - -lemma cont2cont_eta: "cont c1 ==> cont (%x. LAM y. c1 x$y)" -by (simp only: eta_cfun) - text {* cont2cont tactic *} lemmas cont_lemmas1 = @@ -427,19 +297,13 @@ Addsimprocs [cont_proc]; *} -text {* HINT: @{text cont_tac} is now installed in simplifier in Lift.ML ! *} - (*val cont_tac = (fn i => (resolve_tac cont_lemmas i));*) (*val cont_tacR = (fn i => (REPEAT (cont_tac i)));*) -text {* function application @{text "_[_]"} is strict in its first arguments *} +text {* function application is strict in its first argument *} -lemma strict_Rep_CFun1 [simp]: "\\x = \" -by (simp add: inst_cfun_pcpo beta_cfun) - -text {* Instantiate the simplifier *} - -declare beta_cfun [simp] +lemma Rep_CFun_strict1 [simp]: "\\x = \" +by (simp add: Rep_CFun_strict) text {* some lemmata for functions with flat/chfin domain/range types *} @@ -498,7 +362,7 @@ "\y. (f::'a::flat \ 'b::pcpo)\(g\y) = y \ \x y::'b. x \ y \ x = \ \ x = y" apply clarify -apply (drule_tac fo=g in monofun_cfun_arg) +apply (drule_tac f=g in monofun_cfun_arg) apply (drule ax_flat [rule_format]) apply (erule disjE) apply (simp add: injection_defined_rev) @@ -585,32 +449,23 @@ lemma Istrictify2: "x \ \ \ Istrictify f x = f\x" by (simp add: Istrictify_def) -lemma monofun_Istrictify1: "monofun (\f. Istrictify f x)" -apply (rule monofunI [rule_format]) -apply (simp add: Istrictify_def monofun_cfun_fun) +lemma cont_Istrictify1: "cont (\f. Istrictify f x)" +apply (case_tac "x = \") +apply (simp add: Istrictify1) +apply (simp add: Istrictify2) done lemma monofun_Istrictify2: "monofun (\x. Istrictify f x)" -apply (rule monofunI [rule_format]) +apply (rule monofunI) apply (simp add: Istrictify_def monofun_cfun_arg) apply clarify apply (simp add: eq_UU_iff) done -lemma contlub_Istrictify1: "contlub (\f. Istrictify f x)" -apply (rule contlubI [rule_format]) -apply (case_tac "x = \") -apply (simp add: Istrictify1) -apply (simp add: thelub_const) -apply (simp add: Istrictify2) -apply (erule contlub_cfun_fun) -done - lemma contlub_Istrictify2: "contlub (\x. Istrictify f x)" -apply (rule contlubI [rule_format]) +apply (rule contlubI) apply (case_tac "lub (range Y) = \") -apply (simp add: Istrictify1 chain_UU_I) -apply (simp add: thelub_const) +apply (simp add: Istrictify1 chain_UU_I thelub_const) apply (simp add: Istrictify2) apply (simp add: contlub_cfun_arg) apply (rule lub_equal2) @@ -622,9 +477,6 @@ apply (erule monofun_Istrictify2 [THEN ch2ch_monofun]) done -lemmas cont_Istrictify1 = - monocontlub2cont [OF monofun_Istrictify1 contlub_Istrictify1, standard] - lemmas cont_Istrictify2 = monocontlub2cont [OF monofun_Istrictify2 contlub_Istrictify2, standard]