# HG changeset patch # User huffman # Date 1109804282 -3600 # Node ID eb3b1a5c304e65250b58c993801f5ff6d7f37636 # Parent 2454493bd77bc5a0be33f2dab3275c5bf96e3ed5 converted to new-style theory diff -r 2454493bd77b -r eb3b1a5c304e src/HOLCF/Cfun1.ML --- a/src/HOLCF/Cfun1.ML Wed Mar 02 23:28:17 2005 +0100 +++ b/src/HOLCF/Cfun1.ML Wed Mar 02 23:58:02 2005 +0100 @@ -1,90 +1,16 @@ -(* Title: HOLCF/Cfun1.ML - ID: $Id$ - Author: Franz Regensburger -The type -> of continuous functions. -*) - -(* ------------------------------------------------------------------------ *) -(* derive old type definition rules for Abs_CFun & Rep_CFun *) -(* Rep_CFun and Abs_CFun should be replaced by Rep_Cfun anf Abs_Cfun in future *) -(* ------------------------------------------------------------------------ *) -Goal "Rep_CFun fo : CFun"; -by (rtac Rep_CFun 1); -qed "Rep_Cfun"; - -Goal "Abs_CFun (Rep_CFun fo) = fo"; -by (rtac Rep_CFun_inverse 1); -qed "Rep_Cfun_inverse"; - -Goal "f:CFun==>Rep_CFun(Abs_CFun f)=f"; -by (etac Abs_CFun_inverse 1); -qed "Abs_Cfun_inverse"; - -(* ------------------------------------------------------------------------ *) -(* less_cfun is a partial order on type 'a -> 'b *) -(* ------------------------------------------------------------------------ *) - -Goalw [less_cfun_def] "(f::'a->'b) << f"; -by (rtac refl_less 1); -qed "refl_less_cfun"; - -Goalw [less_cfun_def] - "[|(f1::'a->'b) << f2; f2 << f1|] ==> f1 = f2"; -by (rtac injD 1); -by (rtac antisym_less 2); -by (atac 3); -by (atac 2); -by (rtac inj_inverseI 1); -by (rtac Rep_Cfun_inverse 1); -qed "antisym_less_cfun"; +(* legacy ML bindings *) -Goalw [less_cfun_def] - "[|(f1::'a->'b) << f2; f2 << f3|] ==> f1 << f3"; -by (etac trans_less 1); -by (atac 1); -qed "trans_less_cfun"; - -(* ------------------------------------------------------------------------ *) -(* lemmas about application of continuous functions *) -(* ------------------------------------------------------------------------ *) - -Goal "[| f=g; x=y |] ==> f$x = g$y"; -by (Asm_simp_tac 1); -qed "cfun_cong"; - -Goal "f=g ==> f$x = g$x"; -by (Asm_simp_tac 1); -qed "cfun_fun_cong"; - -Goal "x=y ==> f$x = f$y"; -by (Asm_simp_tac 1); -qed "cfun_arg_cong"; - - -(* ------------------------------------------------------------------------ *) -(* additional lemma about the isomorphism between -> and Cfun *) -(* ------------------------------------------------------------------------ *) - -Goal "cont f ==> Rep_CFun (Abs_CFun f) = f"; -by (rtac Abs_Cfun_inverse 1); -by (rewtac CFun_def); -by (etac (mem_Collect_eq RS ssubst) 1); -qed "Abs_Cfun_inverse2"; - -(* ------------------------------------------------------------------------ *) -(* simplification of application *) -(* ------------------------------------------------------------------------ *) - -Goal "cont f ==> (Abs_CFun f)$x = f x"; -by (etac (Abs_Cfun_inverse2 RS fun_cong) 1); -qed "Cfunapp2"; - -(* ------------------------------------------------------------------------ *) -(* beta - equality for continuous functions *) -(* ------------------------------------------------------------------------ *) - -Goal "cont(c1) ==> (LAM x .c1 x)$u = c1 u"; -by (rtac Cfunapp2 1); -by (atac 1); -qed "beta_cfun"; +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 refl_less_cfun = thm "refl_less_cfun"; +val antisym_less_cfun = thm "antisym_less_cfun"; +val trans_less_cfun = thm "trans_less_cfun"; +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 beta_cfun = thm "beta_cfun"; diff -r 2454493bd77b -r eb3b1a5c304e src/HOLCF/Cfun1.thy --- a/src/HOLCF/Cfun1.thy Wed Mar 02 23:28:17 2005 +0100 +++ b/src/HOLCF/Cfun1.thy Wed Mar 02 23:58:02 2005 +0100 @@ -1,19 +1,21 @@ (* Title: HOLCF/Cfun1.thy ID: $Id$ Author: Franz Regensburger + License: GPL (GNU GENERAL PUBLIC LICENSE) Definition of the type -> of continuous functions. *) -Cfun1 = Cont + +theory Cfun1 = Cont: -default cpo +defaultsort cpo -typedef (CFun) ('a, 'b) "->" (infixr 0) = "{f::'a => 'b. cont f}" (CfunI) +typedef (CFun) ('a, 'b) "->" (infixr 0) = "{f::'a => 'b. cont f}" +by (rule exI, rule CfunI) (* to make << defineable *) -instance "->" :: (cpo,cpo)sq_ord +instance "->" :: (cpo,cpo)sq_ord .. syntax Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("_$_" [999,1000] 999) @@ -23,15 +25,114 @@ less_cfun :: "[('a -> 'b),('a -> 'b)]=>bool" syntax (xsymbols) - "->" :: [type, type] => type ("(_ \\/ _)" [1,0]0) + "->" :: "[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) + ("(3\_./ _)" [0, 10] 10) + Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\_)" [999,1000] 999) syntax (HTML output) - Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\\_)" [999,1000] 999) + Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\_)" [999,1000] 999) + +defs (overloaded) + less_cfun_def: "(op <<) == (% fo1 fo2. Rep_CFun fo1 << Rep_CFun fo2 )" + +(* Title: HOLCF/Cfun1.ML + ID: $Id$ + Author: Franz Regensburger + License: GPL (GNU GENERAL PUBLIC LICENSE) + +The type -> of continuous functions. +*) + +(* ------------------------------------------------------------------------ *) +(* derive old type definition rules for Abs_CFun & Rep_CFun + *) +(* Rep_CFun and Abs_CFun should be replaced by Rep_Cfun anf Abs_Cfun in future + *) +(* ------------------------------------------------------------------------ *) + +lemma Rep_Cfun: "Rep_CFun fo : CFun" +apply (rule Rep_CFun) +done + +lemma Rep_Cfun_inverse: "Abs_CFun (Rep_CFun fo) = fo" +apply (rule Rep_CFun_inverse) +done + +lemma Abs_Cfun_inverse: "f:CFun==>Rep_CFun(Abs_CFun f)=f" +apply (erule Abs_CFun_inverse) +done + +(* ------------------------------------------------------------------------ *) +(* less_cfun is a partial order on type 'a -> 'b *) +(* ------------------------------------------------------------------------ *) + +lemma refl_less_cfun: "(f::'a->'b) << f" + +apply (unfold less_cfun_def) +apply (rule refl_less) +done -defs - less_cfun_def "(op <<) == (% fo1 fo2. Rep_CFun fo1 << Rep_CFun fo2 )" +lemma antisym_less_cfun: + "[|(f1::'a->'b) << f2; f2 << f1|] ==> f1 = f2" +apply (unfold less_cfun_def) +apply (rule injD) +apply (rule_tac [2] antisym_less) +prefer 3 apply (assumption) +prefer 2 apply (assumption) +apply (rule inj_on_inverseI) +apply (rule Rep_Cfun_inverse) +done + +lemma trans_less_cfun: + "[|(f1::'a->'b) << f2; f2 << f3|] ==> f1 << f3" +apply (unfold less_cfun_def) +apply (erule trans_less) +apply assumption +done + +(* ------------------------------------------------------------------------ *) +(* lemmas about application of continuous functions *) +(* ------------------------------------------------------------------------ *) + +lemma cfun_cong: "[| f=g; x=y |] ==> f$x = g$y" +apply (simp (no_asm_simp)) +done + +lemma cfun_fun_cong: "f=g ==> f$x = g$x" +apply (simp (no_asm_simp)) +done + +lemma cfun_arg_cong: "x=y ==> f$x = f$y" +apply (simp (no_asm_simp)) +done + + +(* ------------------------------------------------------------------------ *) +(* additional lemma about the isomorphism between -> and 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 + +(* ------------------------------------------------------------------------ *) +(* simplification of application *) +(* ------------------------------------------------------------------------ *) + +lemma Cfunapp2: "cont f ==> (Abs_CFun f)$x = f x" +apply (erule Abs_Cfun_inverse2 [THEN fun_cong]) +done + +(* ------------------------------------------------------------------------ *) +(* beta - equality for continuous functions *) +(* ------------------------------------------------------------------------ *) + +lemma beta_cfun: "cont(c1) ==> (LAM x .c1 x)$u = c1 u" +apply (rule Cfunapp2) +apply assumption +done end diff -r 2454493bd77b -r eb3b1a5c304e src/HOLCF/Cfun2.ML --- a/src/HOLCF/Cfun2.ML Wed Mar 02 23:28:17 2005 +0100 +++ b/src/HOLCF/Cfun2.ML Wed Mar 02 23:58:02 2005 +0100 @@ -1,247 +1,30 @@ -(* Title: HOLCF/Cfun2 - ID: $Id$ - Author: Franz Regensburger -Class Instance ->::(cpo,cpo)po -*) - -(* for compatibility with old HOLCF-Version *) -Goal "(op <<)=(%f1 f2. Rep_CFun f1 << Rep_CFun f2)"; -by (fold_goals_tac [less_cfun_def]); -by (rtac refl 1); -qed "inst_cfun_po"; - -(* ------------------------------------------------------------------------ *) -(* access to less_cfun in class po *) -(* ------------------------------------------------------------------------ *) - -Goal "( f1 << f2 ) = (Rep_CFun(f1) << Rep_CFun(f2))"; -by (simp_tac (simpset() addsimps [inst_cfun_po]) 1); -qed "less_cfun"; - -(* ------------------------------------------------------------------------ *) -(* Type 'a ->'b is pointed *) -(* ------------------------------------------------------------------------ *) - -Goal "Abs_CFun(% x. UU) << f"; -by (stac less_cfun 1); -by (stac Abs_Cfun_inverse2 1); -by (rtac cont_const 1); -by (rtac minimal_fun 1); -qed "minimal_cfun"; - -bind_thm ("UU_cfun_def",minimal_cfun RS minimal2UU RS sym); - -Goal "? x::'a->'b::pcpo.!y. x< 'b *) -(* this is continuity of Rep_CFun in its 'second' argument *) -(* cont_Rep_CFun2 ==> monofun_Rep_CFun2 & contlub_Rep_CFun2 *) -(* ------------------------------------------------------------------------ *) - -Goal "cont(Rep_CFun(fo))"; -by (res_inst_tac [("P","cont")] CollectD 1); -by (fold_goals_tac [CFun_def]); -by (rtac Rep_Cfun 1); -qed "cont_Rep_CFun2"; - -bind_thm ("monofun_Rep_CFun2", cont_Rep_CFun2 RS cont2mono); -(* monofun(Rep_CFun(?fo1)) *) - - -bind_thm ("contlub_Rep_CFun2", cont_Rep_CFun2 RS cont2contlub); -(* contlub(Rep_CFun(?fo1)) *) - -(* ------------------------------------------------------------------------ *) -(* expanded thms cont_Rep_CFun2, contlub_Rep_CFun2 *) -(* looks nice with mixfix syntac *) -(* ------------------------------------------------------------------------ *) - -bind_thm ("cont_cfun_arg", (cont_Rep_CFun2 RS contE RS spec RS mp)); -(* chain(?x1) ==> range (%i. ?fo3$(?x1 i)) <<| ?fo3$(lub (range ?x1)) *) - -bind_thm ("contlub_cfun_arg", (contlub_Rep_CFun2 RS contlubE RS spec RS mp)); -(* chain(?x1) ==> ?fo4$(lub (range ?x1)) = lub (range (%i. ?fo4$(?x1 i))) *) - - -(* ------------------------------------------------------------------------ *) -(* Rep_CFun is monotone in its 'first' argument *) -(* ------------------------------------------------------------------------ *) - -Goalw [monofun] "monofun(Rep_CFun)"; -by (strip_tac 1); -by (etac (less_cfun RS subst) 1); -qed "monofun_Rep_CFun1"; - - -(* ------------------------------------------------------------------------ *) -(* monotonicity of application Rep_CFun in mixfix syntax [_]_ *) -(* ------------------------------------------------------------------------ *) - -Goal "f1 << f2 ==> f1$x << f2$x"; -by (res_inst_tac [("x","x")] spec 1); -by (rtac (less_fun RS subst) 1); -by (etac (monofun_Rep_CFun1 RS monofunE RS spec RS spec RS mp) 1); -qed "monofun_cfun_fun"; - - -bind_thm ("monofun_cfun_arg", monofun_Rep_CFun2 RS monofunE RS spec RS spec RS mp); -(* ?x2 << ?x1 ==> ?fo5$?x2 << ?fo5$?x1 *) - -Goal "chain Y ==> chain (%i. f\\(Y i))"; -by (rtac chainI 1); -by (rtac monofun_cfun_arg 1); -by (etac chainE 1); -qed "chain_monofun"; - - -(* ------------------------------------------------------------------------ *) -(* monotonicity of Rep_CFun in both arguments in mixfix syntax [_]_ *) -(* ------------------------------------------------------------------------ *) - -Goal "[|f1< f1$x1 << f2$x2"; -by (rtac trans_less 1); -by (etac monofun_cfun_arg 1); -by (etac monofun_cfun_fun 1); -qed "monofun_cfun"; - - -Goal "f$x = UU ==> f$UU = UU"; -by (rtac (eq_UU_iff RS iffD2) 1); -by (etac subst 1); -by (rtac (minimal RS monofun_cfun_arg) 1); -qed "strictI"; - - -(* ------------------------------------------------------------------------ *) -(* ch2ch - rules for the type 'a -> 'b *) -(* use MF2 lemmas from Cont.ML *) -(* ------------------------------------------------------------------------ *) +(* legacy ML bindings *) -Goal "chain(Y) ==> chain(%i. f$(Y i))"; -by (etac (monofun_Rep_CFun2 RS ch2ch_MF2R) 1); -qed "ch2ch_Rep_CFunR"; - - -bind_thm ("ch2ch_Rep_CFunL", monofun_Rep_CFun1 RS ch2ch_MF2L); -(* chain(?F) ==> chain (%i. ?F i$?x) *) - - -(* ------------------------------------------------------------------------ *) -(* the lub of a chain of continous functions is monotone *) -(* use MF2 lemmas from Cont.ML *) -(* ------------------------------------------------------------------------ *) - -Goal "chain(F) ==> monofun(% x. lub(range(% j.(F j)$x)))"; -by (rtac lub_MF2_mono 1); -by (rtac monofun_Rep_CFun1 1); -by (rtac (monofun_Rep_CFun2 RS allI) 1); -by (atac 1); -qed "lub_cfun_mono"; - -(* ------------------------------------------------------------------------ *) -(* a lemma about the exchange of lubs for type 'a -> 'b *) -(* use MF2 lemmas from Cont.ML *) -(* ------------------------------------------------------------------------ *) - -Goal "[| chain(F); chain(Y) |] ==>\ -\ lub(range(%j. lub(range(%i. F(j)$(Y i))))) =\ -\ lub(range(%i. lub(range(%j. F(j)$(Y i)))))"; -by (rtac ex_lubMF2 1); -by (rtac monofun_Rep_CFun1 1); -by (rtac (monofun_Rep_CFun2 RS allI) 1); -by (atac 1); -by (atac 1); -qed "ex_lubcfun"; - -(* ------------------------------------------------------------------------ *) -(* the lub of a chain of cont. functions is continuous *) -(* ------------------------------------------------------------------------ *) - -Goal "chain(F) ==> cont(% x. lub(range(% j. F(j)$x)))"; -by (rtac monocontlub2cont 1); -by (etac lub_cfun_mono 1); -by (rtac contlubI 1); -by (strip_tac 1); -by (stac (contlub_cfun_arg RS ext) 1); -by (atac 1); -by (etac ex_lubcfun 1); -by (atac 1); -qed "cont_lubcfun"; - -(* ------------------------------------------------------------------------ *) -(* type 'a -> 'b is chain complete *) -(* ------------------------------------------------------------------------ *) - -Goal "chain(CCF) ==> range(CCF) <<| (LAM x. lub(range(% i. CCF(i)$x)))"; -by (rtac is_lubI 1); -by (rtac ub_rangeI 1); -by (stac less_cfun 1); -by (stac Abs_Cfun_inverse2 1); -by (etac cont_lubcfun 1); -by (rtac (lub_fun RS is_lubD1 RS ub_rangeD) 1); -by (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1); -by (stac less_cfun 1); -by (stac Abs_Cfun_inverse2 1); -by (etac cont_lubcfun 1); -by (rtac (lub_fun RS is_lub_lub) 1); -by (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1); -by (etac (monofun_Rep_CFun1 RS ub2ub_monofun) 1); -qed "lub_cfun"; - -bind_thm ("thelub_cfun", lub_cfun RS thelubI); -(* -chain(?CCF1) ==> lub (range ?CCF1) = (LAM x. lub (range (%i. ?CCF1 i$x))) -*) - -Goal "chain(CCF::nat=>('a->'b)) ==> ? x. range(CCF) <<| x"; -by (rtac exI 1); -by (etac lub_cfun 1); -qed "cpo_cfun"; - - -(* ------------------------------------------------------------------------ *) -(* Extensionality in 'a -> 'b *) -(* ------------------------------------------------------------------------ *) - -val prems = Goal "(!!x. f$x = g$x) ==> f = g"; -by (res_inst_tac [("t","f")] (Rep_Cfun_inverse RS subst) 1); -by (res_inst_tac [("t","g")] (Rep_Cfun_inverse RS subst) 1); -by (res_inst_tac [("f","Abs_CFun")] arg_cong 1); -by (rtac ext 1); -by (resolve_tac prems 1); -qed "ext_cfun"; - -(* ------------------------------------------------------------------------ *) -(* Monotonicity of Abs_CFun *) -(* ------------------------------------------------------------------------ *) - -Goal "[| cont(f); cont(g); f< Abs_CFun(f)< 'b *) -(* ------------------------------------------------------------------------ *) - -val prems = Goal "(!!x. f$x << g$x) ==> f << g"; -by (res_inst_tac [("t","f")] (Rep_Cfun_inverse RS subst) 1); -by (res_inst_tac [("t","g")] (Rep_Cfun_inverse RS subst) 1); -by (rtac semi_monofun_Abs_CFun 1); -by (rtac cont_Rep_CFun2 1); -by (rtac cont_Rep_CFun2 1); -by (rtac (less_fun RS iffD2) 1); -by (rtac allI 1); -by (resolve_tac prems 1); -qed "less_cfun2"; - - +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"; +val cont_cfun_arg = thm "cont_cfun_arg"; +val contlub_cfun_arg = thm "contlub_cfun_arg"; +val monofun_Rep_CFun1 = thm "monofun_Rep_CFun1"; +val monofun_cfun_fun = thm "monofun_cfun_fun"; +val monofun_cfun_arg = thm "monofun_cfun_arg"; +val chain_monofun = thm "chain_monofun"; +val monofun_cfun = thm "monofun_cfun"; +val strictI = thm "strictI"; +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 lub_cfun = thm "lub_cfun"; +val thelub_cfun = thm "thelub_cfun"; +val cpo_cfun = thm "cpo_cfun"; +val ext_cfun = thm "ext_cfun"; +val semi_monofun_Abs_CFun = thm "semi_monofun_Abs_CFun"; +val less_cfun2 = thm "less_cfun2"; diff -r 2454493bd77b -r eb3b1a5c304e src/HOLCF/Cfun2.thy --- a/src/HOLCF/Cfun2.thy Wed Mar 02 23:28:17 2005 +0100 +++ b/src/HOLCF/Cfun2.thy Wed Mar 02 23:58:02 2005 +0100 @@ -1,13 +1,267 @@ (* Title: HOLCF/Cfun2.thy ID: $Id$ Author: Franz Regensburger + License: GPL (GNU GENERAL PUBLIC LICENSE) Class Instance ->::(cpo,cpo)po *) -Cfun2 = Cfun1 + +theory Cfun2 = Cfun1: + +instance "->"::(cpo,cpo)po +apply (intro_classes) +apply (rule refl_less_cfun) +apply (rule antisym_less_cfun, assumption+) +apply (rule trans_less_cfun, assumption+) +done + +(* Title: HOLCF/Cfun2 + ID: $Id$ + Author: Franz Regensburger + License: GPL (GNU GENERAL PUBLIC LICENSE) + +Class Instance ->::(cpo,cpo)po +*) + +(* 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 + +(* ------------------------------------------------------------------------ *) +(* access to less_cfun in class po *) +(* ------------------------------------------------------------------------ *) + +lemma less_cfun: "( f1 << f2 ) = (Rep_CFun(f1) << Rep_CFun(f2))" +apply (simp (no_asm) add: inst_cfun_po) +done + +(* ------------------------------------------------------------------------ *) +(* Type 'a ->'b is pointed *) +(* ------------------------------------------------------------------------ *) + +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< 'b *) +(* this is continuity of Rep_CFun in its 'second' argument *) +(* cont_Rep_CFun2 ==> monofun_Rep_CFun2 & contlub_Rep_CFun2 *) +(* ------------------------------------------------------------------------ *) + +lemma cont_Rep_CFun2: "cont(Rep_CFun(fo))" +apply (rule_tac P = "cont" in CollectD) +apply (fold CFun_def) +apply (rule Rep_Cfun) +done + +lemmas monofun_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2mono, standard] +(* monofun(Rep_CFun(?fo1)) *) + + +lemmas contlub_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2contlub, standard] +(* contlub(Rep_CFun(?fo1)) *) + +(* ------------------------------------------------------------------------ *) +(* expanded thms cont_Rep_CFun2, contlub_Rep_CFun2 *) +(* looks nice with mixfix syntac *) +(* ------------------------------------------------------------------------ *) + +lemmas cont_cfun_arg = cont_Rep_CFun2 [THEN contE, THEN spec, THEN mp] +(* chain(?x1) ==> range (%i. ?fo3$(?x1 i)) <<| ?fo3$(lub (range ?x1)) *) + +lemmas contlub_cfun_arg = contlub_Rep_CFun2 [THEN contlubE, THEN spec, THEN mp] +(* chain(?x1) ==> ?fo4$(lub (range ?x1)) = lub (range (%i. ?fo4$(?x1 i))) *) + + +(* ------------------------------------------------------------------------ *) +(* Rep_CFun is monotone in its 'first' argument *) +(* ------------------------------------------------------------------------ *) + +lemma monofun_Rep_CFun1: "monofun(Rep_CFun)" +apply (unfold monofun) +apply (intro strip) +apply (erule less_cfun [THEN subst]) +done + + +(* ------------------------------------------------------------------------ *) +(* monotonicity of application Rep_CFun in mixfix syntax [_]_ *) +(* ------------------------------------------------------------------------ *) + +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, THEN spec, THEN spec, THEN mp]) +done + + +lemmas monofun_cfun_arg = monofun_Rep_CFun2 [THEN monofunE, THEN spec, THEN spec, THEN mp, standard] +(* ?x2 << ?x1 ==> ?fo5$?x2 << ?fo5$?x1 *) + +lemma chain_monofun: "chain Y ==> chain (%i. f\(Y i))" +apply (rule chainI) +apply (rule monofun_cfun_arg) +apply (erule chainE) +done + + +(* ------------------------------------------------------------------------ *) +(* monotonicity of Rep_CFun in both arguments in mixfix syntax [_]_ *) +(* ------------------------------------------------------------------------ *) + +lemma monofun_cfun: "[|f1< f1$x1 << f2$x2" +apply (rule trans_less) +apply (erule monofun_cfun_arg) +apply (erule monofun_cfun_fun) +done + -instance "->"::(cpo,cpo)po (refl_less_cfun,antisym_less_cfun,trans_less_cfun) +lemma strictI: "f$x = UU ==> f$UU = UU" +apply (rule eq_UU_iff [THEN iffD2]) +apply (erule subst) +apply (rule minimal [THEN monofun_cfun_arg]) +done + + +(* ------------------------------------------------------------------------ *) +(* ch2ch - rules for the type 'a -> 'b *) +(* use MF2 lemmas from Cont.ML *) +(* ------------------------------------------------------------------------ *) + +lemma ch2ch_Rep_CFunR: "chain(Y) ==> chain(%i. f$(Y i))" +apply (erule monofun_Rep_CFun2 [THEN ch2ch_MF2R]) +done + + +lemmas ch2ch_Rep_CFunL = monofun_Rep_CFun1 [THEN ch2ch_MF2L, standard] +(* chain(?F) ==> chain (%i. ?F i$?x) *) + + +(* ------------------------------------------------------------------------ *) +(* the lub of a chain of continous functions is monotone *) +(* use MF2 lemmas from Cont.ML *) +(* ------------------------------------------------------------------------ *) + +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 +done + +(* ------------------------------------------------------------------------ *) +(* a lemma about the exchange of lubs for type 'a -> 'b *) +(* use MF2 lemmas from Cont.ML *) +(* ------------------------------------------------------------------------ *) + +lemma ex_lubcfun: "[| 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 + +(* ------------------------------------------------------------------------ *) +(* 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) +apply (intro strip) +apply (subst contlub_cfun_arg [THEN ext]) +apply assumption +apply (erule ex_lubcfun) +apply assumption +done + +(* ------------------------------------------------------------------------ *) +(* type '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]) +done + +lemmas thelub_cfun = lub_cfun [THEN thelubI, standard] +(* +chain(?CCF1) ==> lub (range ?CCF1) = (LAM x. lub (range (%i. ?CCF1 i$x))) +*) + +lemma cpo_cfun: "chain(CCF::nat=>('a->'b)) ==> ? x. range(CCF) <<| x" +apply (rule exI) +apply (erule lub_cfun) +done + + +(* ------------------------------------------------------------------------ *) +(* Extensionality in 'a -> 'b *) +(* ------------------------------------------------------------------------ *) + +lemma ext_cfun: "(!!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_tac f = "Abs_CFun" in arg_cong) +apply (rule ext) +apply simp +done + +(* ------------------------------------------------------------------------ *) +(* Monotonicity of 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 (rule allI) +apply simp +done end diff -r 2454493bd77b -r eb3b1a5c304e src/HOLCF/Cfun3.ML --- a/src/HOLCF/Cfun3.ML Wed Mar 02 23:28:17 2005 +0100 +++ b/src/HOLCF/Cfun3.ML Wed Mar 02 23:58:02 2005 +0100 @@ -1,498 +1,52 @@ -(* Title: HOLCF/Cfun3 - ID: $Id$ - Author: Franz Regensburger -Class instance of -> for class pcpo -*) - -(* for compatibility with old HOLCF-Version *) -Goal "UU = Abs_CFun(%x. UU)"; -by (simp_tac (HOL_ss addsimps [UU_def,UU_cfun_def]) 1); -qed "inst_cfun_pcpo"; - -(* ------------------------------------------------------------------------ *) -(* the contlub property for Rep_CFun its 'first' argument *) -(* ------------------------------------------------------------------------ *) - -Goal "contlub(Rep_CFun)"; -by (rtac contlubI 1); -by (strip_tac 1); -by (rtac (expand_fun_eq RS iffD2) 1); -by (strip_tac 1); -by (stac thelub_cfun 1); -by (atac 1); -by (stac Cfunapp2 1); -by (etac cont_lubcfun 1); -by (stac thelub_fun 1); -by (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1); -by (rtac refl 1); -qed "contlub_Rep_CFun1"; - - -(* ------------------------------------------------------------------------ *) -(* the cont property for Rep_CFun in its first argument *) -(* ------------------------------------------------------------------------ *) - -Goal "cont(Rep_CFun)"; -by (rtac monocontlub2cont 1); -by (rtac monofun_Rep_CFun1 1); -by (rtac contlub_Rep_CFun1 1); -qed "cont_Rep_CFun1"; - - -(* ------------------------------------------------------------------------ *) -(* contlub, cont properties of Rep_CFun in its first argument in mixfix _[_] *) -(* ------------------------------------------------------------------------ *) - -Goal -"chain(FY) ==>\ -\ lub(range FY)$x = lub(range (%i. FY(i)$x))"; -by (rtac trans 1); -by (etac (contlub_Rep_CFun1 RS contlubE RS spec RS mp RS fun_cong) 1); -by (stac thelub_fun 1); -by (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1); -by (rtac refl 1); -qed "contlub_cfun_fun"; - - -Goal -"chain(FY) ==>\ -\ range(%i. FY(i)$x) <<| lub(range FY)$x"; -by (rtac thelubE 1); -by (etac ch2ch_Rep_CFunL 1); -by (etac (contlub_cfun_fun RS sym) 1); -qed "cont_cfun_fun"; - - -(* ------------------------------------------------------------------------ *) -(* contlub, cont properties of Rep_CFun in both argument in mixfix _[_] *) -(* ------------------------------------------------------------------------ *) - -Goal -"[|chain(FY);chain(TY)|] ==>\ -\ (lub(range FY))$(lub(range TY)) = lub(range(%i. FY(i)$(TY i)))"; -by (rtac contlub_CF2 1); -by (rtac cont_Rep_CFun1 1); -by (rtac allI 1); -by (rtac cont_Rep_CFun2 1); -by (atac 1); -by (atac 1); -qed "contlub_cfun"; - -Goal -"[|chain(FY);chain(TY)|] ==>\ -\ range(%i.(FY i)$(TY i)) <<| (lub (range FY))$(lub(range TY))"; -by (rtac thelubE 1); -by (rtac (monofun_Rep_CFun1 RS ch2ch_MF2LR) 1); -by (rtac allI 1); -by (rtac monofun_Rep_CFun2 1); -by (atac 1); -by (atac 1); -by (etac (contlub_cfun RS sym) 1); -by (atac 1); -qed "cont_cfun"; - - -(* ------------------------------------------------------------------------ *) -(* cont2cont lemma for Rep_CFun *) -(* ------------------------------------------------------------------------ *) - -Goal "[|cont(%x. ft x);cont(%x. tt x)|] ==> cont(%x. (ft x)$(tt x))"; -by (best_tac (claset() addIs [cont2cont_app2, cont_const, cont_Rep_CFun1, - cont_Rep_CFun2]) 1); -qed "cont2cont_Rep_CFun"; - - - -(* ------------------------------------------------------------------------ *) -(* cont2mono Lemma for %x. LAM y. c1(x)(y) *) -(* ------------------------------------------------------------------------ *) - -val [p1,p2] = Goal - "[| !!x. cont(c1 x); !!y. monofun(%x. c1 x y)|] ==> monofun(%x. LAM y. c1 x y)"; -by (rtac monofunI 1); -by (strip_tac 1); -by (stac less_cfun 1); -by (stac less_fun 1); -by (rtac allI 1); -by (stac beta_cfun 1); -by (rtac p1 1); -by (stac beta_cfun 1); -by (rtac p1 1); -by (etac (p2 RS monofunE RS spec RS spec RS mp) 1); -qed "cont2mono_LAM"; - -(* ------------------------------------------------------------------------ *) -(* cont2cont Lemma for %x. LAM y. c1 x y) *) -(* ------------------------------------------------------------------------ *) - -val [p1,p2] = Goal - "[| !!x. cont(c1 x); !!y. cont(%x. c1 x y) |] ==> cont(%x. LAM y. c1 x y)"; -by (rtac monocontlub2cont 1); -by (rtac (p1 RS cont2mono_LAM) 1); -by (rtac (p2 RS cont2mono) 1); -by (rtac contlubI 1); -by (strip_tac 1); -by (stac thelub_cfun 1); -by (rtac (p1 RS cont2mono_LAM RS ch2ch_monofun) 1); -by (rtac (p2 RS cont2mono) 1); -by (atac 1); -by (res_inst_tac [("f","Abs_CFun")] arg_cong 1); -by (rtac ext 1); -by (stac (p1 RS beta_cfun RS ext) 1); -by (etac (p2 RS cont2contlub RS contlubE RS spec RS mp) 1); -qed "cont2cont_LAM"; - -(* ------------------------------------------------------------------------ *) -(* cont2cont tactic *) -(* ------------------------------------------------------------------------ *) - -val cont_lemmas1 = [cont_const, cont_id, cont_Rep_CFun2, - cont2cont_Rep_CFun,cont2cont_LAM]; - -Addsimps cont_lemmas1; - -(* HINT: 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)));*) - -(* ------------------------------------------------------------------------ *) -(* function application _[_] is strict in its first arguments *) -(* ------------------------------------------------------------------------ *) - -Goal "(UU::'a::cpo->'b)$x = (UU::'b)"; -by (stac inst_cfun_pcpo 1); -by (stac beta_cfun 1); -by (Simp_tac 1); -by (rtac refl 1); -qed "strict_Rep_CFun1"; - - -(* ------------------------------------------------------------------------ *) -(* results about strictify *) -(* ------------------------------------------------------------------------ *) - -Goalw [Istrictify_def] - "Istrictify(f)(UU)= (UU)"; -by (Simp_tac 1); -qed "Istrictify1"; - -Goalw [Istrictify_def] - "~x=UU ==> Istrictify(f)(x)=f$x"; -by (Asm_simp_tac 1); -qed "Istrictify2"; - -Goal "monofun(Istrictify)"; -by (rtac monofunI 1); -by (strip_tac 1); -by (rtac (less_fun RS iffD2) 1); -by (strip_tac 1); -by (res_inst_tac [("Q","xa=UU")] (excluded_middle RS disjE) 1); -by (stac Istrictify2 1); -by (atac 1); -by (stac Istrictify2 1); -by (atac 1); -by (rtac monofun_cfun_fun 1); -by (atac 1); -by (hyp_subst_tac 1); -by (stac Istrictify1 1); -by (stac Istrictify1 1); -by (rtac refl_less 1); -qed "monofun_Istrictify1"; - -Goal "monofun(Istrictify(f))"; -by (rtac monofunI 1); -by (strip_tac 1); -by (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1); -by (stac Istrictify2 1); -by (etac notUU_I 1); -by (atac 1); -by (stac Istrictify2 1); -by (atac 1); -by (rtac monofun_cfun_arg 1); -by (atac 1); -by (hyp_subst_tac 1); -by (stac Istrictify1 1); -by (rtac minimal 1); -qed "monofun_Istrictify2"; - - -Goal "contlub(Istrictify)"; -by (rtac contlubI 1); -by (strip_tac 1); -by (rtac (expand_fun_eq RS iffD2) 1); -by (strip_tac 1); -by (stac thelub_fun 1); -by (etac (monofun_Istrictify1 RS ch2ch_monofun) 1); -by (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1); -by (stac Istrictify2 1); -by (atac 1); -by (stac (Istrictify2 RS ext) 1); -by (atac 1); -by (stac thelub_cfun 1); -by (atac 1); -by (stac beta_cfun 1); -by (rtac cont_lubcfun 1); -by (atac 1); -by (rtac refl 1); -by (hyp_subst_tac 1); -by (stac Istrictify1 1); -by (stac (Istrictify1 RS ext) 1); -by (rtac (chain_UU_I_inverse RS sym) 1); -by (rtac (refl RS allI) 1); -qed "contlub_Istrictify1"; +(* legacy ML bindings *) -Goal "contlub(Istrictify(f::'a -> 'b))"; -by (rtac contlubI 1); -by (strip_tac 1); -by (case_tac "lub(range(Y))=(UU::'a)" 1); -by (asm_simp_tac (simpset() addsimps [Istrictify1, chain_UU_I_inverse, chain_UU_I, Istrictify1]) 1); -by (stac Istrictify2 1); -by (atac 1); -by (res_inst_tac [("s","lub(range(%i. f$(Y i)))")] trans 1); -by (rtac contlub_cfun_arg 1); -by (atac 1); -by (rtac lub_equal2 1); -by (best_tac (claset() addIs [ch2ch_monofun, monofun_Istrictify2]) 3); -by (best_tac (claset() addIs [ch2ch_monofun, monofun_Rep_CFun2]) 2); -by (rtac (chain_mono2 RS exE) 1); -by (atac 2); -by (etac chain_UU_I_inverse2 1); -by (blast_tac (claset() addIs [Istrictify2 RS sym]) 1); -qed "contlub_Istrictify2"; - - -bind_thm ("cont_Istrictify1", contlub_Istrictify1 RS - (monofun_Istrictify1 RS monocontlub2cont)); - -bind_thm ("cont_Istrictify2", contlub_Istrictify2 RS - (monofun_Istrictify2 RS monocontlub2cont)); - - -Goalw [strictify_def] "strictify$f$UU=UU"; -by (stac beta_cfun 1); -by (simp_tac (simpset() addsimps [cont_Istrictify2,cont_Istrictify1, cont2cont_CF1L]) 1); -by (stac beta_cfun 1); -by (rtac cont_Istrictify2 1); -by (rtac Istrictify1 1); -qed "strictify1"; - -Goalw [strictify_def] "~x=UU ==> strictify$f$x=f$x"; -by (stac beta_cfun 1); -by (simp_tac (simpset() addsimps [cont_Istrictify2,cont_Istrictify1, cont2cont_CF1L]) 1); -by (stac beta_cfun 1); -by (rtac cont_Istrictify2 1); -by (etac Istrictify2 1); -qed "strictify2"; - - -(* ------------------------------------------------------------------------ *) -(* Instantiate the simplifier *) -(* ------------------------------------------------------------------------ *) - -Addsimps [minimal,refl_less,beta_cfun,strict_Rep_CFun1,strictify1, strictify2]; - - -(* ------------------------------------------------------------------------ *) -(* use cont_tac as autotac. *) -(* ------------------------------------------------------------------------ *) - -(* HINT: cont_tac is now installed in simplifier in Lift.ML ! *) -(*simpset_ref() := simpset() addsolver (K (DEPTH_SOLVE_1 o cont_tac));*) - -(* ------------------------------------------------------------------------ *) -(* some lemmata for functions with flat/chfin domain/range types *) -(* ------------------------------------------------------------------------ *) - -Goal "chain (Y::nat => 'a::cpo->'b::chfin) \ -\ ==> !s. ? n. lub(range(Y))$s = Y n$s"; -by (rtac allI 1); -by (stac contlub_cfun_fun 1); -by (atac 1); -by (fast_tac (HOL_cs addSIs [thelubI,chfin,lub_finch2,chfin2finch,ch2ch_Rep_CFunL])1); -qed "chfin_Rep_CFunR"; - -(* ------------------------------------------------------------------------ *) -(* continuous isomorphisms are strict *) -(* a prove for embedding projection pairs is similar *) -(* ------------------------------------------------------------------------ *) - -Goal -"!!f g.[|!y. f$(g$y)=(y::'b) ; !x. g$(f$x)=(x::'a) |] \ -\ ==> f$UU=UU & g$UU=UU"; -by (rtac conjI 1); -by (rtac UU_I 1); -by (res_inst_tac [("s","f$(g$(UU::'b))"),("t","UU::'b")] subst 1); -by (etac spec 1); -by (rtac (minimal RS monofun_cfun_arg) 1); -by (rtac UU_I 1); -by (res_inst_tac [("s","g$(f$(UU::'a))"),("t","UU::'a")] subst 1); -by (etac spec 1); -by (rtac (minimal RS monofun_cfun_arg) 1); -qed "iso_strict"; - - -Goal "[|!x. rep$(ab$x)=x;!y. ab$(rep$y)=y; z~=UU|] ==> rep$z ~= UU"; -by (etac contrapos_nn 1); -by (dres_inst_tac [("f","ab")] cfun_arg_cong 1); -by (etac box_equals 1); -by (fast_tac HOL_cs 1); -by (etac (iso_strict RS conjunct1) 1); -by (atac 1); -qed "isorep_defined"; - -Goal "[|!x. rep$(ab$x) = x;!y. ab$(rep$y)=y ; z~=UU|] ==> ab$z ~= UU"; -by (etac contrapos_nn 1); -by (dres_inst_tac [("f","rep")] cfun_arg_cong 1); -by (etac box_equals 1); -by (fast_tac HOL_cs 1); -by (etac (iso_strict RS conjunct2) 1); -by (atac 1); -qed "isoabs_defined"; - -(* ------------------------------------------------------------------------ *) -(* propagation of flatness and chainfiniteness by continuous isomorphisms *) -(* ------------------------------------------------------------------------ *) +val Istrictify_def = thm "Istrictify_def"; +val strictify_def = thm "strictify_def"; +val ID_def = thm "ID_def"; +val oo_def = thm "oo_def"; +val inst_cfun_pcpo = thm "inst_cfun_pcpo"; +val contlub_Rep_CFun1 = thm "contlub_Rep_CFun1"; +val cont_Rep_CFun1 = thm "cont_Rep_CFun1"; +val contlub_cfun_fun = thm "contlub_cfun_fun"; +val cont_cfun_fun = thm "cont_cfun_fun"; +val contlub_cfun = thm "contlub_cfun"; +val cont_cfun = thm "cont_cfun"; +val cont2cont_Rep_CFun = thm "cont2cont_Rep_CFun"; +val cont2mono_LAM = thm "cont2mono_LAM"; +val cont2cont_LAM = thm "cont2cont_LAM"; +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 chfin_Rep_CFunR = thm "chfin_Rep_CFunR"; +val iso_strict = thm "iso_strict"; +val isorep_defined = thm "isorep_defined"; +val isoabs_defined = thm "isoabs_defined"; +val chfin2chfin = thm "chfin2chfin"; +val flat2flat = thm "flat2flat"; +val flat_codom = thm "flat_codom"; +val ID1 = thm "ID1"; +val cfcomp1 = thm "cfcomp1"; +val cfcomp2 = thm "cfcomp2"; +val ID2 = thm "ID2"; +val ID3 = thm "ID3"; +val assoc_oo = thm "assoc_oo"; -Goal "!!f g.[|! Y::nat=>'a. chain Y --> (? n. max_in_chain n Y); \ -\ !y. f$(g$y)=(y::'b) ; !x. g$(f$x)=(x::'a::chfin) |] \ -\ ==> ! Y::nat=>'b. chain Y --> (? n. max_in_chain n Y)"; -by (rewtac max_in_chain_def); -by (strip_tac 1); -by (rtac exE 1); -by (res_inst_tac [("P","chain(%i. g$(Y i))")] mp 1); -by (etac spec 1); -by (etac ch2ch_Rep_CFunR 1); -by (rtac exI 1); -by (strip_tac 1); -by (res_inst_tac [("s","f$(g$(Y x))"),("t","Y(x)")] subst 1); -by (etac spec 1); -by (res_inst_tac [("s","f$(g$(Y j))"),("t","Y(j)")] subst 1); -by (etac spec 1); -by (rtac cfun_arg_cong 1); -by (rtac mp 1); -by (etac spec 1); -by (atac 1); -qed "chfin2chfin"; - - -Goal "!!f g.[|!x y::'a. x< x=UU | x=y; \ -\ !y. f$(g$y)=(y::'b); !x. g$(f$x)=(x::'a)|] ==> !x y::'b. x< x=UU | x=y"; -by (strip_tac 1); -by (rtac disjE 1); -by (res_inst_tac [("P","g$x< f$(UU::'a)=(UU::'b) | (!z. f$(z::'a)=c)"; -by (case_tac "f$(x::'a)=(UU::'b)" 1); -by (rtac disjI1 1); -by (rtac UU_I 1); -by (res_inst_tac [("s","f$(x)"),("t","UU::'b")] subst 1); -by (atac 1); -by (rtac (minimal RS monofun_cfun_arg) 1); -by (case_tac "f$(UU::'a)=(UU::'b)" 1); -by (etac disjI1 1); -by (rtac disjI2 1); -by (rtac allI 1); -by (hyp_subst_tac 1); -by (res_inst_tac [("a","f$(UU::'a)")] (refl RS box_equals) 1); -by (res_inst_tac [("fo5","f")] ((minimal RS monofun_cfun_arg) RS (ax_flat RS spec RS spec RS mp) RS disjE) 1); -by (contr_tac 1); -by (atac 1); -by (res_inst_tac [("fo5","f")] ((minimal RS monofun_cfun_arg) RS (ax_flat RS spec RS spec RS mp) RS disjE) 1); -by (contr_tac 1); -by (atac 1); -qed "flat_codom"; - - -(* ------------------------------------------------------------------------ *) -(* Access to definitions *) -(* ------------------------------------------------------------------------ *) - - -Goalw [ID_def] "ID$x=x"; -by (stac beta_cfun 1); -by (rtac cont_id 1); -by (rtac refl 1); -qed "ID1"; - -Goalw [oo_def] "(f oo g)=(LAM x. f$(g$x))"; -by (stac beta_cfun 1); -by (Simp_tac 1); -by (stac beta_cfun 1); -by (Simp_tac 1); -by (rtac refl 1); -qed "cfcomp1"; - -Goal "(f oo g)$x=f$(g$x)"; -by (stac cfcomp1 1); -by (stac beta_cfun 1); -by (Simp_tac 1); -by (rtac refl 1); -qed "cfcomp2"; - - -(* ------------------------------------------------------------------------ *) -(* Show that interpretation of (pcpo,_->_) is a category *) -(* The class of objects is interpretation of syntactical class pcpo *) -(* The class of arrows between objects 'a and 'b is interpret. of 'a -> 'b *) -(* The identity arrow is interpretation of ID *) -(* The composition of f and g is interpretation of oo *) -(* ------------------------------------------------------------------------ *) - - -Goal "f oo ID = f "; -by (rtac ext_cfun 1); -by (stac cfcomp2 1); -by (stac ID1 1); -by (rtac refl 1); -qed "ID2"; - -Goal "ID oo f = f "; -by (rtac ext_cfun 1); -by (stac cfcomp2 1); -by (stac ID1 1); -by (rtac refl 1); -qed "ID3"; - - -Goal "f oo (g oo h) = (f oo g) oo h"; -by (rtac ext_cfun 1); -by (res_inst_tac [("s","f$(g$(h$x))")] trans 1); -by (stac cfcomp2 1); -by (stac cfcomp2 1); -by (rtac refl 1); -by (stac cfcomp2 1); -by (stac cfcomp2 1); -by (rtac refl 1); -qed "assoc_oo"; - -(* ------------------------------------------------------------------------ *) -(* Merge the different rewrite rules for the simplifier *) -(* ------------------------------------------------------------------------ *) - -Addsimps ([ID1,ID2,ID3,cfcomp2]); - - +structure Cfun3 = +struct + val thy = the_context (); + val Istrictify_def = Istrictify_def; + val strictify_def = strictify_def; + val ID_def = ID_def; + val oo_def = oo_def; +end; diff -r 2454493bd77b -r eb3b1a5c304e src/HOLCF/Cfun3.thy --- a/src/HOLCF/Cfun3.thy Wed Mar 02 23:28:17 2005 +0100 +++ b/src/HOLCF/Cfun3.thy Wed Mar 02 23:58:02 2005 +0100 @@ -1,25 +1,29 @@ (* Title: HOLCF/Cfun3.thy ID: $Id$ Author: Franz Regensburger + License: GPL (GNU GENERAL PUBLIC LICENSE) Class instance of -> for class pcpo *) -Cfun3 = Cfun2 + +theory Cfun3 = Cfun2: + +instance "->" :: (cpo,cpo)cpo +by (intro_classes, rule cpo_cfun) -instance "->" :: (cpo,cpo)cpo (cpo_cfun) -instance "->" :: (cpo,pcpo)pcpo (least_cfun) +instance "->" :: (cpo,pcpo)pcpo +by (intro_classes, rule least_cfun) -default pcpo +defaultsort pcpo consts Istrictify :: "('a->'b)=>'a=>'b" strictify :: "('a->'b)->'a->'b" defs -Istrictify_def "Istrictify f x == if x=UU then UU else f$x" -strictify_def "strictify == (LAM f x. Istrictify f x)" +Istrictify_def: "Istrictify f x == if x=UU then UU else f$x" +strictify_def: "strictify == (LAM f x. Istrictify f x)" consts ID :: "('a::cpo) -> 'a" @@ -31,7 +35,512 @@ defs - ID_def "ID ==(LAM x. x)" - oo_def "cfcomp == (LAM f g x. f$(g$x))" + ID_def: "ID ==(LAM x. x)" + oo_def: "cfcomp == (LAM f g x. f$(g$x))" + +(* Title: HOLCF/Cfun3 + ID: $Id$ + Author: Franz Regensburger + License: GPL (GNU GENERAL PUBLIC LICENSE) + +Class instance of -> for class pcpo +*) + +(* for compatibility with old HOLCF-Version *) +lemma inst_cfun_pcpo: "UU = Abs_CFun(%x. UU)" +apply (simp add: UU_def UU_cfun_def) +done + +(* ------------------------------------------------------------------------ *) +(* the contlub property for Rep_CFun its 'first' argument *) +(* ------------------------------------------------------------------------ *) + +lemma contlub_Rep_CFun1: "contlub(Rep_CFun)" +apply (rule contlubI) +apply (intro strip) +apply (rule expand_fun_eq [THEN iffD2]) +apply (intro strip) +apply (subst thelub_cfun) +apply assumption +apply (subst Cfunapp2) +apply (erule cont_lubcfun) +apply (subst thelub_fun) +apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun]) +apply (rule refl) +done + + +(* ------------------------------------------------------------------------ *) +(* the cont property for Rep_CFun in its first argument *) +(* ------------------------------------------------------------------------ *) + +lemma cont_Rep_CFun1: "cont(Rep_CFun)" +apply (rule monocontlub2cont) +apply (rule monofun_Rep_CFun1) +apply (rule contlub_Rep_CFun1) +done + + +(* ------------------------------------------------------------------------ *) +(* contlub, cont properties of Rep_CFun in its first argument in mixfix _[_] *) +(* ------------------------------------------------------------------------ *) + +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 + + +(* ------------------------------------------------------------------------ *) +(* contlub, cont properties of Rep_CFun in both argument in mixfix _[_] *) +(* ------------------------------------------------------------------------ *) + +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 + + +(* ------------------------------------------------------------------------ *) +(* cont2cont lemma for 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 + + + +(* ------------------------------------------------------------------------ *) +(* cont2mono Lemma for %x. LAM y. c1(x)(y) *) +(* ------------------------------------------------------------------------ *) + +lemma cont2mono_LAM: +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) +apply (intro strip) +apply (subst less_cfun) +apply (subst less_fun) +apply (rule allI) +apply (subst beta_cfun) +apply (rule p1) +apply (subst beta_cfun) +apply (rule p1) +apply (erule p2 [THEN monofunE, THEN spec, THEN spec, THEN mp]) +done + +(* ------------------------------------------------------------------------ *) +(* cont2cont Lemma for %x. LAM y. c1 x y) *) +(* ------------------------------------------------------------------------ *) + +lemma cont2cont_LAM: +assumes p1: "!!x. cont(c1 x)" +assumes p2: "!!y. cont(%x. c1 x y)" +shows "cont(%x. LAM y. c1 x y)" +apply (rule monocontlub2cont) +apply (rule p1 [THEN cont2mono_LAM]) +apply (rule p2 [THEN cont2mono]) +apply (rule contlubI) +apply (intro strip) +apply (subst thelub_cfun) +apply (rule p1 [THEN cont2mono_LAM, THEN ch2ch_monofun]) +apply (rule p2 [THEN cont2mono]) +apply assumption +apply (rule_tac f = "Abs_CFun" in arg_cong) +apply (rule ext) +apply (subst p1 [THEN beta_cfun, THEN ext]) +apply (erule p2 [THEN cont2contlub, THEN contlubE, THEN spec, THEN mp]) +done + +(* ------------------------------------------------------------------------ *) +(* cont2cont tactic *) +(* ------------------------------------------------------------------------ *) + +lemmas cont_lemmas1 = cont_const cont_id cont_Rep_CFun2 + cont2cont_Rep_CFun cont2cont_LAM + +declare cont_lemmas1 [simp] + +(* HINT: 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)));*) + +(* ------------------------------------------------------------------------ *) +(* function application _[_] is strict in its first arguments *) +(* ------------------------------------------------------------------------ *) + +lemma strict_Rep_CFun1: "(UU::'a::cpo->'b)$x = (UU::'b)" +apply (subst inst_cfun_pcpo) +apply (subst beta_cfun) +apply (simp (no_asm)) +apply (rule refl) +done + + +(* ------------------------------------------------------------------------ *) +(* results about strictify *) +(* ------------------------------------------------------------------------ *) + +lemma Istrictify1: + "Istrictify(f)(UU)= (UU)" +apply (unfold Istrictify_def) +apply (simp (no_asm)) +done + +lemma Istrictify2: + "~x=UU ==> Istrictify(f)(x)=f$x" +apply (unfold Istrictify_def) +apply (simp (no_asm_simp)) +done + +lemma monofun_Istrictify1: "monofun(Istrictify)" +apply (rule monofunI) +apply (intro strip) +apply (rule less_fun [THEN iffD2]) +apply (intro strip) +apply (rule_tac Q = "xa=UU" in excluded_middle [THEN disjE]) +apply (subst Istrictify2) +apply assumption +apply (subst Istrictify2) +apply assumption +apply (rule monofun_cfun_fun) +apply assumption +apply (erule ssubst) +apply (subst Istrictify1) +apply (subst Istrictify1) +apply (rule refl_less) +done + +lemma monofun_Istrictify2: "monofun(Istrictify(f))" +apply (rule monofunI) +apply (intro strip) +apply (rule_tac Q = "x=UU" in excluded_middle [THEN disjE]) +apply (simplesubst Istrictify2) +apply (erule notUU_I) +apply assumption +apply (subst Istrictify2) +apply assumption +apply (rule monofun_cfun_arg) +apply assumption +apply (erule ssubst) +apply (subst Istrictify1) +apply (rule minimal) +done + + +lemma contlub_Istrictify1: "contlub(Istrictify)" +apply (rule contlubI) +apply (intro strip) +apply (rule expand_fun_eq [THEN iffD2]) +apply (intro strip) +apply (subst thelub_fun) +apply (erule monofun_Istrictify1 [THEN ch2ch_monofun]) +apply (rule_tac Q = "x=UU" in excluded_middle [THEN disjE]) +apply (subst Istrictify2) +apply assumption +apply (subst Istrictify2 [THEN ext]) +apply assumption +apply (subst thelub_cfun) +apply assumption +apply (subst beta_cfun) +apply (rule cont_lubcfun) +apply assumption +apply (rule refl) +apply (erule ssubst) +apply (subst Istrictify1) +apply (subst Istrictify1 [THEN ext]) +apply (rule chain_UU_I_inverse [symmetric]) +apply (rule refl [THEN allI]) +done + +lemma contlub_Istrictify2: "contlub(Istrictify(f::'a -> 'b))" +apply (rule contlubI) +apply (intro strip) +apply (case_tac "lub (range (Y))= (UU::'a) ") +apply (simp (no_asm_simp) add: Istrictify1 chain_UU_I_inverse chain_UU_I Istrictify1) +apply (subst Istrictify2) +apply assumption +apply (rule_tac s = "lub (range (%i. f$ (Y i))) " in trans) +apply (rule contlub_cfun_arg) +apply assumption +apply (rule lub_equal2) +prefer 3 apply (best intro: ch2ch_monofun monofun_Istrictify2) +prefer 2 apply (best intro: ch2ch_monofun monofun_Rep_CFun2) +apply (rule chain_mono2 [THEN exE]) +prefer 2 apply (assumption) +apply (erule chain_UU_I_inverse2) +apply (blast intro: Istrictify2 [symmetric]) +done + + +lemmas cont_Istrictify1 = contlub_Istrictify1 [THEN monofun_Istrictify1 [THEN monocontlub2cont], standard] + +lemmas cont_Istrictify2 = contlub_Istrictify2 [THEN monofun_Istrictify2 [THEN monocontlub2cont], standard] + + +lemma strictify1: "strictify$f$UU=UU" +apply (unfold strictify_def) +apply (subst beta_cfun) +apply (simp (no_asm) add: cont_Istrictify2 cont_Istrictify1 cont2cont_CF1L) +apply (subst beta_cfun) +apply (rule cont_Istrictify2) +apply (rule Istrictify1) +done + +lemma strictify2: "~x=UU ==> strictify$f$x=f$x" +apply (unfold strictify_def) +apply (subst beta_cfun) +apply (simp (no_asm) add: cont_Istrictify2 cont_Istrictify1 cont2cont_CF1L) +apply (subst beta_cfun) +apply (rule cont_Istrictify2) +apply (erule Istrictify2) +done + + +(* ------------------------------------------------------------------------ *) +(* Instantiate the simplifier *) +(* ------------------------------------------------------------------------ *) + +declare minimal [simp] refl_less [simp] beta_cfun [simp] strict_Rep_CFun1 [simp] strictify1 [simp] strictify2 [simp] + + +(* ------------------------------------------------------------------------ *) +(* use cont_tac as autotac. *) +(* ------------------------------------------------------------------------ *) + +(* HINT: cont_tac is now installed in simplifier in Lift.ML ! *) +(*simpset_ref() := simpset() addsolver (K (DEPTH_SOLVE_1 o cont_tac));*) + +(* ------------------------------------------------------------------------ *) +(* some lemmata for functions with flat/chfin domain/range types *) +(* ------------------------------------------------------------------------ *) + +lemma chfin_Rep_CFunR: "chain (Y::nat => 'a::cpo->'b::chfin) + ==> !s. ? n. lub(range(Y))$s = Y n$s" +apply (rule allI) +apply (subst contlub_cfun_fun) +apply assumption +apply (fast intro!: thelubI chfin lub_finch2 chfin2finch ch2ch_Rep_CFunL) +done + +(* ------------------------------------------------------------------------ *) +(* continuous isomorphisms are strict *) +(* a prove for embedding projection pairs is similar *) +(* ------------------------------------------------------------------------ *) + +lemma iso_strict: +"!!f g.[|!y. f$(g$y)=(y::'b) ; !x. g$(f$x)=(x::'a) |] + ==> f$UU=UU & g$UU=UU" +apply (rule conjI) +apply (rule UU_I) +apply (rule_tac s = "f$ (g$ (UU::'b))" and t = "UU::'b" in subst) +apply (erule spec) +apply (rule minimal [THEN monofun_cfun_arg]) +apply (rule UU_I) +apply (rule_tac s = "g$ (f$ (UU::'a))" and t = "UU::'a" in subst) +apply (erule spec) +apply (rule minimal [THEN monofun_cfun_arg]) +done + + +lemma isorep_defined: "[|!x. rep$(ab$x)=x;!y. ab$(rep$y)=y; z~=UU|] ==> rep$z ~= UU" +apply (erule contrapos_nn) +apply (drule_tac f = "ab" in cfun_arg_cong) +apply (erule box_equals) +apply fast +apply (erule iso_strict [THEN conjunct1]) +apply assumption +done + +lemma isoabs_defined: "[|!x. rep$(ab$x) = x;!y. ab$(rep$y)=y ; z~=UU|] ==> ab$z ~= UU" +apply (erule contrapos_nn) +apply (drule_tac f = "rep" in cfun_arg_cong) +apply (erule box_equals) +apply fast +apply (erule iso_strict [THEN conjunct2]) +apply assumption +done + +(* ------------------------------------------------------------------------ *) +(* propagation of flatness and chainfiniteness by continuous isomorphisms *) +(* ------------------------------------------------------------------------ *) + +lemma chfin2chfin: "!!f g.[|! Y::nat=>'a. chain Y --> (? n. max_in_chain n Y); + !y. f$(g$y)=(y::'b) ; !x. g$(f$x)=(x::'a::chfin) |] + ==> ! Y::nat=>'b. chain Y --> (? n. max_in_chain n Y)" +apply (unfold max_in_chain_def) +apply (intro strip) +apply (rule exE) +apply (rule_tac P = "chain (%i. g$ (Y i))" in mp) +apply (erule spec) +apply (erule ch2ch_Rep_CFunR) +apply (rule exI) +apply (intro strip) +apply (rule_tac s = "f$ (g$ (Y x))" and t = "Y (x) " in subst) +apply (erule spec) +apply (rule_tac s = "f$ (g$ (Y j))" and t = "Y (j) " in subst) +apply (erule spec) +apply (rule cfun_arg_cong) +apply (rule mp) +apply (erule spec) +apply assumption +done + + +lemma flat2flat: "!!f g.[|!x y::'a. x< x=UU | x=y; + !y. f$(g$y)=(y::'b); !x. g$(f$x)=(x::'a)|] ==> !x y::'b. x< x=UU | x=y" +apply (intro strip) +apply (rule disjE) +apply (rule_tac P = "g$x< f$(UU::'a)=(UU::'b) | (!z. f$(z::'a)=c)" +apply (case_tac "f$ (x::'a) = (UU::'b) ") +apply (rule disjI1) +apply (rule UU_I) +apply (rule_tac s = "f$ (x) " and t = "UU::'b" in subst) +apply assumption +apply (rule minimal [THEN monofun_cfun_arg]) +apply (case_tac "f$ (UU::'a) = (UU::'b) ") +apply (erule disjI1) +apply (rule disjI2) +apply (rule allI) +apply (erule subst) +apply (rule_tac a = "f$ (UU::'a) " in refl [THEN box_equals]) +apply (rule_tac fo5 = "f" in minimal [THEN monofun_cfun_arg, THEN ax_flat [THEN spec, THEN spec, THEN mp], THEN disjE]) +apply simp +apply assumption +apply (rule_tac fo5 = "f" in minimal [THEN monofun_cfun_arg, THEN ax_flat [THEN spec, THEN spec, THEN mp], THEN disjE]) +apply simp +apply assumption +done + + +(* ------------------------------------------------------------------------ *) +(* Access to definitions *) +(* ------------------------------------------------------------------------ *) + + +lemma ID1: "ID$x=x" +apply (unfold ID_def) +apply (subst beta_cfun) +apply (rule cont_id) +apply (rule refl) +done + +lemma cfcomp1: "(f oo g)=(LAM x. f$(g$x))" +apply (unfold oo_def) +apply (subst beta_cfun) +apply (simp (no_asm)) +apply (subst beta_cfun) +apply (simp (no_asm)) +apply (rule refl) +done + +lemma cfcomp2: "(f oo g)$x=f$(g$x)" +apply (subst cfcomp1) +apply (subst beta_cfun) +apply (simp (no_asm)) +apply (rule refl) +done + + +(* ------------------------------------------------------------------------ *) +(* Show that interpretation of (pcpo,_->_) is a category *) +(* The class of objects is interpretation of syntactical class pcpo *) +(* The class of arrows between objects 'a and 'b is interpret. of 'a -> 'b *) +(* The identity arrow is interpretation of ID *) +(* The composition of f and g is interpretation of oo *) +(* ------------------------------------------------------------------------ *) + + +lemma ID2: "f oo ID = f " +apply (rule ext_cfun) +apply (subst cfcomp2) +apply (subst ID1) +apply (rule refl) +done + +lemma ID3: "ID oo f = f " +apply (rule ext_cfun) +apply (subst cfcomp2) +apply (subst ID1) +apply (rule refl) +done + + +lemma assoc_oo: "f oo (g oo h) = (f oo g) oo h" +apply (rule ext_cfun) +apply (rule_tac s = "f$ (g$ (h$x))" in trans) +apply (subst cfcomp2) +apply (subst cfcomp2) +apply (rule refl) +apply (subst cfcomp2) +apply (subst cfcomp2) +apply (rule refl) +done + +(* ------------------------------------------------------------------------ *) +(* Merge the different rewrite rules for the simplifier *) +(* ------------------------------------------------------------------------ *) + +declare ID1[simp] ID2[simp] ID3[simp] cfcomp2[simp] end