# HG changeset patch # User huffman # Date 1120175407 -7200 # Node ID 645b9560f3fd36aeedafcd5915a9f3c49e68de72 # Parent f3fcfa388ecb3d7b0e936adc68dd39436aaebc8e cleaned up; reorganized and added section headings diff -r f3fcfa388ecb -r 645b9560f3fd src/HOLCF/Cont.thy --- a/src/HOLCF/Cont.thy Fri Jul 01 01:48:37 2005 +0200 +++ b/src/HOLCF/Cont.thy Fri Jul 01 01:50:07 2005 +0200 @@ -18,27 +18,17 @@ defaultsort po -consts - monofun :: "('a => 'b) => bool" -- "monotonicity" - contlub :: "('a::cpo => 'b::cpo) => bool" -- "first cont. def" - cont :: "('a::cpo => 'b::cpo) => bool" -- "secnd cont. def" +subsection {* Definitions *} -defs - -monofun_def: "monofun(f) == ! x y. x << y --> f(x) << f(y)" +constdefs + monofun :: "('a \ 'b) \ bool" -- "monotonicity" + "monofun f \ \x y. x \ y \ f x \ f y" -contlub_def: "contlub(f) == ! Y. chain(Y) --> - f(lub(range(Y))) = lub(range(% i. f(Y(i))))" - -cont_def: "cont(f) == ! Y. chain(Y) --> - range(% i. f(Y(i))) <<| f(lub(range(Y)))" + contlub :: "('a::cpo \ 'b::cpo) \ bool" -- "first cont. def" + "contlub f \ \Y. chain Y \ f (\i. Y i) = (\i. f (Y i))" -text {* - the main purpose of cont.thy is to show: - @{prop "monofun(f) & contlub(f) == cont(f)"} -*} - -text {* access to definition *} + cont :: "('a::cpo \ 'b::cpo) \ bool" -- "secnd cont. def" + "cont f \ \Y. chain Y \ range (\i. f (Y i)) <<| f (\i. Y i)" lemma contlubI: "\\Y. chain Y \ f (\i. Y i) = (\i. f (Y i))\ \ contlub f" @@ -64,6 +54,22 @@ "\monofun f; x \ y\ \ f x \ f y" by (simp add: monofun_def) +text {* + The following results are about application for functions in @{typ "'a=>'b"} +*} + +lemma monofun_fun_fun: "f \ g \ f x \ g x" +by (simp add: less_fun_def) + +lemma monofun_fun_arg: "\monofun f; x \ y\ \ f x \ f y" +by (rule monofunE) + +lemma monofun_fun: "\monofun f; monofun g; f \ g; x \ y\ \ f x \ g y" +by (rule trans_less [OF monofun_fun_arg monofun_fun_fun]) + + +subsection {* @{prop "monofun f \ contlub f \ cont f"} *} + text {* monotone functions map chains to chains *} lemma ch2ch_monofun: "\monofun f; chain Y\ \ chain (\i. f (Y i))" @@ -81,7 +87,7 @@ apply (erule ub_rangeD) done -text {* left to right: @{prop "monofun(f) & contlub(f) ==> cont(f)"} *} +text {* left to right: @{prop "monofun f \ contlub f \ cont f"} *} lemma monocontlub2cont: "\monofun f; contlub f\ \ cont f" apply (rule contI) @@ -117,26 +123,38 @@ text {* right to left: @{prop "cont f \ monofun f \ contlub f"} *} text {* part2: @{prop "cont f \ contlub f"} *} -lemma cont2contlub: "cont(f) ==> contlub(f)" +lemma cont2contlub: "cont f \ contlub f" apply (rule contlubI) apply (rule thelubI [symmetric]) apply (erule contE) apply assumption done -text {* monotone functions map finite chains to finite chains *} +subsection {* Continuity of basic functions *} + +text {* The identity function is continuous *} -lemma monofun_finch2finch: - "\monofun f; finite_chain Y\ \ finite_chain (\n. f (Y n))" -apply (unfold finite_chain_def) -apply (simp add: ch2ch_monofun) -apply (force simp add: max_in_chain_def) +lemma cont_id: "cont (\x. x)" +apply (rule contI) +apply (erule thelubE) +apply (rule refl) done -text {* The same holds for continuous functions *} +text {* constant functions are continuous *} + +lemma cont_const: "cont (\x. c)" +apply (rule contI) +apply (rule lub_const) +done -lemmas cont_finch2finch = cont2mono [THEN monofun_finch2finch, standard] -(* [| cont ?f; finite_chain ?Y |] ==> finite_chain (%n. ?f (?Y n)) *) +text {* if-then-else is continuous *} + +lemma cont_if: "\cont f; cont g\ \ cont (\x. if b then f x else g x)" +by (induct b) simp_all + +subsection {* Propagation of monotonicity and continuity *} + +text {* the lub of a chain of monotone functions is monotone *} lemma monofun_lub_fun: "\chain (F::nat \ 'a \ 'b::cpo); \i. monofun (F i)\ @@ -176,45 +194,20 @@ "\chain F; \i. cont (F i)\ \ cont (\x. \i. F i x)" by (simp add: thelub_fun [symmetric] cont_lub_fun) -text {* - The following results are about application for functions in @{typ "'a=>'b"} -*} - -lemma monofun_fun_fun: "f \ g \ f x \ g x" -by (simp add: less_fun_def) - -lemma monofun_fun_arg: "\monofun f; x \ y\ \ f x \ f y" -by (rule monofunE) - -lemma monofun_fun: "\monofun f; monofun g; f \ g; x \ y\ \ f x \ g y" -by (rule trans_less [OF monofun_fun_arg monofun_fun_fun]) - -text {* - The following results are about the propagation of monotonicity and - continuity -*} - lemma mono2mono_MF1L: "monofun f \ monofun (\x. f x y)" apply (rule monofunI) -apply (erule monofun_fun_arg [THEN monofun_fun_fun]) -apply assumption +apply (erule (1) monofun_fun_arg [THEN monofun_fun_fun]) done lemma cont2cont_CF1L: "cont f \ cont (\x. f x y)" apply (rule monocontlub2cont) apply (erule cont2mono [THEN mono2mono_MF1L]) apply (rule contlubI) -apply (frule asm_rl) -apply (erule cont2contlub [THEN contlubE, THEN ssubst]) -apply assumption -apply (subst thelub_fun) -apply (rule ch2ch_monofun) -apply (erule cont2mono) -apply assumption -apply (rule refl) +apply (simp add: cont2contlub [THEN contlubE]) +apply (simp add: thelub_fun cont2mono [THEN ch2ch_monofun]) done -(********* Note "(%x.%y.c1 x y) = c1" ***********) +text {* Note @{text "(\x. \y. f x y) = f"} *} lemma mono2mono_MF1L_rev: "\y. monofun (\x. f x y) \ monofun f" apply (rule monofunI) @@ -223,35 +216,29 @@ done lemma cont2cont_CF1L_rev: "\y. cont (\x. f x y) \ cont f" +apply (subgoal_tac "monofun f") apply (rule monocontlub2cont) -apply (rule cont2mono [THEN allI, THEN mono2mono_MF1L_rev]) -apply (erule spec) +apply assumption apply (rule contlubI) apply (rule ext) -apply (subst thelub_fun) -apply (rule cont2mono [THEN allI, THEN mono2mono_MF1L_rev, THEN ch2ch_monofun]) -apply (erule spec) -apply assumption +apply (simp add: thelub_fun ch2ch_monofun) apply (blast dest: cont2contlub [THEN contlubE]) +apply (simp add: mono2mono_MF1L_rev cont2mono) done -lemma cont2cont_lambda: "(\y. cont (\x. f x y)) \ cont (\x y. f x y)" +lemma cont2cont_lambda: "(\y. cont (\x. f x y)) \ cont (\x. (\y. f x y))" apply (rule cont2cont_CF1L_rev) apply simp done -text {* - What D.A.Schmidt calls continuity of abstraction - never used here -*} +text {* What D.A.Schmidt calls continuity of abstraction; never used here *} lemma contlub_abstraction: "\chain Y; \y. cont (\x.(c::'a::cpo\'b::type\'c::cpo) x y)\ \ (\y. \i. c (Y i) y) = (\i. (\y. c (Y i) y))" apply (rule thelub_fun [symmetric]) apply (rule cont2mono [THEN ch2ch_monofun]) -apply (erule cont2cont_CF1L_rev) -apply assumption +apply (erule (1) cont2cont_CF1L_rev) done lemma mono2mono_app: @@ -269,45 +256,36 @@ apply (rule diag_lub) apply (erule ch2ch_fun) apply (drule spec) -apply (erule cont2mono [THEN ch2ch_monofun], assumption) -apply (erule cont2mono [THEN ch2ch_monofun], assumption) -apply (erule cont2mono [THEN ch2ch_monofun], assumption) +apply (erule (1) cont2mono [THEN ch2ch_monofun]) +apply (erule (1) cont2mono [THEN ch2ch_monofun]) +apply (erule (1) cont2mono [THEN ch2ch_monofun]) done lemma cont2cont_app: "\cont f; \x. cont (f x); cont t\ \ cont (\x. (f x) (t x))" by (blast intro: monocontlub2cont mono2mono_app cont2mono cont2contlub_app) -lemmas cont2cont_app2 = cont2cont_app[OF _ allI] -(* [| cont ?ft; !!x. cont (?ft x); cont ?tt |] ==> *) -(* cont (%x. ?ft x (?tt x)) *) +lemmas cont2cont_app2 = cont2cont_app [rule_format] +lemma cont2cont_app3: "\cont f; cont t\ \ cont (\x. f (t x))" +by (rule cont2cont_app2 [OF cont_const]) + +subsection {* Finite chains and flat pcpos *} -text {* The identity function is continuous *} +text {* monotone functions map finite chains to finite chains *} -lemma cont_id: "cont (\x. x)" -apply (rule contI) -apply (erule thelubE) -apply (rule refl) +lemma monofun_finch2finch: + "\monofun f; finite_chain Y\ \ finite_chain (\n. f (Y n))" +apply (unfold finite_chain_def) +apply (simp add: ch2ch_monofun) +apply (force simp add: max_in_chain_def) done -text {* constant functions are continuous *} - -lemma cont_const: "cont (\x. c)" -apply (rule contI) -apply (rule lub_const) -done +text {* The same holds for continuous functions *} -lemma cont2cont_app3: "[|cont(f); cont(t) |] ==> cont(%x. f(t(x)))" -by (best intro: cont2cont_app2 cont_const) - -text {* some properties of flat *} - -lemma flatdom2monofun: "f \ = \ \ monofun (f::'a::flat \ 'b::pcpo)" -apply (rule monofunI) -apply (drule ax_flat [rule_format]) -apply auto -done +lemma cont_finch2finch: + "\cont f; finite_chain Y\ \ finite_chain (\n. f (Y n))" +by (rule cont2mono [THEN monofun_finch2finch]) lemma chfindom_monofun2cont: "monofun f \ cont (f::'a::chfin \ 'b::pcpo)" apply (rule monocontlub2cont) @@ -320,7 +298,15 @@ apply (force simp add: max_in_chain_def) done -lemmas flatdom_strict2cont = flatdom2monofun [THEN chfindom_monofun2cont, standard] -(* f UU = UU ==> cont (f::'a=>'b::pcpo)" *) +text {* some properties of flat *} + +lemma flatdom_strict2mono: "f \ = \ \ monofun (f::'a::flat \ 'b::pcpo)" +apply (rule monofunI) +apply (drule ax_flat [rule_format]) +apply auto +done + +lemma flatdom_strict2cont: "f \ = \ \ cont (f::'a::flat \ 'b::pcpo)" +by (rule flatdom_strict2mono [THEN chfindom_monofun2cont]) end