# HG changeset patch # User huffman # Date 1199296660 -3600 # Node ID 6b3c79acac1f4db4dcc487570ce14c9163760b67 # Parent dbe118fe31808b3e52b85c8a2f6ac14499fa908f move lemmas from Cont.thy to Ffun.thy; reorder dependency of Cont.thy and Ffun.thy; add dcpo instance proofs for function type diff -r dbe118fe3180 -r 6b3c79acac1f src/HOLCF/Adm.thy --- a/src/HOLCF/Adm.thy Wed Jan 02 18:54:21 2008 +0100 +++ b/src/HOLCF/Adm.thy Wed Jan 02 18:57:40 2008 +0100 @@ -140,6 +140,8 @@ text {* admissibility and continuity *} +declare range_composition [simp del] + lemma adm_less: "\cont u; cont v\ \ adm (\x. u x \ v x)" apply (rule admI) apply (simp add: cont2contlubE) @@ -165,7 +167,7 @@ apply (drule_tac x=0 in spec) apply (erule contrapos_nn) apply (erule rev_trans_less) -apply (erule cont2mono [THEN monofun_fun_arg]) +apply (erule cont2mono [THEN monofunE]) apply (erule is_ub_thelub) done diff -r dbe118fe3180 -r 6b3c79acac1f src/HOLCF/Cfun.thy --- a/src/HOLCF/Cfun.thy Wed Jan 02 18:54:21 2008 +0100 +++ b/src/HOLCF/Cfun.thy Wed Jan 02 18:57:40 2008 +0100 @@ -8,7 +8,7 @@ header {* The type of continuous functions *} theory Cfun -imports Pcpodef +imports Pcpodef Ffun uses ("Tools/cont_proc.ML") begin @@ -492,7 +492,7 @@ lemma monofun_strictify2: "monofun (\x. if x = \ then \ else f\x)" apply (rule monofunI) -apply (auto simp add: monofun_cfun_arg eq_UU_iff [symmetric]) +apply (auto simp add: monofun_cfun_arg) done (*FIXME: long proof*) diff -r dbe118fe3180 -r 6b3c79acac1f src/HOLCF/Cont.thy --- a/src/HOLCF/Cont.thy Wed Jan 02 18:54:21 2008 +0100 +++ b/src/HOLCF/Cont.thy Wed Jan 02 18:57:40 2008 +0100 @@ -8,7 +8,7 @@ header {* Continuity and monotonicity *} theory Cont -imports Ffun +imports Pcpo begin text {* @@ -56,19 +56,6 @@ "\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"} *} @@ -185,124 +172,6 @@ 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)\ - \ monofun (\i. F i)" -apply (rule monofunI) -apply (simp add: thelub_fun) -apply (rule lub_mono [rule_format]) -apply (erule ch2ch_fun) -apply (erule ch2ch_fun) -apply (simp add: monofunE) -done - -text {* the lub of a chain of continuous functions is continuous *} - -declare range_composition [simp del] - -lemma contlub_lub_fun: - "\chain F; \i. cont (F i)\ \ contlub (\i. F i)" -apply (rule contlubI) -apply (simp add: thelub_fun) -apply (simp add: cont2contlubE) -apply (rule ex_lub) -apply (erule ch2ch_fun) -apply (simp add: ch2ch_cont) -done - -lemma cont_lub_fun: - "\chain F; \i. cont (F i)\ \ cont (\i. F i)" -apply (rule monocontlub2cont) -apply (erule monofun_lub_fun) -apply (simp add: cont2mono) -apply (erule (1) contlub_lub_fun) -done - -lemma cont2cont_lub: - "\chain F; \i. cont (F i)\ \ cont (\x. \i. F i x)" -by (simp add: thelub_fun [symmetric] cont_lub_fun) - -lemma mono2mono_fun: "monofun f \ monofun (\x. f x y)" -apply (rule monofunI) -apply (erule (1) monofun_fun_arg [THEN monofun_fun_fun]) -done - -lemma cont2cont_fun: "cont f \ cont (\x. f x y)" -apply (rule monocontlub2cont) -apply (erule cont2mono [THEN mono2mono_fun]) -apply (rule contlubI) -apply (simp add: cont2contlubE) -apply (simp add: thelub_fun ch2ch_cont) -done - -text {* Note @{text "(\x. \y. f x y) = f"} *} - -lemma mono2mono_lambda: "(\y. monofun (\x. f x y)) \ monofun f" -apply (rule monofunI) -apply (rule less_fun_ext) -apply (blast dest: monofunE) -done - -lemma cont2cont_lambda: "(\y. cont (\x. f x y)) \ cont f" -apply (subgoal_tac "monofun f") -apply (rule monocontlub2cont) -apply assumption -apply (rule contlubI) -apply (rule ext) -apply (simp add: thelub_fun ch2ch_monofun) -apply (blast dest: cont2contlubE) -apply (simp add: mono2mono_lambda cont2mono) -done - -text {* What D.A.Schmidt calls continuity of abstraction; never used here *} - -lemma contlub_lambda: - "(\x::'a::type. chain (\i. S i x::'b::cpo)) - \ (\x. \i. S i x) = (\i. (\x. S i x))" -by (simp add: thelub_fun ch2ch_lambda) - -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 ch2ch_cont) -apply (simp add: cont2cont_lambda) -apply assumption -done - -lemma mono2mono_app: - "\monofun f; \x. monofun (f x); monofun t\ \ monofun (\x. (f x) (t x))" -apply (rule monofunI) -apply (simp add: monofun_fun monofunE) -done - -lemma cont2contlub_app: - "\cont f; \x. cont (f x); cont t\ \ contlub (\x. (f x) (t x))" -apply (rule contlubI) -apply (subgoal_tac "chain (\i. f (Y i))") -apply (subgoal_tac "chain (\i. t (Y i))") -apply (simp add: cont2contlubE thelub_fun) -apply (rule diag_lub) -apply (erule ch2ch_fun) -apply (drule spec) -apply (erule (1) ch2ch_cont) -apply (erule (1) ch2ch_cont) -apply (erule (1) ch2ch_cont) -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 [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 {* monotone functions map finite chains to finite chains *} diff -r dbe118fe3180 -r 6b3c79acac1f src/HOLCF/Ffun.thy --- a/src/HOLCF/Ffun.thy Wed Jan 02 18:54:21 2008 +0100 +++ b/src/HOLCF/Ffun.thy Wed Jan 02 18:57:40 2008 +0100 @@ -10,7 +10,7 @@ header {* Class instances for the full function space *} theory Ffun -imports Pcpo +imports Cont begin subsection {* Full function space is a partial order *} @@ -49,6 +49,11 @@ subsection {* Full function space is chain complete *} +text {* function application is monotone *} + +lemma monofun_app: "monofun (\f. f x)" +by (rule monofunI, simp add: less_fun_def) + text {* chains of functions yield chains in the po range *} lemma ch2ch_fun: "chain S \ chain (\i. S i x)" @@ -61,7 +66,7 @@ text {* upper bounds of function chains yield upper bound in the po range *} lemma ub2ub_fun: - "range (S::nat \ 'a \ 'b::po) <| u \ range (\i. S i x) <| u x" + "range (S::nat \ 'a::type \ 'b::po) <| u \ range (\i. S i x) <| u x" by (auto simp add: is_ub_def less_fun_def) text {* Type @{typ "'a::type => 'b::cpo"} is chain complete *} @@ -80,11 +85,32 @@ apply (erule ub2ub_fun) done +lemma lub_fun': + fixes S :: "('a::type \ 'b::dcpo) set" + assumes S: "directed S" + shows "S <<| (\x. \f\S. f x)" +apply (rule is_lubI) +apply (rule is_ubI) +apply (rule less_fun_ext) +apply (rule is_ub_thelub') +apply (rule dir2dir_monofun [OF monofun_app S]) +apply (erule imageI) +apply (rule less_fun_ext) +apply (rule is_lub_thelub') +apply (rule dir2dir_monofun [OF monofun_app S]) +apply (erule ub2ub_monofun' [OF monofun_app]) +done + lemma thelub_fun: "chain (S::nat \ 'a::type \ 'b::cpo) \ lub (range S) = (\x. \i. S i x)" by (rule lub_fun [THEN thelubI]) +lemma thelub_fun': + "directed (S::('a::type \ 'b::dcpo) set) + \ lub S = (\x. \f\S. f x)" +by (rule lub_fun' [THEN thelubI]) + lemma cpo_fun: "chain (S::nat \ 'a::type \ 'b::cpo) \ \x. range S <<| x" by (rule exI, erule lub_fun) @@ -92,12 +118,19 @@ instance "fun" :: (type, cpo) cpo by intro_classes (rule cpo_fun) +lemma dcpo_fun: + "directed (S::('a::type \ 'b::dcpo) set) \ \x. S <<| x" +by (rule exI, erule lub_fun') + +instance "fun" :: (type, dcpo) dcpo +by intro_classes (rule dcpo_fun) + subsection {* Full function space is pointed *} lemma minimal_fun: "(\x. \) \ f" by (simp add: less_fun_def) -lemma least_fun: "\x::'a \ 'b::pcpo. \y. x \ y" +lemma least_fun: "\x::'a::type \ 'b::pcpo. \y. x \ y" apply (rule_tac x = "\x. \" in exI) apply (rule minimal_fun [THEN allI]) done @@ -113,5 +146,136 @@ lemma app_strict [simp]: "\ x = \" by (simp add: inst_fun_pcpo) +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 {* 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)\ + \ monofun (\i. F i)" +apply (rule monofunI) +apply (simp add: thelub_fun) +apply (rule lub_mono [rule_format]) +apply (erule ch2ch_fun) +apply (erule ch2ch_fun) +apply (simp add: monofunE) +done + +text {* the lub of a chain of continuous functions is continuous *} + +declare range_composition [simp del] + +lemma contlub_lub_fun: + "\chain F; \i. cont (F i)\ \ contlub (\i. F i)" +apply (rule contlubI) +apply (simp add: thelub_fun) +apply (simp add: cont2contlubE) +apply (rule ex_lub) +apply (erule ch2ch_fun) +apply (simp add: ch2ch_cont) +done + +lemma cont_lub_fun: + "\chain F; \i. cont (F i)\ \ cont (\i. F i)" +apply (rule monocontlub2cont) +apply (erule monofun_lub_fun) +apply (simp add: cont2mono) +apply (erule (1) contlub_lub_fun) +done + +lemma cont2cont_lub: + "\chain F; \i. cont (F i)\ \ cont (\x. \i. F i x)" +by (simp add: thelub_fun [symmetric] cont_lub_fun) + +lemma mono2mono_fun: "monofun f \ monofun (\x. f x y)" +apply (rule monofunI) +apply (erule (1) monofun_fun_arg [THEN monofun_fun_fun]) +done + +lemma cont2cont_fun: "cont f \ cont (\x. f x y)" +apply (rule monocontlub2cont) +apply (erule cont2mono [THEN mono2mono_fun]) +apply (rule contlubI) +apply (simp add: cont2contlubE) +apply (simp add: thelub_fun ch2ch_cont) +done + +text {* Note @{text "(\x. \y. f x y) = f"} *} + +lemma mono2mono_lambda: "(\y. monofun (\x. f x y)) \ monofun f" +apply (rule monofunI) +apply (rule less_fun_ext) +apply (blast dest: monofunE) +done + +lemma cont2cont_lambda: "(\y. cont (\x. f x y)) \ cont f" +apply (subgoal_tac "monofun f") +apply (rule monocontlub2cont) +apply assumption +apply (rule contlubI) +apply (rule ext) +apply (simp add: thelub_fun ch2ch_monofun) +apply (blast dest: cont2contlubE) +apply (simp add: mono2mono_lambda cont2mono) +done + +text {* What D.A.Schmidt calls continuity of abstraction; never used here *} + +lemma contlub_lambda: + "(\x::'a::type. chain (\i. S i x::'b::cpo)) + \ (\x. \i. S i x) = (\i. (\x. S i x))" +by (simp add: thelub_fun ch2ch_lambda) + +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 ch2ch_cont) +apply (simp add: cont2cont_lambda) +apply assumption +done + +lemma mono2mono_app: + "\monofun f; \x. monofun (f x); monofun t\ \ monofun (\x. (f x) (t x))" +apply (rule monofunI) +apply (simp add: monofun_fun monofunE) +done + +lemma cont2contlub_app: + "\cont f; \x. cont (f x); cont t\ \ contlub (\x. (f x) (t x))" +apply (rule contlubI) +apply (subgoal_tac "chain (\i. f (Y i))") +apply (subgoal_tac "chain (\i. t (Y i))") +apply (simp add: cont2contlubE thelub_fun) +apply (rule diag_lub) +apply (erule ch2ch_fun) +apply (drule spec) +apply (erule (1) ch2ch_cont) +apply (erule (1) ch2ch_cont) +apply (erule (1) ch2ch_cont) +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 [rule_format] + +lemma cont2cont_app3: "\cont f; cont t\ \ cont (\x. f (t x))" +by (rule cont2cont_app2 [OF cont_const]) + end