# HG changeset patch # User huffman # Date 1109802497 -3600 # Node ID 2454493bd77bc5a0be33f2dab3275c5bf96e3ed5 # Parent c899efea601fd0e4a932a83fec96b3024c659561 converted to new-style theory diff -r c899efea601f -r 2454493bd77b src/HOLCF/Cont.ML --- a/src/HOLCF/Cont.ML Wed Mar 02 23:15:16 2005 +0100 +++ b/src/HOLCF/Cont.ML Wed Mar 02 23:28:17 2005 +0100 @@ -1,525 +1,48 @@ -(* Title: HOLCF/Cont.ML - ID: $Id$ - Author: Franz Regensburger -Results about continuity and monotonicity -*) - -(* ------------------------------------------------------------------------ *) -(* access to definition *) -(* ------------------------------------------------------------------------ *) - -Goalw [contlub] - "! Y. chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))==>\ -\ contlub(f)"; -by (atac 1); -qed "contlubI"; - -Goalw [contlub] - " contlub(f)==>\ -\ ! Y. chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))"; -by (atac 1); -qed "contlubE"; - - -Goalw [cont] - "! Y. chain(Y) --> range(% i. f(Y(i))) <<| f(lub(range(Y))) ==> cont(f)"; -by (atac 1); -qed "contI"; - -Goalw [cont] - "cont(f) ==> ! Y. chain(Y) --> range(% i. f(Y(i))) <<| f(lub(range(Y)))"; -by (atac 1); -qed "contE"; - - -Goalw [monofun] - "! x y. x << y --> f(x) << f(y) ==> monofun(f)"; -by (atac 1); -qed "monofunI"; - -Goalw [monofun] - "monofun(f) ==> ! x y. x << y --> f(x) << f(y)"; -by (atac 1); -qed "monofunE"; - -(* ------------------------------------------------------------------------ *) -(* the main purpose of cont.thy is to show: *) -(* monofun(f) & contlub(f) <==> cont(f) *) -(* ------------------------------------------------------------------------ *) - -(* ------------------------------------------------------------------------ *) -(* monotone functions map chains to chains *) -(* ------------------------------------------------------------------------ *) - -Goal - "[| monofun(f); chain(Y) |] ==> chain(%i. f(Y(i)))"; -by (rtac chainI 1); -by (etac (monofunE RS spec RS spec RS mp) 1); -by (etac (chainE) 1); -qed "ch2ch_monofun"; - -(* ------------------------------------------------------------------------ *) -(* monotone functions map upper bound to upper bounds *) -(* ------------------------------------------------------------------------ *) - -Goal - "[| monofun(f); range(Y) <| u|] ==> range(%i. f(Y(i))) <| f(u)"; -by (rtac ub_rangeI 1); -by (etac (monofunE RS spec RS spec RS mp) 1); -by (etac (ub_rangeD) 1); -qed "ub2ub_monofun"; - -(* ------------------------------------------------------------------------ *) -(* left to right: monofun(f) & contlub(f) ==> cont(f) *) -(* ------------------------------------------------------------------------ *) - -Goalw [cont] - "[|monofun(f);contlub(f)|] ==> cont(f)"; -by (strip_tac 1); -by (rtac thelubE 1); -by (etac ch2ch_monofun 1); -by (atac 1); -by (etac (contlubE RS spec RS mp RS sym) 1); -by (atac 1); -qed "monocontlub2cont"; - -(* ------------------------------------------------------------------------ *) -(* first a lemma about binary chains *) -(* ------------------------------------------------------------------------ *) - -Goal "[| cont(f); x << y |] \ -\ ==> range(%i::nat. f(if i = 0 then x else y)) <<| f(y)"; -by (rtac subst 1); -by (etac (contE RS spec RS mp) 2); -by (etac bin_chain 2); -by (res_inst_tac [("y","y")] arg_cong 1); -by (etac (lub_bin_chain RS thelubI) 1); -qed "binchain_cont"; - -(* ------------------------------------------------------------------------ *) -(* right to left: cont(f) ==> monofun(f) & contlub(f) *) -(* part1: cont(f) ==> monofun(f *) -(* ------------------------------------------------------------------------ *) - -Goalw [monofun] "cont(f) ==> monofun(f)"; -by (strip_tac 1); -by (dtac (binchain_cont RS is_ub_lub) 1); -by (auto_tac (claset(), simpset() addsplits [split_if_asm])); -qed "cont2mono"; - -(* ------------------------------------------------------------------------ *) -(* right to left: cont(f) ==> monofun(f) & contlub(f) *) -(* part2: cont(f) ==> contlub(f) *) -(* ------------------------------------------------------------------------ *) - -Goalw [contlub] "cont(f) ==> contlub(f)"; -by (strip_tac 1); -by (rtac (thelubI RS sym) 1); -by (etac (contE RS spec RS mp) 1); -by (atac 1); -qed "cont2contlub"; - -(* ------------------------------------------------------------------------ *) -(* monotone functions map finite chains to finite chains *) -(* ------------------------------------------------------------------------ *) - -Goalw [finite_chain_def] - "[| monofun f; finite_chain Y |] ==> finite_chain (%n. f (Y n))"; -by (force_tac (claset() addSEs [ch2ch_monofun], - simpset() addsimps [max_in_chain_def]) 1); -qed "monofun_finch2finch"; - -(* ------------------------------------------------------------------------ *) -(* The same holds for continuous functions *) -(* ------------------------------------------------------------------------ *) - -bind_thm ("cont_finch2finch", cont2mono RS monofun_finch2finch); -(* [| cont ?f; finite_chain ?Y |] ==> finite_chain (%n. ?f (?Y n)) *) - - -(* ------------------------------------------------------------------------ *) -(* The following results are about a curried function that is monotone *) -(* in both arguments *) -(* ------------------------------------------------------------------------ *) - -Goal -"[|monofun(MF2); chain(F)|] ==> chain(%i. MF2 (F i) x)"; -by (etac (ch2ch_monofun RS ch2ch_fun) 1); -by (atac 1); -qed "ch2ch_MF2L"; - - -Goal -"[|monofun(MF2(f)); chain(Y)|] ==> chain(%i. MF2 f (Y i))"; -by (etac ch2ch_monofun 1); -by (atac 1); -qed "ch2ch_MF2R"; - -Goal -"[|monofun(MF2); !f. monofun(MF2(f)); chain(F); chain(Y)|] ==> \ -\ chain(%i. MF2(F(i))(Y(i)))"; -by (rtac chainI 1); -by (rtac trans_less 1); -by (etac (ch2ch_MF2L RS chainE) 1); -by (atac 1); -by ((rtac (monofunE RS spec RS spec RS mp) 1) THEN (etac spec 1)); -by (etac (chainE) 1); -qed "ch2ch_MF2LR"; - - -Goal -"[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\ -\ !f. monofun(MF2(f)::('b::po=>'c::cpo));\ -\ chain(F);chain(Y)|] ==> \ -\ chain(%j. lub(range(%i. MF2 (F j) (Y i))))"; -by (rtac (lub_mono RS chainI) 1); -by ((rtac ch2ch_MF2R 1) THEN (etac spec 1)); -by (atac 1); -by ((rtac ch2ch_MF2R 1) THEN (etac spec 1)); -by (atac 1); -by (strip_tac 1); -by (rtac (chainE) 1); -by (etac ch2ch_MF2L 1); -by (atac 1); -qed "ch2ch_lubMF2R"; - - -Goal -"[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\ -\ !f. monofun(MF2(f)::('b::po=>'c::cpo));\ -\ chain(F);chain(Y)|] ==> \ -\ chain(%i. lub(range(%j. MF2 (F j) (Y i))))"; -by (rtac (lub_mono RS chainI) 1); -by (etac ch2ch_MF2L 1); -by (atac 1); -by (etac ch2ch_MF2L 1); -by (atac 1); -by (strip_tac 1); -by (rtac (chainE) 1); -by ((rtac ch2ch_MF2R 1) THEN (etac spec 1)); -by (atac 1); -qed "ch2ch_lubMF2L"; - - -Goal -"[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\ -\ !f. monofun(MF2(f)::('b::po=>'c::cpo));\ -\ chain(F)|] ==> \ -\ monofun(% x. lub(range(% j. MF2 (F j) (x))))"; -by (rtac monofunI 1); -by (strip_tac 1); -by (rtac lub_mono 1); -by (etac ch2ch_MF2L 1); -by (atac 1); -by (etac ch2ch_MF2L 1); -by (atac 1); -by (strip_tac 1); -by ((rtac (monofunE RS spec RS spec RS mp) 1) THEN (etac spec 1)); -by (atac 1); -qed "lub_MF2_mono"; - -Goal -"[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\ -\ !f. monofun(MF2(f)::('b::po=>'c::cpo));\ -\ chain(F); chain(Y)|] ==> \ -\ lub(range(%j. lub(range(%i. MF2(F j) (Y i))))) =\ -\ lub(range(%i. lub(range(%j. MF2(F j) (Y i)))))"; -by (rtac antisym_less 1); -by (rtac (ub_rangeI RSN (2,is_lub_thelub)) 1); -by (etac ch2ch_lubMF2R 1); -by (REPEAT (atac 1)); -by (strip_tac 1); -by (rtac lub_mono 1); -by ((rtac ch2ch_MF2R 1) THEN (etac spec 1)); -by (atac 1); -by (etac ch2ch_lubMF2L 1); -by (REPEAT (atac 1)); -by (strip_tac 1); -by (rtac is_ub_thelub 1); -by (etac ch2ch_MF2L 1); -by (atac 1); -by (rtac (ub_rangeI RSN (2,is_lub_thelub)) 1); -by (etac ch2ch_lubMF2L 1); -by (REPEAT (atac 1)); -by (strip_tac 1); -by (rtac lub_mono 1); -by (etac ch2ch_MF2L 1); -by (atac 1); -by (etac ch2ch_lubMF2R 1); -by (REPEAT (atac 1)); -by (strip_tac 1); -by (rtac is_ub_thelub 1); -by ((rtac ch2ch_MF2R 1) THEN (etac spec 1)); -by (atac 1); -qed "ex_lubMF2"; - +(* legacy ML bindings *) -Goal -"[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\ -\ !f. monofun(MF2(f)::('b::po=>'c::cpo));\ -\ chain(FY);chain(TY)|] ==>\ -\ lub(range(%i. lub(range(%j. MF2(FY(j))(TY(i)))))) =\ -\ lub(range(%i. MF2(FY(i))(TY(i))))"; -by (rtac antisym_less 1); -by (rtac (ub_rangeI RSN (2,is_lub_thelub)) 1); -by (etac ch2ch_lubMF2L 1); -by (REPEAT (atac 1)); -by (strip_tac 1 ); -by (rtac lub_mono3 1); -by (etac ch2ch_MF2L 1); -by (REPEAT (atac 1)); -by (etac ch2ch_MF2LR 1); -by (REPEAT (atac 1)); -by (rtac allI 1); -by (res_inst_tac [("m","i"),("n","ia")] nat_less_cases 1); -by (res_inst_tac [("x","ia")] exI 1); -by (rtac (chain_mono) 1); -by (etac allE 1); -by (etac ch2ch_MF2R 1); -by (REPEAT (atac 1)); -by (hyp_subst_tac 1); -by (res_inst_tac [("x","ia")] exI 1); -by (rtac refl_less 1); -by (res_inst_tac [("x","i")] exI 1); -by (rtac (chain_mono) 1); -by (etac ch2ch_MF2L 1); -by (REPEAT (atac 1)); -by (rtac lub_mono 1); -by (etac ch2ch_MF2LR 1); -by (REPEAT(atac 1)); -by (etac ch2ch_lubMF2L 1); -by (REPEAT (atac 1)); -by (strip_tac 1 ); -by (rtac is_ub_thelub 1); -by (etac ch2ch_MF2L 1); -by (atac 1); -qed "diag_lubMF2_1"; - -Goal -"[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\ -\ !f. monofun(MF2(f)::('b::po=>'c::cpo));\ -\ chain(FY);chain(TY)|] ==>\ -\ lub(range(%j. lub(range(%i. MF2(FY(j))(TY(i)))))) =\ -\ lub(range(%i. MF2(FY(i))(TY(i))))"; -by (rtac trans 1); -by (rtac ex_lubMF2 1); -by (REPEAT (atac 1)); -by (etac diag_lubMF2_1 1); -by (REPEAT (atac 1)); -qed "diag_lubMF2_2"; - - -(* ------------------------------------------------------------------------ *) -(* The following results are about a curried function that is continuous *) -(* in both arguments *) -(* ------------------------------------------------------------------------ *) - -val [prem1,prem2,prem3,prem4] = goal thy -"[| cont(CF2); !f. cont(CF2(f)); chain(FY); chain(TY)|] ==>\ -\ CF2(lub(range(FY)))(lub(range(TY))) = lub(range(%i. CF2(FY(i))(TY(i))))"; -by (cut_facts_tac [prem1,prem2,prem3, prem4] 1); -by (stac (prem1 RS cont2contlub RS contlubE RS spec RS mp) 1); -by (assume_tac 1); -by (stac thelub_fun 1); -by (rtac (prem1 RS (cont2mono RS ch2ch_monofun)) 1); -by (assume_tac 1); -by (rtac trans 1); -by (rtac ((prem2 RS spec RS cont2contlub) RS contlubE RS spec RS mp RS ext RS arg_cong RS arg_cong) 1); -by (rtac prem4 1); -by (blast_tac (claset() addIs [diag_lubMF2_2, cont2mono]) 1); -qed "contlub_CF2"; - -(* ------------------------------------------------------------------------ *) -(* The following results are about application for functions in 'a=>'b *) -(* ------------------------------------------------------------------------ *) - -Goal "f1 << f2 ==> f1(x) << f2(x)"; -by (etac (less_fun RS iffD1 RS spec) 1); -qed "monofun_fun_fun"; - -Goal "[|monofun(f); x1 << x2|] ==> f(x1) << f(x2)"; -by (etac (monofunE RS spec RS spec RS mp) 1); -by (atac 1); -qed "monofun_fun_arg"; - -Goal "[|monofun(f1); monofun(f2); f1 << f2; x1 << x2|] ==> f1(x1) << f2(x2)"; -by (rtac trans_less 1); -by (etac monofun_fun_arg 1); -by (atac 1); -by (etac monofun_fun_fun 1); -qed "monofun_fun"; - - -(* ------------------------------------------------------------------------ *) -(* The following results are about the propagation of monotonicity and *) -(* continuity *) -(* ------------------------------------------------------------------------ *) - -Goal "[|monofun(c1)|] ==> monofun(%x. c1 x y)"; -by (rtac monofunI 1); -by (strip_tac 1); -by (etac (monofun_fun_arg RS monofun_fun_fun) 1); -by (atac 1); -qed "mono2mono_MF1L"; - -Goal "[|cont(c1)|] ==> cont(%x. c1 x y)"; -by (rtac monocontlub2cont 1); -by (etac (cont2mono RS mono2mono_MF1L) 1); -by (rtac contlubI 1); -by (strip_tac 1); -by (ftac asm_rl 1); -by (etac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1); -by (atac 1); -by (stac thelub_fun 1); -by (rtac ch2ch_monofun 1); -by (etac cont2mono 1); -by (atac 1); -by (rtac refl 1); -qed "cont2cont_CF1L"; - -(********* Note "(%x.%y.c1 x y) = c1" ***********) - -Goal "!y. monofun(%x. c1 x y) ==> monofun(c1)"; -by (rtac monofunI 1); -by (strip_tac 1); -by (rtac (less_fun RS iffD2) 1); -by (blast_tac (claset() addDs [monofunE]) 1); -qed "mono2mono_MF1L_rev"; - -Goal "!y. cont(%x. c1 x y) ==> cont(c1)"; -by (rtac monocontlub2cont 1); -by (rtac (cont2mono RS allI RS mono2mono_MF1L_rev ) 1); -by (etac spec 1); -by (rtac contlubI 1); -by (strip_tac 1); -by (rtac ext 1); -by (stac thelub_fun 1); -by (rtac (cont2mono RS allI RS mono2mono_MF1L_rev RS ch2ch_monofun) 1); -by (etac spec 1); -by (atac 1); -by (blast_tac (claset() addDs [ cont2contlub RS contlubE]) 1); -qed "cont2cont_CF1L_rev"; - -(* ------------------------------------------------------------------------ *) -(* What D.A.Schmidt calls continuity of abstraction *) -(* never used here *) -(* ------------------------------------------------------------------------ *) - -Goal -"[|chain(Y::nat=>'a);!y. cont(%x.(c::'a::cpo=>'b::cpo=>'c::cpo) x y)|] ==>\ -\ (%y. lub(range(%i. c (Y i) y))) = (lub(range(%i.%y. c (Y i) y)))"; -by (rtac trans 1); -by (rtac (cont2contlub RS contlubE RS spec RS mp) 2); -by (atac 3); -by (etac cont2cont_CF1L_rev 2); -by (rtac ext 1); -by (rtac (cont2contlub RS contlubE RS spec RS mp RS sym) 1); -by (etac spec 1); -by (atac 1); -qed "contlub_abstraction"; - -Goal "[|monofun(ft);!x. monofun(ft(x));monofun(tt)|] ==>\ -\ monofun(%x.(ft(x))(tt(x)))"; -by (rtac monofunI 1); -by (strip_tac 1); -by (res_inst_tac [("f1.0","ft(x)"),("f2.0","ft(y)")] monofun_fun 1); -by (etac spec 1); -by (etac spec 1); -by (etac (monofunE RS spec RS spec RS mp) 1); -by (atac 1); -by (etac (monofunE RS spec RS spec RS mp) 1); -by (atac 1); -qed "mono2mono_app"; - - -Goal "[|cont(ft);!x. cont(ft(x));cont(tt)|] ==> contlub(%x.(ft(x))(tt(x)))"; -by (rtac contlubI 1); -by (strip_tac 1); -by (res_inst_tac [("f3","tt")] (contlubE RS spec RS mp RS ssubst) 1); -by (etac cont2contlub 1); -by (atac 1); -by (rtac contlub_CF2 1); -by (REPEAT (atac 1)); -by (etac (cont2mono RS ch2ch_monofun) 1); -by (atac 1); -qed "cont2contlub_app"; - - -Goal "[|cont(ft); !x. cont(ft(x)); cont(tt)|] ==> cont(%x.(ft(x))(tt(x)))"; -by (blast_tac (claset() addIs [monocontlub2cont, mono2mono_app, cont2mono, - cont2contlub_app]) 1); -qed "cont2cont_app"; - - -bind_thm ("cont2cont_app2", allI RSN (2,cont2cont_app)); -(* [| cont ?ft; !!x. cont (?ft x); cont ?tt |] ==> *) -(* cont (%x. ?ft x (?tt x)) *) - - -(* ------------------------------------------------------------------------ *) -(* The identity function is continuous *) -(* ------------------------------------------------------------------------ *) - -Goal "cont(% x. x)"; -by (rtac contI 1); -by (strip_tac 1); -by (etac thelubE 1); -by (rtac refl 1); -qed "cont_id"; - -(* ------------------------------------------------------------------------ *) -(* constant functions are continuous *) -(* ------------------------------------------------------------------------ *) - -Goalw [cont] "cont(%x. c)"; -by (strip_tac 1); -by (blast_tac (claset() addIs [is_lubI, ub_rangeI] addDs [ub_rangeD]) 1); -qed "cont_const"; - - -Goal "[|cont(f); cont(t) |] ==> cont(%x. f(t(x)))"; -by (best_tac (claset() addIs [ cont2cont_app2, cont_const]) 1); -qed "cont2cont_app3"; - -(* ------------------------------------------------------------------------ *) -(* A non-emptyness result for Cfun *) -(* ------------------------------------------------------------------------ *) - -Goal "?x:Collect cont"; -by (rtac CollectI 1); -by (rtac cont_const 1); -qed "CfunI"; - -(* ------------------------------------------------------------------------ *) -(* some properties of flat *) -(* ------------------------------------------------------------------------ *) - -Goalw [monofun] "f UU = UU ==> monofun (f::'a::flat=>'b::pcpo)"; -by (strip_tac 1); -by (dtac (ax_flat RS spec RS spec RS mp) 1); -by (fast_tac ((HOL_cs addss (simpset() addsimps [minimal]))) 1); -qed "flatdom2monofun"; - - -Goal "monofun f ==> cont(f::'a::chfin=>'b::pcpo)"; -by (rtac monocontlub2cont 1); -by ( atac 1); -by (rtac contlubI 1); -by (strip_tac 1); -by (ftac chfin2finch 1); -by (rtac antisym_less 1); -by ( force_tac (HOL_cs addIs [is_ub_thelub,ch2ch_monofun], - HOL_ss addsimps [finite_chain_def,maxinch_is_thelub]) 1); -by (dtac (monofun_finch2finch COMP swap_prems_rl) 1); -by ( atac 1); -by (asm_full_simp_tac (HOL_ss addsimps [finite_chain_def]) 1); -by (etac conjE 1); -by (etac exE 1); -by (asm_full_simp_tac (HOL_ss addsimps [maxinch_is_thelub]) 1); -by (etac (monofunE RS spec RS spec RS mp) 1); -by (etac is_ub_thelub 1); -qed "chfindom_monofun2cont"; - -bind_thm ("flatdom_strict2cont",flatdom2monofun RS chfindom_monofun2cont); -(* f UU = UU ==> cont (f::'a=>'b::pcpo)" *) +val monofun = thm "monofun"; +val contlub = thm "contlub"; +val cont = thm "cont"; +val contlubI = thm "contlubI"; +val contlubE = thm "contlubE"; +val contI = thm "contI"; +val contE = thm "contE"; +val monofunI = thm "monofunI"; +val monofunE = thm "monofunE"; +val ch2ch_monofun = thm "ch2ch_monofun"; +val ub2ub_monofun = thm "ub2ub_monofun"; +val monocontlub2cont = thm "monocontlub2cont"; +val binchain_cont = thm "binchain_cont"; +val cont2mono = thm "cont2mono"; +val cont2contlub = thm "cont2contlub"; +val monofun_finch2finch = thm "monofun_finch2finch"; +val cont_finch2finch = thm "cont_finch2finch"; +val ch2ch_MF2L = thm "ch2ch_MF2L"; +val ch2ch_MF2R = thm "ch2ch_MF2R"; +val ch2ch_MF2LR = thm "ch2ch_MF2LR"; +val ch2ch_lubMF2R = thm "ch2ch_lubMF2R"; +val ch2ch_lubMF2L = thm "ch2ch_lubMF2L"; +val lub_MF2_mono = thm "lub_MF2_mono"; +val ex_lubMF2 = thm "ex_lubMF2"; +val diag_lubMF2_1 = thm "diag_lubMF2_1"; +val diag_lubMF2_2 = thm "diag_lubMF2_2"; +val contlub_CF2 = thm "contlub_CF2"; +val monofun_fun_fun = thm "monofun_fun_fun"; +val monofun_fun_arg = thm "monofun_fun_arg"; +val mono2mono_MF1L = thm "mono2mono_MF1L"; +val cont2cont_CF1L = thm "cont2cont_CF1L"; +val mono2mono_MF1L_rev = thm "mono2mono_MF1L_rev"; +val cont2cont_CF1L_rev = thm "cont2cont_CF1L_rev"; +val contlub_abstraction = thm "contlub_abstraction"; +val mono2mono_app = thm "mono2mono_app"; +val cont2contlub_app = thm "cont2contlub_app"; +val cont2cont_app = thm "cont2cont_app"; +val cont2cont_app2 = thm "cont2cont_app2"; +val cont_id = thm "cont_id"; +val cont_const = thm "cont_const"; +val cont2cont_app3 = thm "cont2cont_app3"; +val CfunI = thm "CfunI"; +val flatdom2monofun = thm "flatdom2monofun"; +val chfindom_monofun2cont = thm "chfindom_monofun2cont"; +val flatdom_strict2cont = thm "flatdom_strict2cont"; diff -r c899efea601f -r 2454493bd77b src/HOLCF/Cont.thy --- a/src/HOLCF/Cont.thy Wed Mar 02 23:15:16 2005 +0100 +++ b/src/HOLCF/Cont.thy Wed Mar 02 23:28:17 2005 +0100 @@ -1,11 +1,12 @@ (* Title: HOLCF/cont.thy ID: $Id$ Author: Franz Regensburger + License: GPL (GNU GENERAL PUBLIC LICENSE) Results about continuity and monotonicity *) -Cont = Fun3 + +theory Cont = Fun3: (* @@ -15,7 +16,7 @@ *) -default po +defaultsort po consts monofun :: "('a => 'b) => bool" (* monotonicity *) @@ -24,12 +25,12 @@ defs -monofun "monofun(f) == ! x y. x << y --> f(x) << f(y)" +monofun: "monofun(f) == ! x y. x << y --> f(x) << f(y)" -contlub "contlub(f) == ! Y. chain(Y) --> +contlub: "contlub(f) == ! Y. chain(Y) --> f(lub(range(Y))) = lub(range(% i. f(Y(i))))" -cont "cont(f) == ! Y. chain(Y) --> +cont: "cont(f) == ! Y. chain(Y) --> range(% i. f(Y(i))) <<| f(lub(range(Y)))" (* ------------------------------------------------------------------------ *) @@ -37,4 +38,545 @@ (* monofun(f) & contlub(f) <==> cont(f) *) (* ------------------------------------------------------------------------ *) +(* Title: HOLCF/Cont.ML + ID: $Id$ + Author: Franz Regensburger + License: GPL (GNU GENERAL PUBLIC LICENSE) + +Results about continuity and monotonicity +*) + +(* ------------------------------------------------------------------------ *) +(* access to definition *) +(* ------------------------------------------------------------------------ *) + +lemma contlubI: + "! Y. chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))==> + contlub(f)" +apply (unfold contlub) +apply assumption +done + +lemma contlubE: + " contlub(f)==> + ! Y. chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))" +apply (unfold contlub) +apply assumption +done + + +lemma contI: + "! Y. chain(Y) --> range(% i. f(Y(i))) <<| f(lub(range(Y))) ==> cont(f)" + +apply (unfold cont) +apply assumption +done + +lemma contE: + "cont(f) ==> ! Y. chain(Y) --> range(% i. f(Y(i))) <<| f(lub(range(Y)))" +apply (unfold cont) +apply assumption +done + + +lemma monofunI: + "! x y. x << y --> f(x) << f(y) ==> monofun(f)" +apply (unfold monofun) +apply assumption +done + +lemma monofunE: + "monofun(f) ==> ! x y. x << y --> f(x) << f(y)" +apply (unfold monofun) +apply assumption +done + +(* ------------------------------------------------------------------------ *) +(* the main purpose of cont.thy is to show: *) +(* monofun(f) & contlub(f) <==> cont(f) *) +(* ------------------------------------------------------------------------ *) + +(* ------------------------------------------------------------------------ *) +(* monotone functions map chains to chains *) +(* ------------------------------------------------------------------------ *) + +lemma ch2ch_monofun: + "[| monofun(f); chain(Y) |] ==> chain(%i. f(Y(i)))" +apply (rule chainI) +apply (erule monofunE [THEN spec, THEN spec, THEN mp]) +apply (erule chainE) +done + +(* ------------------------------------------------------------------------ *) +(* monotone functions map upper bound to upper bounds *) +(* ------------------------------------------------------------------------ *) + +lemma ub2ub_monofun: + "[| monofun(f); range(Y) <| u|] ==> range(%i. f(Y(i))) <| f(u)" +apply (rule ub_rangeI) +apply (erule monofunE [THEN spec, THEN spec, THEN mp]) +apply (erule ub_rangeD) +done + +(* ------------------------------------------------------------------------ *) +(* left to right: monofun(f) & contlub(f) ==> cont(f) *) +(* ------------------------------------------------------------------------ *) + +lemma monocontlub2cont: + "[|monofun(f);contlub(f)|] ==> cont(f)" +apply (unfold cont) +apply (intro strip) +apply (rule thelubE) +apply (erule ch2ch_monofun) +apply assumption +apply (erule contlubE [THEN spec, THEN mp, symmetric]) +apply assumption +done + +(* ------------------------------------------------------------------------ *) +(* first a lemma about binary chains *) +(* ------------------------------------------------------------------------ *) + +lemma binchain_cont: "[| cont(f); x << y |] + ==> range(%i::nat. f(if i = 0 then x else y)) <<| f(y)" +apply (rule subst) +prefer 2 apply (erule contE [THEN spec, THEN mp]) +apply (erule bin_chain) +apply (rule_tac y = "y" in arg_cong) +apply (erule lub_bin_chain [THEN thelubI]) +done + +(* ------------------------------------------------------------------------ *) +(* right to left: cont(f) ==> monofun(f) & contlub(f) *) +(* part1: cont(f) ==> monofun(f *) +(* ------------------------------------------------------------------------ *) + +lemma cont2mono: "cont(f) ==> monofun(f)" +apply (unfold monofun) +apply (intro strip) +apply (drule binchain_cont [THEN is_ub_lub]) +apply (auto split add: split_if_asm) +done + +(* ------------------------------------------------------------------------ *) +(* right to left: cont(f) ==> monofun(f) & contlub(f) *) +(* part2: cont(f) ==> contlub(f) *) +(* ------------------------------------------------------------------------ *) + +lemma cont2contlub: "cont(f) ==> contlub(f)" +apply (unfold contlub) +apply (intro strip) +apply (rule thelubI [symmetric]) +apply (erule contE [THEN spec, THEN mp]) +apply assumption +done + +(* ------------------------------------------------------------------------ *) +(* monotone functions map finite chains to finite chains *) +(* ------------------------------------------------------------------------ *) + +lemma monofun_finch2finch: + "[| monofun f; finite_chain Y |] ==> finite_chain (%n. f (Y n))" +apply (unfold finite_chain_def) +apply (force elim!: ch2ch_monofun simp add: max_in_chain_def) +done + +(* ------------------------------------------------------------------------ *) +(* The same holds for continuous functions *) +(* ------------------------------------------------------------------------ *) + +lemmas cont_finch2finch = cont2mono [THEN monofun_finch2finch, standard] +(* [| cont ?f; finite_chain ?Y |] ==> finite_chain (%n. ?f (?Y n)) *) + + +(* ------------------------------------------------------------------------ *) +(* The following results are about a curried function that is monotone *) +(* in both arguments *) +(* ------------------------------------------------------------------------ *) + +lemma ch2ch_MF2L: +"[|monofun(MF2); chain(F)|] ==> chain(%i. MF2 (F i) x)" +apply (erule ch2ch_monofun [THEN ch2ch_fun]) +apply assumption +done + + +lemma ch2ch_MF2R: +"[|monofun(MF2(f)); chain(Y)|] ==> chain(%i. MF2 f (Y i))" +apply (erule ch2ch_monofun) +apply assumption +done + +lemma ch2ch_MF2LR: +"[|monofun(MF2); !f. monofun(MF2(f)); chain(F); chain(Y)|] ==> + chain(%i. MF2(F(i))(Y(i)))" +apply (rule chainI) +apply (rule trans_less) +apply (erule ch2ch_MF2L [THEN chainE]) +apply assumption +apply (rule monofunE [THEN spec, THEN spec, THEN mp], erule spec) +apply (erule chainE) +done + + +lemma ch2ch_lubMF2R: +"[|monofun(MF2::('a::po=>'b::po=>'c::cpo)); + !f. monofun(MF2(f)::('b::po=>'c::cpo)); + chain(F);chain(Y)|] ==> + chain(%j. lub(range(%i. MF2 (F j) (Y i))))" +apply (rule lub_mono [THEN chainI]) +apply (rule ch2ch_MF2R, erule spec) +apply assumption +apply (rule ch2ch_MF2R, erule spec) +apply assumption +apply (intro strip) +apply (rule chainE) +apply (erule ch2ch_MF2L) +apply assumption +done + + +lemma ch2ch_lubMF2L: +"[|monofun(MF2::('a::po=>'b::po=>'c::cpo)); + !f. monofun(MF2(f)::('b::po=>'c::cpo)); + chain(F);chain(Y)|] ==> + chain(%i. lub(range(%j. MF2 (F j) (Y i))))" +apply (rule lub_mono [THEN chainI]) +apply (erule ch2ch_MF2L) +apply assumption +apply (erule ch2ch_MF2L) +apply assumption +apply (intro strip) +apply (rule chainE) +apply (rule ch2ch_MF2R, erule spec) +apply assumption +done + + +lemma lub_MF2_mono: +"[|monofun(MF2::('a::po=>'b::po=>'c::cpo)); + !f. monofun(MF2(f)::('b::po=>'c::cpo)); + chain(F)|] ==> + monofun(% x. lub(range(% j. MF2 (F j) (x))))" +apply (rule monofunI) +apply (intro strip) +apply (rule lub_mono) +apply (erule ch2ch_MF2L) +apply assumption +apply (erule ch2ch_MF2L) +apply assumption +apply (intro strip) +apply (rule monofunE [THEN spec, THEN spec, THEN mp], erule spec) +apply assumption +done + +lemma ex_lubMF2: +"[|monofun(MF2::('a::po=>'b::po=>'c::cpo)); + !f. monofun(MF2(f)::('b::po=>'c::cpo)); + chain(F); chain(Y)|] ==> + lub(range(%j. lub(range(%i. MF2(F j) (Y i))))) = + lub(range(%i. lub(range(%j. MF2(F j) (Y i)))))" +apply (rule antisym_less) +apply (rule is_lub_thelub[OF _ ub_rangeI]) +apply (erule ch2ch_lubMF2R) +apply (assumption+) +apply (rule lub_mono) +apply (rule ch2ch_MF2R, erule spec) +apply assumption +apply (erule ch2ch_lubMF2L) +apply (assumption+) +apply (intro strip) +apply (rule is_ub_thelub) +apply (erule ch2ch_MF2L) +apply assumption +apply (rule is_lub_thelub[OF _ ub_rangeI]) +apply (erule ch2ch_lubMF2L) +apply (assumption+) +apply (rule lub_mono) +apply (erule ch2ch_MF2L) +apply assumption +apply (erule ch2ch_lubMF2R) +apply (assumption+) +apply (intro strip) +apply (rule is_ub_thelub) +apply (rule ch2ch_MF2R, erule spec) +apply assumption +done + + +lemma diag_lubMF2_1: +"[|monofun(MF2::('a::po=>'b::po=>'c::cpo)); + !f. monofun(MF2(f)::('b::po=>'c::cpo)); + chain(FY);chain(TY)|] ==> + lub(range(%i. lub(range(%j. MF2(FY(j))(TY(i)))))) = + lub(range(%i. MF2(FY(i))(TY(i))))" +apply (rule antisym_less) +apply (rule is_lub_thelub[OF _ ub_rangeI]) +apply (erule ch2ch_lubMF2L) +apply (assumption+) +apply (rule lub_mono3) +apply (erule ch2ch_MF2L) +apply (assumption+) +apply (erule ch2ch_MF2LR) +apply (assumption+) +apply (rule allI) +apply (rule_tac m = "i" and n = "ia" in nat_less_cases) +apply (rule_tac x = "ia" in exI) +apply (rule chain_mono) +apply (erule allE) +apply (erule ch2ch_MF2R) +apply (assumption+) +apply (erule ssubst) +apply (rule_tac x = "ia" in exI) +apply (rule refl_less) +apply (rule_tac x = "i" in exI) +apply (rule chain_mono) +apply (erule ch2ch_MF2L) +apply (assumption+) +apply (rule lub_mono) +apply (erule ch2ch_MF2LR) +apply (assumption+) +apply (erule ch2ch_lubMF2L) +apply (assumption+) +apply (intro strip) +apply (rule is_ub_thelub) +apply (erule ch2ch_MF2L) +apply assumption +done + +lemma diag_lubMF2_2: +"[|monofun(MF2::('a::po=>'b::po=>'c::cpo)); + !f. monofun(MF2(f)::('b::po=>'c::cpo)); + chain(FY);chain(TY)|] ==> + lub(range(%j. lub(range(%i. MF2(FY(j))(TY(i)))))) = + lub(range(%i. MF2(FY(i))(TY(i))))" +apply (rule trans) +apply (rule ex_lubMF2) +apply (assumption+) +apply (erule diag_lubMF2_1) +apply (assumption+) +done + + +(* ------------------------------------------------------------------------ *) +(* The following results are about a curried function that is continuous *) +(* in both arguments *) +(* ------------------------------------------------------------------------ *) + +lemma contlub_CF2: +assumes prem1: "cont(CF2)" +assumes prem2: "!f. cont(CF2(f))" +assumes prem3: "chain(FY)" +assumes prem4: "chain(TY)" +shows "CF2(lub(range(FY)))(lub(range(TY))) = lub(range(%i. CF2(FY(i))(TY(i))))" +apply (subst prem1 [THEN cont2contlub, THEN contlubE, THEN spec, THEN mp]) +apply assumption +apply (subst thelub_fun) +apply (rule prem1 [THEN cont2mono [THEN ch2ch_monofun]]) +apply assumption +apply (rule trans) +apply (rule prem2 [THEN spec, THEN cont2contlub, THEN contlubE, THEN spec, THEN mp, THEN ext, THEN arg_cong, THEN arg_cong]) +apply (rule prem4) +apply (rule diag_lubMF2_2) +apply (auto simp add: cont2mono prems) +done + +(* ------------------------------------------------------------------------ *) +(* The following results are about application for functions in 'a=>'b *) +(* ------------------------------------------------------------------------ *) + +lemma monofun_fun_fun: "f1 << f2 ==> f1(x) << f2(x)" +apply (erule less_fun [THEN iffD1, THEN spec]) +done + +lemma monofun_fun_arg: "[|monofun(f); x1 << x2|] ==> f(x1) << f(x2)" +apply (erule monofunE [THEN spec, THEN spec, THEN mp]) +apply assumption +done + +lemma monofun_fun: "[|monofun(f1); monofun(f2); f1 << f2; x1 << x2|] ==> f1(x1) << f2(x2)" +apply (rule trans_less) +apply (erule monofun_fun_arg) +apply assumption +apply (erule monofun_fun_fun) +done + + +(* ------------------------------------------------------------------------ *) +(* The following results are about the propagation of monotonicity and *) +(* continuity *) +(* ------------------------------------------------------------------------ *) + +lemma mono2mono_MF1L: "[|monofun(c1)|] ==> monofun(%x. c1 x y)" +apply (rule monofunI) +apply (intro strip) +apply (erule monofun_fun_arg [THEN monofun_fun_fun]) +apply assumption +done + +lemma cont2cont_CF1L: "[|cont(c1)|] ==> cont(%x. c1 x y)" +apply (rule monocontlub2cont) +apply (erule cont2mono [THEN mono2mono_MF1L]) +apply (rule contlubI) +apply (intro strip) +apply (frule asm_rl) +apply (erule cont2contlub [THEN contlubE, THEN spec, THEN mp, THEN ssubst]) +apply assumption +apply (subst thelub_fun) +apply (rule ch2ch_monofun) +apply (erule cont2mono) +apply assumption +apply (rule refl) +done + +(********* Note "(%x.%y.c1 x y) = c1" ***********) + +lemma mono2mono_MF1L_rev: "!y. monofun(%x. c1 x y) ==> monofun(c1)" +apply (rule monofunI) +apply (intro strip) +apply (rule less_fun [THEN iffD2]) +apply (blast dest: monofunE) +done + +lemma cont2cont_CF1L_rev: "!y. cont(%x. c1 x y) ==> cont(c1)" +apply (rule monocontlub2cont) +apply (rule cont2mono [THEN allI, THEN mono2mono_MF1L_rev]) +apply (erule spec) +apply (rule contlubI) +apply (intro strip) +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 (blast dest: cont2contlub [THEN contlubE]) +done + +(* ------------------------------------------------------------------------ *) +(* What D.A.Schmidt calls continuity of abstraction *) +(* never used here *) +(* ------------------------------------------------------------------------ *) + +lemma contlub_abstraction: +"[|chain(Y::nat=>'a);!y. cont(%x.(c::'a::cpo=>'b::cpo=>'c::cpo) x y)|] ==> + (%y. lub(range(%i. c (Y i) y))) = (lub(range(%i.%y. c (Y i) y)))" +apply (rule trans) +prefer 2 apply (rule cont2contlub [THEN contlubE, THEN spec, THEN mp]) +prefer 2 apply (assumption) +apply (erule cont2cont_CF1L_rev) +apply (rule ext) +apply (rule cont2contlub [THEN contlubE, THEN spec, THEN mp, symmetric]) +apply (erule spec) +apply assumption +done + +lemma mono2mono_app: "[|monofun(ft);!x. monofun(ft(x));monofun(tt)|] ==> + monofun(%x.(ft(x))(tt(x)))" +apply (rule monofunI) +apply (intro strip) +apply (rule_tac ?f1.0 = "ft(x)" and ?f2.0 = "ft(y)" in monofun_fun) +apply (erule spec) +apply (erule spec) +apply (erule monofunE [THEN spec, THEN spec, THEN mp]) +apply assumption +apply (erule monofunE [THEN spec, THEN spec, THEN mp]) +apply assumption +done + + +lemma cont2contlub_app: "[|cont(ft);!x. cont(ft(x));cont(tt)|] ==> contlub(%x.(ft(x))(tt(x)))" +apply (rule contlubI) +apply (intro strip) +apply (rule_tac f3 = "tt" in contlubE [THEN spec, THEN mp, THEN ssubst]) +apply (erule cont2contlub) +apply assumption +apply (rule contlub_CF2) +apply (assumption+) +apply (erule cont2mono [THEN ch2ch_monofun]) +apply assumption +done + + +lemma cont2cont_app: "[|cont(ft); !x. cont(ft(x)); cont(tt)|] ==> cont(%x.(ft(x))(tt(x)))" +apply (blast intro: monocontlub2cont mono2mono_app cont2mono cont2contlub_app) +done + + +lemmas cont2cont_app2 = cont2cont_app[OF _ allI] +(* [| cont ?ft; !!x. cont (?ft x); cont ?tt |] ==> *) +(* cont (%x. ?ft x (?tt x)) *) + + +(* ------------------------------------------------------------------------ *) +(* The identity function is continuous *) +(* ------------------------------------------------------------------------ *) + +lemma cont_id: "cont(% x. x)" +apply (rule contI) +apply (intro strip) +apply (erule thelubE) +apply (rule refl) +done + +(* ------------------------------------------------------------------------ *) +(* constant functions are continuous *) +(* ------------------------------------------------------------------------ *) + +lemma cont_const: "cont(%x. c)" +apply (unfold cont) +apply (intro strip) +apply (blast intro: is_lubI ub_rangeI dest: ub_rangeD) +done + + +lemma cont2cont_app3: "[|cont(f); cont(t) |] ==> cont(%x. f(t(x)))" +apply (best intro: cont2cont_app2 cont_const) +done + +(* ------------------------------------------------------------------------ *) +(* A non-emptyness result for Cfun *) +(* ------------------------------------------------------------------------ *) + +lemma CfunI: "?x:Collect cont" +apply (rule CollectI) +apply (rule cont_const) +done + +(* ------------------------------------------------------------------------ *) +(* some properties of flat *) +(* ------------------------------------------------------------------------ *) + +lemma flatdom2monofun: "f UU = UU ==> monofun (f::'a::flat=>'b::pcpo)" + +apply (unfold monofun) +apply (intro strip) +apply (drule ax_flat [THEN spec, THEN spec, THEN mp]) +apply auto +done + +declare range_composition [simp del] +lemma chfindom_monofun2cont: "monofun f ==> cont(f::'a::chfin=>'b::pcpo)" +apply (rule monocontlub2cont) +apply assumption +apply (rule contlubI) +apply (intro strip) +apply (frule chfin2finch) +apply (rule antisym_less) +apply (clarsimp simp add: finite_chain_def maxinch_is_thelub) +apply (rule is_ub_thelub) +apply (erule ch2ch_monofun) +apply assumption +apply (drule monofun_finch2finch[COMP swap_prems_rl]) +apply assumption +apply (simp add: finite_chain_def) +apply (erule conjE) +apply (erule exE) +apply (simp add: maxinch_is_thelub) +apply (erule monofunE [THEN spec, THEN spec, THEN mp]) +apply (erule is_ub_thelub) +done + +lemmas flatdom_strict2cont = flatdom2monofun [THEN chfindom_monofun2cont, standard] +(* f UU = UU ==> cont (f::'a=>'b::pcpo)" *) + end