paulson@9245: (* Title: HOLCF/Cfun2 nipkow@243: ID: $Id$ clasohm@1461: Author: Franz Regensburger nipkow@243: Copyright 1993 Technische Universitaet Muenchen nipkow@243: paulson@9245: Class Instance ->::(cpo,cpo)po nipkow@243: *) nipkow@243: slotosch@2640: (* for compatibility with old HOLCF-Version *) paulson@9245: val prems = goal thy "(op <<)=(%f1 f2. Rep_CFun f1 << Rep_CFun f2)"; paulson@9245: by (fold_goals_tac [less_cfun_def]); paulson@9245: by (rtac refl 1); paulson@9245: qed "inst_cfun_po"; slotosch@2640: nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: (* access to less_cfun in class po *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy "( f1 << f2 ) = (Rep_CFun(f1) << Rep_CFun(f2))"; paulson@9245: by (simp_tac (simpset() addsimps [inst_cfun_po]) 1); paulson@9245: qed "less_cfun"; nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: (* Type 'a ->'b is pointed *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy "Abs_CFun(% x. UU) << f"; paulson@9245: by (stac less_cfun 1); paulson@9245: by (stac Abs_Cfun_inverse2 1); paulson@9245: by (rtac cont_const 1); paulson@9245: by (rtac minimal_fun 1); paulson@9245: qed "minimal_cfun"; nipkow@243: slotosch@2640: bind_thm ("UU_cfun_def",minimal_cfun RS minimal2UU RS sym); slotosch@2640: paulson@9245: val prems = goal thy "? x::'a->'b::pcpo.!y. x< 'b *) slotosch@5291: (* this is continuity of Rep_CFun in its 'second' argument *) slotosch@5291: (* cont_Rep_CFun2 ==> monofun_Rep_CFun2 & contlub_Rep_CFun2 *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy "cont(Rep_CFun(fo))"; paulson@9245: by (res_inst_tac [("P","cont")] CollectD 1); paulson@9245: by (fold_goals_tac [CFun_def]); paulson@9245: by (rtac Rep_Cfun 1); paulson@9245: qed "cont_Rep_CFun2"; nipkow@243: slotosch@5291: bind_thm ("monofun_Rep_CFun2", cont_Rep_CFun2 RS cont2mono); slotosch@5291: (* monofun(Rep_CFun(?fo1)) *) nipkow@243: nipkow@243: slotosch@5291: bind_thm ("contlub_Rep_CFun2", cont_Rep_CFun2 RS cont2contlub); slotosch@5291: (* contlub(Rep_CFun(?fo1)) *) nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) slotosch@5291: (* expanded thms cont_Rep_CFun2, contlub_Rep_CFun2 *) regensbu@1168: (* looks nice with mixfix syntac *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: slotosch@5291: bind_thm ("cont_cfun_arg", (cont_Rep_CFun2 RS contE RS spec RS mp)); oheimb@4721: (* chain(?x1) ==> range (%i. ?fo3`(?x1 i)) <<| ?fo3`(lub (range ?x1)) *) nipkow@243: slotosch@5291: bind_thm ("contlub_cfun_arg", (contlub_Rep_CFun2 RS contlubE RS spec RS mp)); oheimb@4721: (* chain(?x1) ==> ?fo4`(lub (range ?x1)) = lub (range (%i. ?fo4`(?x1 i))) *) nipkow@243: nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) slotosch@5291: (* Rep_CFun is monotone in its 'first' argument *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goalw thy [monofun] "monofun(Rep_CFun)"; paulson@9245: by (strip_tac 1); paulson@9245: by (etac (less_cfun RS subst) 1); paulson@9245: qed "monofun_Rep_CFun1"; nipkow@243: nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) slotosch@5291: (* monotonicity of application Rep_CFun in mixfix syntax [_]_ *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy "f1 << f2 ==> f1`x << f2`x"; paulson@9245: by (cut_facts_tac prems 1); paulson@9245: by (res_inst_tac [("x","x")] spec 1); paulson@9245: by (rtac (less_fun RS subst) 1); paulson@9245: by (etac (monofun_Rep_CFun1 RS monofunE RS spec RS spec RS mp) 1); paulson@9245: qed "monofun_cfun_fun"; nipkow@243: nipkow@243: slotosch@5291: bind_thm ("monofun_cfun_arg", monofun_Rep_CFun2 RS monofunE RS spec RS spec RS mp); regensbu@1168: (* ?x2 << ?x1 ==> ?fo5`?x2 << ?fo5`?x1 *) nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) slotosch@5291: (* monotonicity of Rep_CFun in both arguments in mixfix syntax [_]_ *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy paulson@9245: "[|f1< f1`x1 << f2`x2"; paulson@9245: by (cut_facts_tac prems 1); paulson@9245: by (rtac trans_less 1); paulson@9245: by (etac monofun_cfun_arg 1); paulson@9245: by (etac monofun_cfun_fun 1); paulson@9245: qed "monofun_cfun"; nipkow@243: nipkow@243: paulson@9245: Goal "f`x = UU ==> f`UU = UU"; paulson@9245: by (rtac (eq_UU_iff RS iffD2) 1); paulson@9245: by (etac subst 1); paulson@9245: by (rtac (minimal RS monofun_cfun_arg) 1); paulson@9245: qed "strictI"; oheimb@1989: oheimb@1989: nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: (* ch2ch - rules for the type 'a -> 'b *) nipkow@243: (* use MF2 lemmas from Cont.ML *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy paulson@9245: "chain(Y) ==> chain(%i. f`(Y i))"; paulson@9245: by (cut_facts_tac prems 1); paulson@9245: by (etac (monofun_Rep_CFun2 RS ch2ch_MF2R) 1); paulson@9245: qed "ch2ch_Rep_CFunR"; nipkow@243: nipkow@243: slotosch@5291: bind_thm ("ch2ch_Rep_CFunL", monofun_Rep_CFun1 RS ch2ch_MF2L); oheimb@4721: (* chain(?F) ==> chain (%i. ?F i`?x) *) nipkow@243: nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: (* the lub of a chain of continous functions is monotone *) nipkow@243: (* use MF2 lemmas from Cont.ML *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy paulson@9245: "chain(F) ==> monofun(% x. lub(range(% j.(F j)`x)))"; paulson@9245: by (cut_facts_tac prems 1); paulson@9245: by (rtac lub_MF2_mono 1); paulson@9245: by (rtac monofun_Rep_CFun1 1); paulson@9245: by (rtac (monofun_Rep_CFun2 RS allI) 1); paulson@9245: by (atac 1); paulson@9245: qed "lub_cfun_mono"; nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: (* a lemma about the exchange of lubs for type 'a -> 'b *) nipkow@243: (* use MF2 lemmas from Cont.ML *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy oheimb@4721: "[| chain(F); chain(Y) |] ==>\ clasohm@1461: \ lub(range(%j. lub(range(%i. F(j)`(Y i))))) =\ paulson@9245: \ lub(range(%i. lub(range(%j. F(j)`(Y i)))))"; paulson@9245: by (cut_facts_tac prems 1); paulson@9245: by (rtac ex_lubMF2 1); paulson@9245: by (rtac monofun_Rep_CFun1 1); paulson@9245: by (rtac (monofun_Rep_CFun2 RS allI) 1); paulson@9245: by (atac 1); paulson@9245: by (atac 1); paulson@9245: qed "ex_lubcfun"; nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: (* the lub of a chain of cont. functions is continuous *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy paulson@9245: "chain(F) ==> cont(% x. lub(range(% j. F(j)`x)))"; paulson@9245: by (cut_facts_tac prems 1); paulson@9245: by (rtac monocontlub2cont 1); paulson@9245: by (etac lub_cfun_mono 1); paulson@9245: by (rtac contlubI 1); paulson@9245: by (strip_tac 1); paulson@9245: by (stac (contlub_cfun_arg RS ext) 1); paulson@9245: by (atac 1); paulson@9245: by (etac ex_lubcfun 1); paulson@9245: by (atac 1); paulson@9245: qed "cont_lubcfun"; nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: (* type 'a -> 'b is chain complete *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy paulson@9245: "chain(CCF) ==> range(CCF) <<| (LAM x. lub(range(% i. CCF(i)`x)))"; paulson@9245: by (cut_facts_tac prems 1); paulson@9245: by (rtac is_lubI 1); paulson@9245: by (rtac conjI 1); paulson@9245: by (rtac ub_rangeI 1); paulson@9245: by (rtac allI 1); paulson@9245: by (stac less_cfun 1); paulson@9245: by (stac Abs_Cfun_inverse2 1); paulson@9245: by (etac cont_lubcfun 1); paulson@9245: by (rtac (lub_fun RS is_lubE RS conjunct1 RS ub_rangeE RS spec) 1); paulson@9245: by (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1); paulson@9245: by (strip_tac 1); paulson@9245: by (stac less_cfun 1); paulson@9245: by (stac Abs_Cfun_inverse2 1); paulson@9245: by (etac cont_lubcfun 1); paulson@9245: by (rtac (lub_fun RS is_lubE RS conjunct2 RS spec RS mp) 1); paulson@9245: by (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1); paulson@9245: by (etac (monofun_Rep_CFun1 RS ub2ub_monofun) 1); paulson@9245: qed "lub_cfun"; nipkow@243: oheimb@1779: bind_thm ("thelub_cfun", lub_cfun RS thelubI); nipkow@243: (* oheimb@4721: chain(?CCF1) ==> lub (range ?CCF1) = (LAM x. lub (range (%i. ?CCF1 i`x))) nipkow@243: *) nipkow@243: paulson@9245: val prems = goal thy paulson@9245: "chain(CCF::nat=>('a->'b)) ==> ? x. range(CCF) <<| x"; paulson@9245: by (cut_facts_tac prems 1); paulson@9245: by (rtac exI 1); paulson@9245: by (etac lub_cfun 1); paulson@9245: qed "cpo_cfun"; nipkow@243: nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: (* Extensionality in 'a -> 'b *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal Cfun1.thy "(!!x. f`x = g`x) ==> f = g"; paulson@9245: by (res_inst_tac [("t","f")] (Rep_Cfun_inverse RS subst) 1); paulson@9245: by (res_inst_tac [("t","g")] (Rep_Cfun_inverse RS subst) 1); paulson@9245: by (res_inst_tac [("f","Abs_CFun")] arg_cong 1); paulson@9245: by (rtac ext 1); paulson@9245: by (resolve_tac prems 1); paulson@9245: qed "ext_cfun"; nipkow@243: nipkow@243: (* ------------------------------------------------------------------------ *) slotosch@5291: (* Monotonicity of Abs_CFun *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy paulson@9245: "[|cont(f);cont(g);f<Abs_CFun(f)< 'b *) nipkow@243: (* ------------------------------------------------------------------------ *) nipkow@243: paulson@9245: val prems = goal thy "(!!x. f`x << g`x) ==> f << g"; paulson@9245: by (res_inst_tac [("t","f")] (Rep_Cfun_inverse RS subst) 1); paulson@9245: by (res_inst_tac [("t","g")] (Rep_Cfun_inverse RS subst) 1); paulson@9245: by (rtac semi_monofun_Abs_CFun 1); paulson@9245: by (rtac cont_Rep_CFun2 1); paulson@9245: by (rtac cont_Rep_CFun2 1); paulson@9245: by (rtac (less_fun RS iffD2) 1); paulson@9245: by (rtac allI 1); paulson@9245: by (resolve_tac prems 1); paulson@9245: qed "less_cfun2"; nipkow@243: nipkow@243: