# HG changeset patch # User paulson # Date 962719091 -7200 # Node ID 428385c4bc50801ad696fa844649ba58616a781a # Parent 7edd3e5f26d455b053436f80026e8928006d3f96 removed most batch-style proofs diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Cfun1.ML --- a/src/HOLCF/Cfun1.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Cfun1.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,129 +3,103 @@ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for Cfun1.thy +The type -> of continuous functions. *) -open Cfun1; - (* ------------------------------------------------------------------------ *) (* 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 *) (* ------------------------------------------------------------------------ *) -qed_goal "Rep_Cfun" thy "Rep_CFun fo : CFun" -(fn prems => - [ - (rtac Rep_CFun 1) - ]); +val prems = goal thy "Rep_CFun fo : CFun"; +by (rtac Rep_CFun 1); +qed "Rep_Cfun"; -qed_goal "Rep_Cfun_inverse" thy "Abs_CFun (Rep_CFun fo) = fo" -(fn prems => - [ - (rtac Rep_CFun_inverse 1) - ]); +val prems = goal thy "Abs_CFun (Rep_CFun fo) = fo"; +by (rtac Rep_CFun_inverse 1); +qed "Rep_Cfun_inverse"; -qed_goal "Abs_Cfun_inverse" thy "f:CFun==>Rep_CFun(Abs_CFun f)=f" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac Abs_CFun_inverse 1) - ]); +val prems = goal thy "f:CFun==>Rep_CFun(Abs_CFun f)=f"; +by (cut_facts_tac prems 1); +by (etac Abs_CFun_inverse 1); +qed "Abs_Cfun_inverse"; (* ------------------------------------------------------------------------ *) (* less_cfun is a partial order on type 'a -> 'b *) (* ------------------------------------------------------------------------ *) -qed_goalw "refl_less_cfun" thy [less_cfun_def] "(f::'a->'b) << f" -(fn prems => - [ - (rtac refl_less 1) - ]); +val prems = goalw thy [less_cfun_def] "(f::'a->'b) << f"; +by (rtac refl_less 1); +qed "refl_less_cfun"; -qed_goalw "antisym_less_cfun" thy [less_cfun_def] - "[|(f1::'a->'b) << f2; f2 << f1|] ==> f1 = f2" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac injD 1), - (rtac antisym_less 2), - (atac 3), - (atac 2), - (rtac inj_inverseI 1), - (rtac Rep_Cfun_inverse 1) - ]); +val prems = goalw thy [less_cfun_def] + "[|(f1::'a->'b) << f2; f2 << f1|] ==> f1 = f2"; +by (cut_facts_tac prems 1); +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"; -qed_goalw "trans_less_cfun" thy [less_cfun_def] - "[|(f1::'a->'b) << f2; f2 << f3|] ==> f1 << f3" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac trans_less 1), - (atac 1) - ]); +val prems = goalw thy [less_cfun_def] + "[|(f1::'a->'b) << f2; f2 << f3|] ==> f1 << f3"; +by (cut_facts_tac prems 1); +by (etac trans_less 1); +by (atac 1); +qed "trans_less_cfun"; (* ------------------------------------------------------------------------ *) (* lemmas about application of continuous functions *) (* ------------------------------------------------------------------------ *) -qed_goal "cfun_cong" thy - "[| f=g; x=y |] ==> f`x = g`y" -(fn prems => - [ - (cut_facts_tac prems 1), - (fast_tac HOL_cs 1) - ]); +val prems = goal thy + "[| f=g; x=y |] ==> f`x = g`y"; +by (cut_facts_tac prems 1); +by (fast_tac HOL_cs 1); +qed "cfun_cong"; -qed_goal "cfun_fun_cong" thy "f=g ==> f`x = g`x" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac cfun_cong 1), - (rtac refl 1) - ]); +val prems = goal thy "f=g ==> f`x = g`x"; +by (cut_facts_tac prems 1); +by (etac cfun_cong 1); +by (rtac refl 1); +qed "cfun_fun_cong"; -qed_goal "cfun_arg_cong" thy "x=y ==> f`x = f`y" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac cfun_cong 1), - (rtac refl 1), - (atac 1) - ]); +val prems = goal thy "x=y ==> f`x = f`y"; +by (cut_facts_tac prems 1); +by (rtac cfun_cong 1); +by (rtac refl 1); +by (atac 1); +qed "cfun_arg_cong"; (* ------------------------------------------------------------------------ *) (* additional lemma about the isomorphism between -> and Cfun *) (* ------------------------------------------------------------------------ *) -qed_goal "Abs_Cfun_inverse2" thy "cont f ==> Rep_CFun (Abs_CFun f) = f" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac Abs_Cfun_inverse 1), - (rewtac CFun_def), - (etac (mem_Collect_eq RS ssubst) 1) - ]); +val prems = goal thy "cont f ==> Rep_CFun (Abs_CFun f) = f"; +by (cut_facts_tac prems 1); +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 *) (* ------------------------------------------------------------------------ *) -qed_goal "Cfunapp2" thy "cont f ==> (Abs_CFun f)`x = f x" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac (Abs_Cfun_inverse2 RS fun_cong) 1) - ]); +val prems = goal thy "cont f ==> (Abs_CFun f)`x = f x"; +by (cut_facts_tac prems 1); +by (etac (Abs_Cfun_inverse2 RS fun_cong) 1); +qed "Cfunapp2"; (* ------------------------------------------------------------------------ *) (* beta - equality for continuous functions *) (* ------------------------------------------------------------------------ *) -qed_goal "beta_cfun" thy - "cont(c1) ==> (LAM x .c1 x)`u = c1 u" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac Cfunapp2 1), - (atac 1) - ]); +val prems = goal thy + "cont(c1) ==> (LAM x .c1 x)`u = c1 u"; +by (cut_facts_tac prems 1); +by (rtac Cfunapp2 1); +by (atac 1); +qed "beta_cfun"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Cfun2.ML --- a/src/HOLCF/Cfun2.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Cfun2.ML Tue Jul 04 15:58:11 2000 +0200 @@ -1,52 +1,42 @@ -(* Title: HOLCF/cfun2.thy +(* Title: HOLCF/Cfun2 ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for cfun2.thy +Class Instance ->::(cpo,cpo)po *) -open Cfun2; - (* for compatibility with old HOLCF-Version *) -qed_goal "inst_cfun_po" thy "(op <<)=(%f1 f2. Rep_CFun f1 << Rep_CFun f2)" - (fn prems => - [ - (fold_goals_tac [less_cfun_def]), - (rtac refl 1) - ]); +val prems = goal thy "(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 *) (* ------------------------------------------------------------------------ *) -qed_goal "less_cfun" thy "( f1 << f2 ) = (Rep_CFun(f1) << Rep_CFun(f2))" -(fn prems => - [ - (simp_tac (simpset() addsimps [inst_cfun_po]) 1) - ]); +val prems = goal thy "( 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 *) (* ------------------------------------------------------------------------ *) -qed_goal "minimal_cfun" thy "Abs_CFun(% x. UU) << f" -(fn prems => - [ - (stac less_cfun 1), - (stac Abs_Cfun_inverse2 1), - (rtac cont_const 1), - (rtac minimal_fun 1) - ]); +val prems = goal thy "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); -qed_goal "least_cfun" thy "? x::'a->'b::pcpo.!y. x< - [ - (res_inst_tac [("x","Abs_CFun(% x. UU)")] exI 1), - (rtac (minimal_cfun RS allI) 1) - ]); +val prems = goal thy "? x::'a->'b::pcpo.!y. x< 'b *) @@ -54,13 +44,11 @@ (* cont_Rep_CFun2 ==> monofun_Rep_CFun2 & contlub_Rep_CFun2 *) (* ------------------------------------------------------------------------ *) -qed_goal "cont_Rep_CFun2" thy "cont(Rep_CFun(fo))" -(fn prems => - [ - (res_inst_tac [("P","cont")] CollectD 1), - (fold_goals_tac [CFun_def]), - (rtac Rep_Cfun 1) - ]); +val prems = goal thy "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)) *) @@ -85,26 +73,22 @@ (* Rep_CFun is monotone in its 'first' argument *) (* ------------------------------------------------------------------------ *) -qed_goalw "monofun_Rep_CFun1" thy [monofun] "monofun(Rep_CFun)" -(fn prems => - [ - (strip_tac 1), - (etac (less_cfun RS subst) 1) - ]); +val prems = goalw thy [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 [_]_ *) (* ------------------------------------------------------------------------ *) -qed_goal "monofun_cfun_fun" thy "f1 << f2 ==> f1`x << f2`x" -(fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("x","x")] spec 1), - (rtac (less_fun RS subst) 1), - (etac (monofun_Rep_CFun1 RS monofunE RS spec RS spec RS mp) 1) - ]); +val prems = goal thy "f1 << f2 ==> f1`x << f2`x"; +by (cut_facts_tac prems 1); +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); @@ -114,22 +98,20 @@ (* monotonicity of Rep_CFun in both arguments in mixfix syntax [_]_ *) (* ------------------------------------------------------------------------ *) -qed_goal "monofun_cfun" thy - "[|f1< f1`x1 << f2`x2" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans_less 1), - (etac monofun_cfun_arg 1), - (etac monofun_cfun_fun 1) - ]); +val prems = goal thy + "[|f1< f1`x1 << f2`x2"; +by (cut_facts_tac prems 1); +by (rtac trans_less 1); +by (etac monofun_cfun_arg 1); +by (etac monofun_cfun_fun 1); +qed "monofun_cfun"; -qed_goal "strictI" thy "f`x = UU ==> f`UU = UU" (fn prems => [ - cut_facts_tac prems 1, - rtac (eq_UU_iff RS iffD2) 1, - etac subst 1, - rtac (minimal RS monofun_cfun_arg) 1]); +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"; (* ------------------------------------------------------------------------ *) @@ -137,13 +119,11 @@ (* use MF2 lemmas from Cont.ML *) (* ------------------------------------------------------------------------ *) -qed_goal "ch2ch_Rep_CFunR" thy - "chain(Y) ==> chain(%i. f`(Y i))" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac (monofun_Rep_CFun2 RS ch2ch_MF2R) 1) - ]); +val prems = goal thy + "chain(Y) ==> chain(%i. f`(Y i))"; +by (cut_facts_tac prems 1); +by (etac (monofun_Rep_CFun2 RS ch2ch_MF2R) 1); +qed "ch2ch_Rep_CFunR"; bind_thm ("ch2ch_Rep_CFunL", monofun_Rep_CFun1 RS ch2ch_MF2L); @@ -155,142 +135,126 @@ (* use MF2 lemmas from Cont.ML *) (* ------------------------------------------------------------------------ *) -qed_goal "lub_cfun_mono" thy - "chain(F) ==> monofun(% x. lub(range(% j.(F j)`x)))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac lub_MF2_mono 1), - (rtac monofun_Rep_CFun1 1), - (rtac (monofun_Rep_CFun2 RS allI) 1), - (atac 1) - ]); +val prems = goal thy + "chain(F) ==> monofun(% x. lub(range(% j.(F j)`x)))"; +by (cut_facts_tac prems 1); +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 *) (* ------------------------------------------------------------------------ *) -qed_goal "ex_lubcfun" thy +val prems = goal thy "[| chain(F); chain(Y) |] ==>\ \ lub(range(%j. lub(range(%i. F(j)`(Y i))))) =\ -\ lub(range(%i. lub(range(%j. F(j)`(Y i)))))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac ex_lubMF2 1), - (rtac monofun_Rep_CFun1 1), - (rtac (monofun_Rep_CFun2 RS allI) 1), - (atac 1), - (atac 1) - ]); +\ lub(range(%i. lub(range(%j. F(j)`(Y i)))))"; +by (cut_facts_tac prems 1); +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 *) (* ------------------------------------------------------------------------ *) -qed_goal "cont_lubcfun" thy - "chain(F) ==> cont(% x. lub(range(% j. F(j)`x)))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac monocontlub2cont 1), - (etac lub_cfun_mono 1), - (rtac contlubI 1), - (strip_tac 1), - (stac (contlub_cfun_arg RS ext) 1), - (atac 1), - (etac ex_lubcfun 1), - (atac 1) - ]); +val prems = goal thy + "chain(F) ==> cont(% x. lub(range(% j. F(j)`x)))"; +by (cut_facts_tac prems 1); +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 *) (* ------------------------------------------------------------------------ *) -qed_goal "lub_cfun" thy - "chain(CCF) ==> range(CCF) <<| (LAM x. lub(range(% i. CCF(i)`x)))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_lubI 1), - (rtac conjI 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (stac less_cfun 1), - (stac Abs_Cfun_inverse2 1), - (etac cont_lubcfun 1), - (rtac (lub_fun RS is_lubE RS conjunct1 RS ub_rangeE RS spec) 1), - (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1), - (strip_tac 1), - (stac less_cfun 1), - (stac Abs_Cfun_inverse2 1), - (etac cont_lubcfun 1), - (rtac (lub_fun RS is_lubE RS conjunct2 RS spec RS mp) 1), - (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1), - (etac (monofun_Rep_CFun1 RS ub2ub_monofun) 1) - ]); +val prems = goal thy + "chain(CCF) ==> range(CCF) <<| (LAM x. lub(range(% i. CCF(i)`x)))"; +by (cut_facts_tac prems 1); +by (rtac is_lubI 1); +by (rtac conjI 1); +by (rtac ub_rangeI 1); +by (rtac allI 1); +by (stac less_cfun 1); +by (stac Abs_Cfun_inverse2 1); +by (etac cont_lubcfun 1); +by (rtac (lub_fun RS is_lubE RS conjunct1 RS ub_rangeE RS spec) 1); +by (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1); +by (strip_tac 1); +by (stac less_cfun 1); +by (stac Abs_Cfun_inverse2 1); +by (etac cont_lubcfun 1); +by (rtac (lub_fun RS is_lubE RS conjunct2 RS spec RS mp) 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))) *) -qed_goal "cpo_cfun" thy - "chain(CCF::nat=>('a->'b)) ==> ? x. range(CCF) <<| x" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac exI 1), - (etac lub_cfun 1) - ]); +val prems = goal thy + "chain(CCF::nat=>('a->'b)) ==> ? x. range(CCF) <<| x"; +by (cut_facts_tac prems 1); +by (rtac exI 1); +by (etac lub_cfun 1); +qed "cpo_cfun"; (* ------------------------------------------------------------------------ *) (* Extensionality in 'a -> 'b *) (* ------------------------------------------------------------------------ *) -qed_goal "ext_cfun" Cfun1.thy "(!!x. f`x = g`x) ==> f = g" - (fn prems => - [ - (res_inst_tac [("t","f")] (Rep_Cfun_inverse RS subst) 1), - (res_inst_tac [("t","g")] (Rep_Cfun_inverse RS subst) 1), - (res_inst_tac [("f","Abs_CFun")] arg_cong 1), - (rtac ext 1), - (resolve_tac prems 1) - ]); +val prems = goal Cfun1.thy "(!!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 *) (* ------------------------------------------------------------------------ *) -qed_goal "semi_monofun_Abs_CFun" thy - "[|cont(f);cont(g);f<Abs_CFun(f)< - [ - (rtac (less_cfun RS iffD2) 1), - (stac Abs_Cfun_inverse2 1), - (resolve_tac prems 1), - (stac Abs_Cfun_inverse2 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); +val prems = goal thy + "[|cont(f);cont(g);f<Abs_CFun(f)< 'b *) (* ------------------------------------------------------------------------ *) -qed_goal "less_cfun2" thy "(!!x. f`x << g`x) ==> f << g" - (fn prems => - [ - (res_inst_tac [("t","f")] (Rep_Cfun_inverse RS subst) 1), - (res_inst_tac [("t","g")] (Rep_Cfun_inverse RS subst) 1), - (rtac semi_monofun_Abs_CFun 1), - (rtac cont_Rep_CFun2 1), - (rtac cont_Rep_CFun2 1), - (rtac (less_fun RS iffD2) 1), - (rtac allI 1), - (resolve_tac prems 1) - ]); +val prems = goal thy "(!!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"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Cfun3.ML --- a/src/HOLCF/Cfun3.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Cfun3.ML Tue Jul 04 15:58:11 2000 +0200 @@ -1,134 +1,111 @@ -(* Title: HOLCF/cfun3.ML +(* Title: HOLCF/Cfun3 ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen + +Class instance of -> for class pcpo *) -open Cfun3; - (* for compatibility with old HOLCF-Version *) -qed_goal "inst_cfun_pcpo" thy "UU = Abs_CFun(%x. UU)" - (fn prems => - [ - (simp_tac (HOL_ss addsimps [UU_def,UU_cfun_def]) 1) - ]); +val prems = goal thy "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 *) (* ------------------------------------------------------------------------ *) -qed_goal "contlub_Rep_CFun1" thy "contlub(Rep_CFun)" -(fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (expand_fun_eq RS iffD2) 1), - (strip_tac 1), - (stac thelub_cfun 1), - (atac 1), - (stac Cfunapp2 1), - (etac cont_lubcfun 1), - (stac thelub_fun 1), - (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1), - (rtac refl 1) - ]); +val prems = goal thy "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 *) (* ------------------------------------------------------------------------ *) -qed_goal "cont_Rep_CFun1" thy "cont(Rep_CFun)" -(fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Rep_CFun1 1), - (rtac contlub_Rep_CFun1 1) - ]); +val prems = goal thy "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 _[_] *) (* ------------------------------------------------------------------------ *) -qed_goal "contlub_cfun_fun" thy +val prems = goal thy "chain(FY) ==>\ -\ lub(range FY)`x = lub(range (%i. FY(i)`x))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans 1), - (etac (contlub_Rep_CFun1 RS contlubE RS spec RS mp RS fun_cong) 1), - (stac thelub_fun 1), - (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1), - (rtac refl 1) - ]); +\ lub(range FY)`x = lub(range (%i. FY(i)`x))"; +by (cut_facts_tac prems 1); +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"; -qed_goal "cont_cfun_fun" thy +val prems = goal thy "chain(FY) ==>\ -\ range(%i. FY(i)`x) <<| lub(range FY)`x" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac thelubE 1), - (etac ch2ch_Rep_CFunL 1), - (etac (contlub_cfun_fun RS sym) 1) - ]); +\ range(%i. FY(i)`x) <<| lub(range FY)`x"; +by (cut_facts_tac prems 1); +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 _[_] *) (* ------------------------------------------------------------------------ *) -qed_goal "contlub_cfun" thy +val prems = goal thy "[|chain(FY);chain(TY)|] ==>\ -\ (lub(range FY))`(lub(range TY)) = lub(range(%i. FY(i)`(TY i)))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac contlub_CF2 1), - (rtac cont_Rep_CFun1 1), - (rtac allI 1), - (rtac cont_Rep_CFun2 1), - (atac 1), - (atac 1) - ]); +\ (lub(range FY))`(lub(range TY)) = lub(range(%i. FY(i)`(TY i)))"; +by (cut_facts_tac prems 1); +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"; -qed_goal "cont_cfun" thy +val prems = goal thy "[|chain(FY);chain(TY)|] ==>\ -\ range(%i.(FY i)`(TY i)) <<| (lub (range FY))`(lub(range TY))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac thelubE 1), - (rtac (monofun_Rep_CFun1 RS ch2ch_MF2LR) 1), - (rtac allI 1), - (rtac monofun_Rep_CFun2 1), - (atac 1), - (atac 1), - (etac (contlub_cfun RS sym) 1), - (atac 1) - ]); +\ range(%i.(FY i)`(TY i)) <<| (lub (range FY))`(lub(range TY))"; +by (cut_facts_tac prems 1); +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 *) (* ------------------------------------------------------------------------ *) -qed_goal "cont2cont_Rep_CFun" thy - "[|cont(%x. ft x);cont(%x. tt x)|] ==> cont(%x. (ft x)`(tt x))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac cont2cont_app2 1), - (rtac cont2cont_app2 1), - (rtac cont_const 1), - (rtac cont_Rep_CFun1 1), - (atac 1), - (rtac cont_Rep_CFun2 1), - (atac 1) - ]); +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"; @@ -136,44 +113,40 @@ (* cont2mono Lemma for %x. LAM y. c1(x)(y) *) (* ------------------------------------------------------------------------ *) -qed_goal "cont2mono_LAM" thy - "[| !!x. cont(c1 x); !!y. monofun(%x. c1 x y)|] ==> monofun(%x. LAM y. c1 x y)" -(fn [p1,p2] => - [ - (rtac monofunI 1), - (strip_tac 1), - (stac less_cfun 1), - (stac less_fun 1), - (rtac allI 1), - (stac beta_cfun 1), - (rtac p1 1), - (stac beta_cfun 1), - (rtac p1 1), - (etac (p2 RS monofunE RS spec RS spec RS mp) 1) - ]); +val [p1,p2] = goal thy + "[| !!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) *) (* ------------------------------------------------------------------------ *) -qed_goal "cont2cont_LAM" thy - "[| !!x. cont(c1 x); !!y. cont(%x. c1 x y) |] ==> cont(%x. LAM y. c1 x y)" -(fn [p1,p2] => - [ - (rtac monocontlub2cont 1), - (rtac (p1 RS cont2mono_LAM) 1), - (rtac (p2 RS cont2mono) 1), - (rtac contlubI 1), - (strip_tac 1), - (stac thelub_cfun 1), - (rtac (p1 RS cont2mono_LAM RS ch2ch_monofun) 1), - (rtac (p2 RS cont2mono) 1), - (atac 1), - (res_inst_tac [("f","Abs_CFun")] arg_cong 1), - (rtac ext 1), - (stac (p1 RS beta_cfun RS ext) 1), - (etac (p2 RS cont2contlub RS contlubE RS spec RS mp) 1) - ]); +val [p1,p2] = goal thy + "[| !!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 *) @@ -193,141 +166,107 @@ (* function application _[_] is strict in its first arguments *) (* ------------------------------------------------------------------------ *) -qed_goal "strict_Rep_CFun1" thy "(UU::'a::cpo->'b)`x = (UU::'b)" - (fn prems => - [ - (stac inst_cfun_pcpo 1), - (stac beta_cfun 1), - (Simp_tac 1), - (rtac refl 1) - ]); +val prems = goal thy "(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 *) (* ------------------------------------------------------------------------ *) -qed_goalw "Istrictify1" thy [Istrictify_def] - "Istrictify(f)(UU)= (UU)" - (fn prems => - [ - (Simp_tac 1) - ]); +val prems = goalw thy [Istrictify_def] + "Istrictify(f)(UU)= (UU)"; +by (Simp_tac 1); +qed "Istrictify1"; -qed_goalw "Istrictify2" thy [Istrictify_def] - "~x=UU ==> Istrictify(f)(x)=f`x" - (fn prems => - [ - (cut_facts_tac prems 1), - (Asm_simp_tac 1) - ]); +val prems = goalw thy [Istrictify_def] + "~x=UU ==> Istrictify(f)(x)=f`x"; +by (cut_facts_tac prems 1); +by (Asm_simp_tac 1); +qed "Istrictify2"; -qed_goal "monofun_Istrictify1" thy "monofun(Istrictify)" - (fn prems => - [ - (rtac monofunI 1), - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (res_inst_tac [("Q","xa=UU")] (excluded_middle RS disjE) 1), - (stac Istrictify2 1), - (atac 1), - (stac Istrictify2 1), - (atac 1), - (rtac monofun_cfun_fun 1), - (atac 1), - (hyp_subst_tac 1), - (stac Istrictify1 1), - (stac Istrictify1 1), - (rtac refl_less 1) - ]); +val prems = goal thy "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"; -qed_goal "monofun_Istrictify2" thy "monofun(Istrictify(f))" - (fn prems => - [ - (rtac monofunI 1), - (strip_tac 1), - (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), - (stac Istrictify2 1), - (etac notUU_I 1), - (atac 1), - (stac Istrictify2 1), - (atac 1), - (rtac monofun_cfun_arg 1), - (atac 1), - (hyp_subst_tac 1), - (stac Istrictify1 1), - (rtac minimal 1) - ]); +val prems = goal thy "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"; -qed_goal "contlub_Istrictify1" thy "contlub(Istrictify)" - (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (expand_fun_eq RS iffD2) 1), - (strip_tac 1), - (stac thelub_fun 1), - (etac (monofun_Istrictify1 RS ch2ch_monofun) 1), - (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), - (stac Istrictify2 1), - (atac 1), - (stac (Istrictify2 RS ext) 1), - (atac 1), - (stac thelub_cfun 1), - (atac 1), - (stac beta_cfun 1), - (rtac cont_lubcfun 1), - (atac 1), - (rtac refl 1), - (hyp_subst_tac 1), - (stac Istrictify1 1), - (stac (Istrictify1 RS ext) 1), - (rtac (chain_UU_I_inverse RS sym) 1), - (rtac (refl RS allI) 1) - ]); +val prems = goal thy "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"; -qed_goal "contlub_Istrictify2" thy "contlub(Istrictify(f::'a -> 'b))" - (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (case_tac "lub(range(Y))=(UU::'a)" 1), - (res_inst_tac [("t","lub(range(Y))")] subst 1), - (rtac sym 1), - (atac 1), - (stac Istrictify1 1), - (rtac sym 1), - (rtac chain_UU_I_inverse 1), - (strip_tac 1), - (res_inst_tac [("t","Y(i)"),("s","UU::'a")] subst 1), - (rtac sym 1), - (rtac (chain_UU_I RS spec) 1), - (atac 1), - (atac 1), - (rtac Istrictify1 1), - (stac Istrictify2 1), - (atac 1), - (res_inst_tac [("s","lub(range(%i. f`(Y i)))")] trans 1), - (rtac contlub_cfun_arg 1), - (atac 1), - (rtac lub_equal2 1), - (rtac (chain_mono2 RS exE) 1), - (atac 2), - (rtac chain_UU_I_inverse2 1), - (atac 1), - (rtac exI 1), - (strip_tac 1), - (rtac (Istrictify2 RS sym) 1), - (fast_tac HOL_cs 1), - (rtac ch2ch_monofun 1), - (rtac monofun_Rep_CFun2 1), - (atac 1), - (rtac ch2ch_monofun 1), - (rtac monofun_Istrictify2 1), - (atac 1) - ]); +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 @@ -337,25 +276,23 @@ (monofun_Istrictify2 RS monocontlub2cont)); -qed_goalw "strictify1" thy [strictify_def] "strictify`f`UU=UU" (fn _ => [ - (stac beta_cfun 1), - (simp_tac (simpset() addsimps [cont_Istrictify2,cont_Istrictify1, - cont2cont_CF1L]) 1), - (stac beta_cfun 1), - (rtac cont_Istrictify2 1), - (rtac Istrictify1 1) - ]); +val _ = goalw thy [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"; -qed_goalw "strictify2" thy [strictify_def] - "~x=UU ==> strictify`f`x=f`x" (fn prems => [ - (stac beta_cfun 1), - (simp_tac (simpset() addsimps [cont_Istrictify2,cont_Istrictify1, - cont2cont_CF1L]) 1), - (stac beta_cfun 1), - (rtac cont_Istrictify2 1), - (rtac Istrictify2 1), - (resolve_tac prems 1) - ]); +val prems = goalw thy [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 (rtac Istrictify2 1); +by (resolve_tac prems 1); +qed "strictify2"; (* ------------------------------------------------------------------------ *) @@ -376,149 +313,134 @@ (* some lemmata for functions with flat/chfin domain/range types *) (* ------------------------------------------------------------------------ *) -qed_goal "chfin_Rep_CFunR" thy - "chain (Y::nat => 'a::cpo->'b::chfin)==> !s. ? n. lub(range(Y))`s = Y n`s" -(fn prems => - [ - cut_facts_tac prems 1, - rtac allI 1, - stac contlub_cfun_fun 1, - atac 1, - fast_tac (HOL_cs addSIs [thelubI,chfin,lub_finch2,chfin2finch,ch2ch_Rep_CFunL])1 - ]); +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 *) (* ------------------------------------------------------------------------ *) -qed_goal "iso_strict" thy +val prems = goal thy "!!f g.[|!y. f`(g`y)=(y::'b) ; !x. g`(f`x)=(x::'a) |] \ -\ ==> f`UU=UU & g`UU=UU" - (fn prems => - [ - (rtac conjI 1), - (rtac UU_I 1), - (res_inst_tac [("s","f`(g`(UU::'b))"),("t","UU::'b")] subst 1), - (etac spec 1), - (rtac (minimal RS monofun_cfun_arg) 1), - (rtac UU_I 1), - (res_inst_tac [("s","g`(f`(UU::'a))"),("t","UU::'a")] subst 1), - (etac spec 1), - (rtac (minimal RS monofun_cfun_arg) 1) - ]); +\ ==> 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"; -qed_goal "isorep_defined" thy - "[|!x. rep`(ab`x)=x;!y. ab`(rep`y)=y; z~=UU|] ==> rep`z ~= UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (etac swap 1), - (dtac notnotD 1), - (dres_inst_tac [("f","ab")] cfun_arg_cong 1), - (etac box_equals 1), - (fast_tac HOL_cs 1), - (etac (iso_strict RS conjunct1) 1), - (atac 1) - ]); +val prems = goal thy + "[|!x. rep`(ab`x)=x;!y. ab`(rep`y)=y; z~=UU|] ==> rep`z ~= UU"; +by (cut_facts_tac prems 1); +by (etac swap 1); +by (dtac notnotD 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"; -qed_goal "isoabs_defined" thy - "[|!x. rep`(ab`x) = x;!y. ab`(rep`y)=y ; z~=UU|] ==> ab`z ~= UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (etac swap 1), - (dtac notnotD 1), - (dres_inst_tac [("f","rep")] cfun_arg_cong 1), - (etac box_equals 1), - (fast_tac HOL_cs 1), - (etac (iso_strict RS conjunct2) 1), - (atac 1) - ]); +val prems = goal thy + "[|!x. rep`(ab`x) = x;!y. ab`(rep`y)=y ; z~=UU|] ==> ab`z ~= UU"; +by (cut_facts_tac prems 1); +by (etac swap 1); +by (dtac notnotD 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 *) (* ------------------------------------------------------------------------ *) -qed_goal "chfin2chfin" thy "!!f g.[|! Y::nat=>'a. chain Y --> (? n. max_in_chain n Y); \ +val prems = goal thy "!!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)" - (fn prems => - [ - (rewtac max_in_chain_def), - (strip_tac 1), - (rtac exE 1), - (res_inst_tac [("P","chain(%i. g`(Y i))")] mp 1), - (etac spec 1), - (etac ch2ch_Rep_CFunR 1), - (rtac exI 1), - (strip_tac 1), - (res_inst_tac [("s","f`(g`(Y x))"),("t","Y(x)")] subst 1), - (etac spec 1), - (res_inst_tac [("s","f`(g`(Y j))"),("t","Y(j)")] subst 1), - (etac spec 1), - (rtac cfun_arg_cong 1), - (rtac mp 1), - (etac spec 1), - (atac 1) - ]); +\ ==> ! 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"; -qed_goal "flat2flat" thy "!!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" - (fn prems => - [ - (strip_tac 1), - (rtac disjE 1), - (res_inst_tac [("P","g`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)" - (fn prems => - [ - (cut_facts_tac prems 1), - (case_tac "f`(x::'a)=(UU::'b)" 1), - (rtac disjI1 1), - (rtac UU_I 1), - (res_inst_tac [("s","f`(x)"),("t","UU::'b")] subst 1), - (atac 1), - (rtac (minimal RS monofun_cfun_arg) 1), - (case_tac "f`(UU::'a)=(UU::'b)" 1), - (etac disjI1 1), - (rtac disjI2 1), - (rtac allI 1), - (hyp_subst_tac 1), - (res_inst_tac [("a","f`(UU::'a)")] (refl RS box_equals) 1), - (res_inst_tac [("fo5","f")] ((minimal RS monofun_cfun_arg) RS - (ax_flat RS spec RS spec RS mp) RS disjE) 1), - (contr_tac 1),(atac 1), - (res_inst_tac [("fo5","f")] ((minimal RS monofun_cfun_arg) RS - (ax_flat RS spec RS spec RS mp) RS disjE) 1), - (contr_tac 1),(atac 1) -]); +val prems = goal thy +"f`(x::'a)=(c::'b::flat) ==> f`(UU::'a)=(UU::'b) | (!z. f`(z::'a)=c)"; +by (cut_facts_tac prems 1); +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"; (* ------------------------------------------------------------------------ *) @@ -526,28 +448,26 @@ (* ------------------------------------------------------------------------ *) -qed_goalw "ID1" thy [ID_def] "ID`x=x" - (fn prems => - [ - (stac beta_cfun 1), - (rtac cont_id 1), - (rtac refl 1) - ]); +val prems = goalw thy [ID_def] "ID`x=x"; +by (stac beta_cfun 1); +by (rtac cont_id 1); +by (rtac refl 1); +qed "ID1"; -qed_goalw "cfcomp1" thy [oo_def] "(f oo g)=(LAM x. f`(g`x))" (fn _ => [ - (stac beta_cfun 1), - (Simp_tac 1), - (stac beta_cfun 1), - (Simp_tac 1), - (rtac refl 1) - ]); +val _ = goalw thy [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"; -qed_goal "cfcomp2" thy "(f oo g)`x=f`(g`x)" (fn _ => [ - (stac cfcomp1 1), - (stac beta_cfun 1), - (Simp_tac 1), - (rtac refl 1) - ]); +val _ = goal thy "(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"; (* ------------------------------------------------------------------------ *) @@ -559,37 +479,31 @@ (* ------------------------------------------------------------------------ *) -qed_goal "ID2" thy "f oo ID = f " - (fn prems => - [ - (rtac ext_cfun 1), - (stac cfcomp2 1), - (stac ID1 1), - (rtac refl 1) - ]); +val prems = goal thy "f oo ID = f "; +by (rtac ext_cfun 1); +by (stac cfcomp2 1); +by (stac ID1 1); +by (rtac refl 1); +qed "ID2"; -qed_goal "ID3" thy "ID oo f = f " - (fn prems => - [ - (rtac ext_cfun 1), - (stac cfcomp2 1), - (stac ID1 1), - (rtac refl 1) - ]); +val prems = goal thy "ID oo f = f "; +by (rtac ext_cfun 1); +by (stac cfcomp2 1); +by (stac ID1 1); +by (rtac refl 1); +qed "ID3"; -qed_goal "assoc_oo" thy "f oo (g oo h) = (f oo g) oo h" - (fn prems => - [ - (rtac ext_cfun 1), - (res_inst_tac [("s","f`(g`(h`x))")] trans 1), - (stac cfcomp2 1), - (stac cfcomp2 1), - (rtac refl 1), - (stac cfcomp2 1), - (stac cfcomp2 1), - (rtac refl 1) - ]); +val prems = goal thy "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 *) diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Cont.ML --- a/src/HOLCF/Cont.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Cont.ML Tue Jul 04 15:58:11 2000 +0200 @@ -10,57 +10,45 @@ (* access to definition *) (* ------------------------------------------------------------------------ *) -qed_goalw "contlubI" thy [contlub] +val prems = goalw thy [contlub] "! Y. chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))==>\ -\ contlub(f)" -(fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); +\ contlub(f)"; +by (cut_facts_tac prems 1); +by (atac 1); +qed "contlubI"; -qed_goalw "contlubE" thy [contlub] +val prems = goalw thy [contlub] " contlub(f)==>\ -\ ! Y. chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))" -(fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); +\ ! Y. chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))"; +by (cut_facts_tac prems 1); +by (atac 1); +qed "contlubE"; -qed_goalw "contI" thy [cont] - "! Y. chain(Y) --> range(% i. f(Y(i))) <<| f(lub(range(Y))) ==> cont(f)" -(fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); +val prems = goalw thy [cont] + "! Y. chain(Y) --> range(% i. f(Y(i))) <<| f(lub(range(Y))) ==> cont(f)"; +by (cut_facts_tac prems 1); +by (atac 1); +qed "contI"; -qed_goalw "contE" thy [cont] - "cont(f) ==> ! Y. chain(Y) --> range(% i. f(Y(i))) <<| f(lub(range(Y)))" -(fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); +val prems = goalw thy [cont] + "cont(f) ==> ! Y. chain(Y) --> range(% i. f(Y(i))) <<| f(lub(range(Y)))"; +by (cut_facts_tac prems 1); +by (atac 1); +qed "contE"; -qed_goalw "monofunI" thy [monofun] - "! x y. x << y --> f(x) << f(y) ==> monofun(f)" -(fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); +val prems = goalw thy [monofun] + "! x y. x << y --> f(x) << f(y) ==> monofun(f)"; +by (cut_facts_tac prems 1); +by (atac 1); +qed "monofunI"; -qed_goalw "monofunE" thy [monofun] - "monofun(f) ==> ! x y. x << y --> f(x) << f(y)" -(fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); +val prems = goalw thy [monofun] + "monofun(f) ==> ! x y. x << y --> f(x) << f(y)"; +by (cut_facts_tac prems 1); +by (atac 1); +qed "monofunE"; (* ------------------------------------------------------------------------ *) (* the main purpose of cont.thy is to show: *) @@ -71,115 +59,91 @@ (* monotone functions map chains to chains *) (* ------------------------------------------------------------------------ *) -qed_goal "ch2ch_monofun" thy - "[| monofun(f); chain(Y) |] ==> chain(%i. f(Y(i)))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac chainI 1), - (rtac allI 1), - (etac (monofunE RS spec RS spec RS mp) 1), - (etac (chainE RS spec) 1) - ]); +val prems = goal thy + "[| monofun(f); chain(Y) |] ==> chain(%i. f(Y(i)))"; +by (cut_facts_tac prems 1); +by (rtac chainI 1); +by (rtac allI 1); +by (etac (monofunE RS spec RS spec RS mp) 1); +by (etac (chainE RS spec) 1); +qed "ch2ch_monofun"; (* ------------------------------------------------------------------------ *) (* monotone functions map upper bound to upper bounds *) (* ------------------------------------------------------------------------ *) -qed_goal "ub2ub_monofun" thy - "[| monofun(f); range(Y) <| u|] ==> range(%i. f(Y(i))) <| f(u)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (etac (monofunE RS spec RS spec RS mp) 1), - (etac (ub_rangeE RS spec) 1) - ]); +val prems = goal thy + "[| monofun(f); range(Y) <| u|] ==> range(%i. f(Y(i))) <| f(u)"; +by (cut_facts_tac prems 1); +by (rtac ub_rangeI 1); +by (rtac allI 1); +by (etac (monofunE RS spec RS spec RS mp) 1); +by (etac (ub_rangeE RS spec) 1); +qed "ub2ub_monofun"; (* ------------------------------------------------------------------------ *) (* left to right: monofun(f) & contlub(f) ==> cont(f) *) (* ------------------------------------------------------------------------ *) -qed_goalw "monocontlub2cont" thy [cont] - "[|monofun(f);contlub(f)|] ==> cont(f)" -(fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (rtac thelubE 1), - (etac ch2ch_monofun 1), - (atac 1), - (etac (contlubE RS spec RS mp RS sym) 1), - (atac 1) - ]); +val prems = goalw thy [cont] + "[|monofun(f);contlub(f)|] ==> cont(f)"; +by (cut_facts_tac prems 1); +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 *) (* ------------------------------------------------------------------------ *) -qed_goal "binchain_cont" thy -"[| cont(f); x << y |] ==> range(%i::nat. f(if i = 0 then x else y)) <<| f(y)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac subst 1), - (etac (contE RS spec RS mp) 2), - (etac bin_chain 2), - (res_inst_tac [("y","y")] arg_cong 1), - (etac (lub_bin_chain RS thelubI) 1) - ]); +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 *) (* ------------------------------------------------------------------------ *) -qed_goalw "cont2mono" thy [monofun] - "cont(f) ==> monofun(f)" -(fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (res_inst_tac [("s","if 0 = 0 then x else y")] subst 1), - (rtac (binchain_cont RS is_ub_lub) 2), - (atac 2), - (atac 2), - (Simp_tac 1) - ]); +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) *) (* ------------------------------------------------------------------------ *) -qed_goalw "cont2contlub" thy [contlub] - "cont(f) ==> contlub(f)" -(fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (rtac (thelubI RS sym) 1), - (etac (contE RS spec RS mp) 1), - (atac 1) - ]); +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 *) +(* monotone functions map finite chains to finite chains *) (* ------------------------------------------------------------------------ *) -qed_goalw "monofun_finch2finch" thy [finite_chain_def] - "[| monofun f; finite_chain Y |] ==> finite_chain (%n. f (Y n))" -(fn prems => - [ - cut_facts_tac prems 1, - safe_tac HOL_cs, - fast_tac (HOL_cs addSEs [ch2ch_monofun]) 1, - fast_tac (HOL_cs addss (HOL_ss addsimps [max_in_chain_def])) 1 - ]); +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 *) +(* The same holds for continuous functions *) (* ------------------------------------------------------------------------ *) bind_thm ("cont_finch2finch", cont2mono RS monofun_finch2finch); @@ -191,199 +155,181 @@ (* in both arguments *) (* ------------------------------------------------------------------------ *) -qed_goal "ch2ch_MF2L" thy -"[|monofun(MF2); chain(F)|] ==> chain(%i. MF2 (F i) x)" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac (ch2ch_monofun RS ch2ch_fun) 1), - (atac 1) - ]); +val prems = goal thy +"[|monofun(MF2); chain(F)|] ==> chain(%i. MF2 (F i) x)"; +by (cut_facts_tac prems 1); +by (etac (ch2ch_monofun RS ch2ch_fun) 1); +by (atac 1); +qed "ch2ch_MF2L"; -qed_goal "ch2ch_MF2R" thy -"[|monofun(MF2(f)); chain(Y)|] ==> chain(%i. MF2 f (Y i))" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac ch2ch_monofun 1), - (atac 1) - ]); +val prems = goal thy +"[|monofun(MF2(f)); chain(Y)|] ==> chain(%i. MF2 f (Y i))"; +by (cut_facts_tac prems 1); +by (etac ch2ch_monofun 1); +by (atac 1); +qed "ch2ch_MF2R"; -qed_goal "ch2ch_MF2LR" thy +val prems = goal thy "[|monofun(MF2); !f. monofun(MF2(f)); chain(F); chain(Y)|] ==> \ -\ chain(%i. MF2(F(i))(Y(i)))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac chainI 1), - (strip_tac 1 ), - (rtac trans_less 1), - (etac (ch2ch_MF2L RS chainE RS spec) 1), - (atac 1), - ((rtac (monofunE RS spec RS spec RS mp) 1) THEN (etac spec 1)), - (etac (chainE RS spec) 1) - ]); +\ chain(%i. MF2(F(i))(Y(i)))"; +by (cut_facts_tac prems 1); +by (rtac chainI 1); +by (strip_tac 1 ); +by (rtac trans_less 1); +by (etac (ch2ch_MF2L RS chainE RS spec) 1); +by (atac 1); +by ((rtac (monofunE RS spec RS spec RS mp) 1) THEN (etac spec 1)); +by (etac (chainE RS spec) 1); +qed "ch2ch_MF2LR"; -qed_goal "ch2ch_lubMF2R" thy +val prems = goal thy "[|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))))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac (lub_mono RS allI RS chainI) 1), - ((rtac ch2ch_MF2R 1) THEN (etac spec 1)), - (atac 1), - ((rtac ch2ch_MF2R 1) THEN (etac spec 1)), - (atac 1), - (strip_tac 1), - (rtac (chainE RS spec) 1), - (etac ch2ch_MF2L 1), - (atac 1) - ]); +\ chain(%j. lub(range(%i. MF2 (F j) (Y i))))"; +by (cut_facts_tac prems 1); +by (rtac (lub_mono RS allI 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 RS spec) 1); +by (etac ch2ch_MF2L 1); +by (atac 1); +qed "ch2ch_lubMF2R"; -qed_goal "ch2ch_lubMF2L" thy +val prems = goal thy "[|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))))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac (lub_mono RS allI RS chainI) 1), - (etac ch2ch_MF2L 1), - (atac 1), - (etac ch2ch_MF2L 1), - (atac 1), - (strip_tac 1), - (rtac (chainE RS spec) 1), - ((rtac ch2ch_MF2R 1) THEN (etac spec 1)), - (atac 1) - ]); +\ chain(%i. lub(range(%j. MF2 (F j) (Y i))))"; +by (cut_facts_tac prems 1); +by (rtac (lub_mono RS allI 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 RS spec) 1); +by ((rtac ch2ch_MF2R 1) THEN (etac spec 1)); +by (atac 1); +qed "ch2ch_lubMF2L"; -qed_goal "lub_MF2_mono" thy +val prems = goal thy "[|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))))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac monofunI 1), - (strip_tac 1), - (rtac lub_mono 1), - (etac ch2ch_MF2L 1), - (atac 1), - (etac ch2ch_MF2L 1), - (atac 1), - (strip_tac 1), - ((rtac (monofunE RS spec RS spec RS mp) 1) THEN (etac spec 1)), - (atac 1) - ]); +\ monofun(% x. lub(range(% j. MF2 (F j) (x))))"; +by (cut_facts_tac prems 1); +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"; -qed_goal "ex_lubMF2" thy +val prems = goal thy "[|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)))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac antisym_less 1), - (rtac (ub_rangeI RSN (2,is_lub_thelub)) 1), - (etac ch2ch_lubMF2R 1), - (REPEAT (atac 1)), - (strip_tac 1), - (rtac lub_mono 1), - ((rtac ch2ch_MF2R 1) THEN (etac spec 1)), - (atac 1), - (etac ch2ch_lubMF2L 1), - (REPEAT (atac 1)), - (strip_tac 1), - (rtac is_ub_thelub 1), - (etac ch2ch_MF2L 1), - (atac 1), - (rtac (ub_rangeI RSN (2,is_lub_thelub)) 1), - (etac ch2ch_lubMF2L 1), - (REPEAT (atac 1)), - (strip_tac 1), - (rtac lub_mono 1), - (etac ch2ch_MF2L 1), - (atac 1), - (etac ch2ch_lubMF2R 1), - (REPEAT (atac 1)), - (strip_tac 1), - (rtac is_ub_thelub 1), - ((rtac ch2ch_MF2R 1) THEN (etac spec 1)), - (atac 1) - ]); +\ lub(range(%i. lub(range(%j. MF2(F j) (Y i)))))"; +by (cut_facts_tac prems 1); +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"; -qed_goal "diag_lubMF2_1" thy +val prems = goal thy "[|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))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac antisym_less 1), - (rtac (ub_rangeI RSN (2,is_lub_thelub)) 1), - (etac ch2ch_lubMF2L 1), - (REPEAT (atac 1)), - (strip_tac 1 ), - (rtac lub_mono3 1), - (etac ch2ch_MF2L 1), - (REPEAT (atac 1)), - (etac ch2ch_MF2LR 1), - (REPEAT (atac 1)), - (rtac allI 1), - (res_inst_tac [("m","i"),("n","ia")] nat_less_cases 1), - (res_inst_tac [("x","ia")] exI 1), - (rtac (chain_mono RS mp) 1), - (etac allE 1), - (etac ch2ch_MF2R 1), - (REPEAT (atac 1)), - (hyp_subst_tac 1), - (res_inst_tac [("x","ia")] exI 1), - (rtac refl_less 1), - (res_inst_tac [("x","i")] exI 1), - (rtac (chain_mono RS mp) 1), - (etac ch2ch_MF2L 1), - (REPEAT (atac 1)), - (rtac lub_mono 1), - (etac ch2ch_MF2LR 1), - (REPEAT(atac 1)), - (etac ch2ch_lubMF2L 1), - (REPEAT (atac 1)), - (strip_tac 1 ), - (rtac is_ub_thelub 1), - (etac ch2ch_MF2L 1), - (atac 1) - ]); +\ lub(range(%i. MF2(FY(i))(TY(i))))"; +by (cut_facts_tac prems 1); +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 RS mp) 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 RS mp) 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"; -qed_goal "diag_lubMF2_2" thy +val prems = goal thy "[|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))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans 1), - (rtac ex_lubMF2 1), - (REPEAT (atac 1)), - (etac diag_lubMF2_1 1), - (REPEAT (atac 1)) - ]); +\ lub(range(%i. MF2(FY(i))(TY(i))))"; +by (cut_facts_tac prems 1); +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"; (* ------------------------------------------------------------------------ *) @@ -391,59 +337,51 @@ (* in both arguments *) (* ------------------------------------------------------------------------ *) -qed_goal "contlub_CF2" thy +val prems = 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))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac ((hd prems) RS cont2contlub RS contlubE RS spec RS mp) 1), - (atac 1), - (stac thelub_fun 1), - (rtac ((hd prems) RS cont2mono RS ch2ch_monofun) 1), - (atac 1), - (rtac trans 1), - (rtac (((hd (tl prems)) RS spec RS cont2contlub) RS contlubE RS spec RS mp RS ext RS arg_cong RS arg_cong) 1), - (atac 1), - (rtac diag_lubMF2_2 1), - (etac cont2mono 1), - (rtac allI 1), - (etac allE 1), - (etac cont2mono 1), - (REPEAT (atac 1)) - ]); +\ CF2(lub(range(FY)))(lub(range(TY))) = lub(range(%i. CF2(FY(i))(TY(i))))"; +by (cut_facts_tac prems 1); +by (stac ((hd prems) RS cont2contlub RS contlubE RS spec RS mp) 1); +by (atac 1); +by (stac thelub_fun 1); +by (rtac ((hd prems) RS cont2mono RS ch2ch_monofun) 1); +by (atac 1); +by (rtac trans 1); +by (rtac (((hd (tl prems)) RS spec RS cont2contlub) RS contlubE RS spec RS mp RS ext RS arg_cong RS arg_cong) 1); +by (atac 1); +by (rtac diag_lubMF2_2 1); +by (etac cont2mono 1); +by (rtac allI 1); +by (etac allE 1); +by (etac cont2mono 1); +by (REPEAT (atac 1)); +qed "contlub_CF2"; (* ------------------------------------------------------------------------ *) (* The following results are about application for functions in 'a=>'b *) (* ------------------------------------------------------------------------ *) -qed_goal "monofun_fun_fun" thy - "f1 << f2 ==> f1(x) << f2(x)" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac (less_fun RS iffD1 RS spec) 1) - ]); +val prems = goal thy + "f1 << f2 ==> f1(x) << f2(x)"; +by (cut_facts_tac prems 1); +by (etac (less_fun RS iffD1 RS spec) 1); +qed "monofun_fun_fun"; -qed_goal "monofun_fun_arg" thy - "[|monofun(f); x1 << x2|] ==> f(x1) << f(x2)" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac (monofunE RS spec RS spec RS mp) 1), - (atac 1) - ]); +val prems = goal thy + "[|monofun(f); x1 << x2|] ==> f(x1) << f(x2)"; +by (cut_facts_tac prems 1); +by (etac (monofunE RS spec RS spec RS mp) 1); +by (atac 1); +qed "monofun_fun_arg"; -qed_goal "monofun_fun" thy -"[|monofun(f1); monofun(f2); f1 << f2; x1 << x2|] ==> f1(x1) << f2(x2)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans_less 1), - (etac monofun_fun_arg 1), - (atac 1), - (etac monofun_fun_fun 1) - ]); +val prems = goal thy +"[|monofun(f1); monofun(f2); f1 << f2; x1 << x2|] ==> f1(x1) << f2(x2)"; +by (cut_facts_tac prems 1); +by (rtac trans_less 1); +by (etac monofun_fun_arg 1); +by (atac 1); +by (etac monofun_fun_fun 1); +qed "monofun_fun"; (* ------------------------------------------------------------------------ *) @@ -451,147 +389,115 @@ (* continuity *) (* ------------------------------------------------------------------------ *) -qed_goal "mono2mono_MF1L" thy - "[|monofun(c1)|] ==> monofun(%x. c1 x y)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac monofunI 1), - (strip_tac 1), - (etac (monofun_fun_arg RS monofun_fun_fun) 1), - (atac 1) - ]); +val prems = goal thy + "[|monofun(c1)|] ==> monofun(%x. c1 x y)"; +by (cut_facts_tac prems 1); +by (rtac monofunI 1); +by (strip_tac 1); +by (etac (monofun_fun_arg RS monofun_fun_fun) 1); +by (atac 1); +qed "mono2mono_MF1L"; -qed_goal "cont2cont_CF1L" thy - "[|cont(c1)|] ==> cont(%x. c1 x y)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac monocontlub2cont 1), - (etac (cont2mono RS mono2mono_MF1L) 1), - (rtac contlubI 1), - (strip_tac 1), - (rtac ((hd prems) RS cont2contlub RS - contlubE RS spec RS mp RS ssubst) 1), - (atac 1), - (stac thelub_fun 1), - (rtac ch2ch_monofun 1), - (etac cont2mono 1), - (atac 1), - (rtac refl 1) - ]); +val prems = goal thy + "[|cont(c1)|] ==> cont(%x. c1 x y)"; +by (cut_facts_tac prems 1); +by (rtac monocontlub2cont 1); +by (etac (cont2mono RS mono2mono_MF1L) 1); +by (rtac contlubI 1); +by (strip_tac 1); +by (rtac ((hd prems) RS 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" ***********) -qed_goal "mono2mono_MF1L_rev" thy - "!y. monofun(%x. c1 x y) ==> monofun(c1)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac monofunI 1), - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (rtac ((hd prems) RS spec RS monofunE RS spec RS spec RS mp) 1), - (atac 1) - ]); +val prems = goal thy + "!y. monofun(%x. c1 x y) ==> monofun(c1)"; +by (cut_facts_tac prems 1); +by (rtac monofunI 1); +by (strip_tac 1); +by (rtac (less_fun RS iffD2) 1); +by (strip_tac 1); +by (rtac ((hd prems) RS spec RS monofunE RS spec RS spec RS mp) 1); +by (atac 1); +qed "mono2mono_MF1L_rev"; -qed_goal "cont2cont_CF1L_rev" thy - "!y. cont(%x. c1 x y) ==> cont(c1)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac monocontlub2cont 1), - (rtac (cont2mono RS allI RS mono2mono_MF1L_rev ) 1), - (etac spec 1), - (rtac contlubI 1), - (strip_tac 1), - (rtac ext 1), - (stac thelub_fun 1), - (rtac (cont2mono RS allI RS mono2mono_MF1L_rev RS ch2ch_monofun) 1), - (etac spec 1), - (atac 1), - (rtac - ((hd prems) RS spec RS cont2contlub RS contlubE RS spec RS mp) 1), - (atac 1) - ]); +val prems = goal thy + "!y. cont(%x. c1 x y) ==> cont(c1)"; +by (cut_facts_tac prems 1); +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 (rtac ((hd prems) RS spec RS cont2contlub RS contlubE RS spec RS mp) 1); +by (atac 1); +qed "cont2cont_CF1L_rev"; (* ------------------------------------------------------------------------ *) (* What D.A.Schmidt calls continuity of abstraction *) (* never used here *) (* ------------------------------------------------------------------------ *) -qed_goal "contlub_abstraction" thy +val prems = goal thy "[|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)))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans 1), - (rtac (cont2contlub RS contlubE RS spec RS mp) 2), - (atac 3), - (etac cont2cont_CF1L_rev 2), - (rtac ext 1), - (rtac (cont2contlub RS contlubE RS spec RS mp RS sym) 1), - (etac spec 1), - (atac 1) - ]); +\ (%y. lub(range(%i. c (Y i) y))) = (lub(range(%i.%y. c (Y i) y)))"; +by (cut_facts_tac prems 1); +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"; -qed_goal "mono2mono_app" thy +val prems = goal thy "[|monofun(ft);!x. monofun(ft(x));monofun(tt)|] ==>\ -\ monofun(%x.(ft(x))(tt(x)))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac monofunI 1), - (strip_tac 1), - (res_inst_tac [("f1.0","ft(x)"),("f2.0","ft(y)")] monofun_fun 1), - (etac spec 1), - (etac spec 1), - (etac (monofunE RS spec RS spec RS mp) 1), - (atac 1), - (etac (monofunE RS spec RS spec RS mp) 1), - (atac 1) - ]); +\ monofun(%x.(ft(x))(tt(x)))"; +by (cut_facts_tac prems 1); +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"; -qed_goal "cont2contlub_app" thy -"[|cont(ft);!x. cont(ft(x));cont(tt)|] ==> contlub(%x.(ft(x))(tt(x)))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac contlubI 1), - (strip_tac 1), - (res_inst_tac [("f3","tt")] (contlubE RS spec RS mp RS ssubst) 1), - (etac cont2contlub 1), - (atac 1), - (rtac contlub_CF2 1), - (REPEAT (atac 1)), - (etac (cont2mono RS ch2ch_monofun) 1), - (atac 1) - ]); +val prems = goal thy +"[|cont(ft);!x. cont(ft(x));cont(tt)|] ==> contlub(%x.(ft(x))(tt(x)))"; +by (cut_facts_tac prems 1); +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"; -qed_goal "cont2cont_app" thy -"[|cont(ft);!x. cont(ft(x));cont(tt)|] ==>\ -\ cont(%x.(ft(x))(tt(x)))" - (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac mono2mono_app 1), - (rtac cont2mono 1), - (resolve_tac prems 1), - (strip_tac 1), - (rtac cont2mono 1), - (cut_facts_tac prems 1), - (etac spec 1), - (rtac cont2mono 1), - (resolve_tac prems 1), - (rtac cont2contlub_app 1), - (resolve_tac prems 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); +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)); @@ -603,69 +509,54 @@ (* The identity function is continuous *) (* ------------------------------------------------------------------------ *) -qed_goal "cont_id" thy "cont(% x. x)" - (fn prems => - [ - (rtac contI 1), - (strip_tac 1), - (etac thelubE 1), - (rtac refl 1) - ]); +val prems = goal thy "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 *) (* ------------------------------------------------------------------------ *) -qed_goalw "cont_const" thy [cont] "cont(%x. c)" - (fn prems => - [ - (strip_tac 1), - (rtac is_lubI 1), - (rtac conjI 1), - (rtac ub_rangeI 1), - (strip_tac 1), - (rtac refl_less 1), - (strip_tac 1), - (dtac ub_rangeE 1), - (etac spec 1) - ]); +val prems = goalw thy [cont] "cont(%x. c)"; +by (strip_tac 1); +by (rtac is_lubI 1); +by (rtac conjI 1); +by (rtac ub_rangeI 1); +by (strip_tac 1); +by (rtac refl_less 1); +by (strip_tac 1); +by (dtac ub_rangeE 1); +by (etac spec 1); +qed "cont_const"; -qed_goal "cont2cont_app3" thy - "[|cont(f);cont(t) |] ==> cont(%x. f(t(x)))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac cont2cont_app2 1), - (rtac cont_const 1), - (atac 1), - (atac 1) - ]); +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 *) (* ------------------------------------------------------------------------ *) -qed_goal "CfunI" thy "?x:Collect cont" - (fn prems => - [ - (rtac CollectI 1), - (rtac cont_const 1) - ]); +val prems = goal thy "?x:Collect cont"; +by (rtac CollectI 1); +by (rtac cont_const 1); +qed "CfunI"; (* ------------------------------------------------------------------------ *) -(* some properties of flat *) +(* some properties of flat *) (* ------------------------------------------------------------------------ *) -qed_goalw "flatdom2monofun" thy [monofun] - "f UU = UU ==> monofun (f::'a::flat=>'b::pcpo)" -(fn prems => - [ - cut_facts_tac prems 1, - strip_tac 1, - dtac (ax_flat RS spec RS spec RS mp) 1, - fast_tac ((HOL_cs addss (simpset() addsimps [minimal]))) 1 - ]); +val prems = goalw thy [monofun] + "f UU = UU ==> monofun (f::'a::flat=>'b::pcpo)"; +by (cut_facts_tac prems 1); +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)"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Cprod1.ML --- a/src/HOLCF/Cprod1.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Cprod1.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,46 +3,37 @@ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for theory Cprod1.thy +Partial ordering for cartesian product of HOL theory Prod.thy *) -open Cprod1; - (* ------------------------------------------------------------------------ *) (* less_cprod is a partial order on 'a * 'b *) (* ------------------------------------------------------------------------ *) -qed_goal "Sel_injective_cprod" Prod.thy - "[|fst x = fst y; snd x = snd y|] ==> x = y" -(fn prems => - [ - (cut_facts_tac prems 1), - (subgoal_tac "(fst x,snd x)=(fst y,snd y)" 1), - (rotate_tac ~1 1), - (asm_full_simp_tac(HOL_ss addsimps[surjective_pairing RS sym])1), - (asm_simp_tac (simpset_of Prod.thy) 1) - ]); +val prems = goal Prod.thy + "[|fst x = fst y; snd x = snd y|] ==> x = y"; +by (cut_facts_tac prems 1); +by (subgoal_tac "(fst x,snd x)=(fst y,snd y)" 1); +by (rotate_tac ~1 1); +by (asm_full_simp_tac(HOL_ss addsimps[surjective_pairing RS sym])1); +by (asm_simp_tac (simpset_of Prod.thy) 1); +qed "Sel_injective_cprod"; -qed_goalw "refl_less_cprod" Cprod1.thy [less_cprod_def] "(p::'a*'b) << p" - (fn prems => [Simp_tac 1]); +val prems = goalw Cprod1.thy [less_cprod_def] "(p::'a*'b) << p"; +by (Simp_tac 1); +qed "refl_less_cprod"; -qed_goalw "antisym_less_cprod" thy [less_cprod_def] - "[|(p1::'a * 'b) << p2;p2 << p1|] ==> p1=p2" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac Sel_injective_cprod 1), - (fast_tac (HOL_cs addIs [antisym_less]) 1), - (fast_tac (HOL_cs addIs [antisym_less]) 1) - ]); +Goalw [less_cprod_def] "[|(p1::'a * 'b) << p2;p2 << p1|] ==> p1=p2"; +by (rtac Sel_injective_cprod 1); +by (fast_tac (HOL_cs addIs [antisym_less]) 1); +by (fast_tac (HOL_cs addIs [antisym_less]) 1); +qed "antisym_less_cprod"; -qed_goalw "trans_less_cprod" thy [less_cprod_def] - "[|(p1::'a*'b) << p2;p2 << p3|] ==> p1 << p3" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - (fast_tac (HOL_cs addIs [trans_less]) 1), - (fast_tac (HOL_cs addIs [trans_less]) 1) - ]); +val prems = goalw thy [less_cprod_def] + "[|(p1::'a*'b) << p2;p2 << p3|] ==> p1 << p3"; +by (cut_facts_tac prems 1); +by (rtac conjI 1); +by (fast_tac (HOL_cs addIs [trans_less]) 1); +by (fast_tac (HOL_cs addIs [trans_less]) 1); +qed "trans_less_cprod"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Cprod2.ML --- a/src/HOLCF/Cprod2.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Cprod2.ML Tue Jul 04 15:58:11 2000 +0200 @@ -1,143 +1,120 @@ -(* Title: HOLCF/cprod2.ML +(* Title: HOLCF/Cprod2 ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for cprod2.thy +Class Instance *::(pcpo,pcpo)po *) -open Cprod2; - (* for compatibility with old HOLCF-Version *) -qed_goal "inst_cprod_po" thy "(op <<)=(%x y. fst x< - [ - (fold_goals_tac [less_cprod_def]), - (rtac refl 1) - ]); +val prems = goal thy "(op <<)=(%x y. fst x< x1 << x2 & y1 << y2" - (fn prems => - [ - (cut_facts_tac prems 1), - (etac conjE 1), - (dtac (fst_conv RS subst) 1), - (dtac (fst_conv RS subst) 1), - (dtac (fst_conv RS subst) 1), - (dtac (snd_conv RS subst) 1), - (dtac (snd_conv RS subst) 1), - (dtac (snd_conv RS subst) 1), - (rtac conjI 1), - (atac 1), - (atac 1) - ]); +val prems = goalw thy [inst_cprod_po RS eq_reflection] + "(x1,y1) << (x2,y2) ==> x1 << x2 & y1 << y2"; +by (cut_facts_tac prems 1); +by (etac conjE 1); +by (dtac (fst_conv RS subst) 1); +by (dtac (fst_conv RS subst) 1); +by (dtac (fst_conv RS subst) 1); +by (dtac (snd_conv RS subst) 1); +by (dtac (snd_conv RS subst) 1); +by (dtac (snd_conv RS subst) 1); +by (rtac conjI 1); +by (atac 1); +by (atac 1); +qed "less_cprod4c"; (* ------------------------------------------------------------------------ *) (* type cprod is pointed *) (* ------------------------------------------------------------------------ *) -qed_goal "minimal_cprod" thy "(UU,UU)< - [ - (simp_tac(simpset() addsimps[inst_cprod_po])1) - ]); +val prems = goal thy "(UU,UU)< - [ - (res_inst_tac [("x","(UU,UU)")] exI 1), - (rtac (minimal_cprod RS allI) 1) - ]); +val prems = goal thy "? x::'a*'b.!y. x< is monotone in both arguments *) (* ------------------------------------------------------------------------ *) -qed_goalw "monofun_pair1" thy [monofun] "monofun Pair" - (fn prems => - [ - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (asm_simp_tac (simpset() addsimps [inst_cprod_po]) 1) - ]); +val prems = goalw thy [monofun] "monofun Pair"; +by (strip_tac 1); +by (rtac (less_fun RS iffD2) 1); +by (strip_tac 1); +by (asm_simp_tac (simpset() addsimps [inst_cprod_po]) 1); +qed "monofun_pair1"; -qed_goalw "monofun_pair2" thy [monofun] "monofun(Pair x)" - (fn prems => - [ - (asm_simp_tac (simpset() addsimps [inst_cprod_po]) 1) - ]); +val prems = goalw thy [monofun] "monofun(Pair x)"; +by (asm_simp_tac (simpset() addsimps [inst_cprod_po]) 1); +qed "monofun_pair2"; -qed_goal "monofun_pair" thy "[|x1< (x1::'a::cpo,y1::'b::cpo)<<(x2,y2)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans_less 1), - (rtac (monofun_pair1 RS monofunE RS spec RS spec RS mp RS - (less_fun RS iffD1 RS spec)) 1), - (rtac (monofun_pair2 RS monofunE RS spec RS spec RS mp) 2), - (atac 1), - (atac 1) - ]); +val prems = goal thy "[|x1< (x1::'a::cpo,y1::'b::cpo)<<(x2,y2)"; +by (cut_facts_tac prems 1); +by (rtac trans_less 1); +by (rtac (monofun_pair1 RS monofunE RS spec RS spec RS mp RS (less_fun RS iffD1 RS spec)) 1); +by (rtac (monofun_pair2 RS monofunE RS spec RS spec RS mp) 2); +by (atac 1); +by (atac 1); +qed "monofun_pair"; (* ------------------------------------------------------------------------ *) (* fst and snd are monotone *) (* ------------------------------------------------------------------------ *) -qed_goalw "monofun_fst" thy [monofun] "monofun fst" - (fn prems => - [ - (strip_tac 1), - (res_inst_tac [("p","x")] PairE 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","y")] PairE 1), - (hyp_subst_tac 1), - (Asm_simp_tac 1), - (etac (less_cprod4c RS conjunct1) 1) - ]); +val prems = goalw thy [monofun] "monofun fst"; +by (strip_tac 1); +by (res_inst_tac [("p","x")] PairE 1); +by (hyp_subst_tac 1); +by (res_inst_tac [("p","y")] PairE 1); +by (hyp_subst_tac 1); +by (Asm_simp_tac 1); +by (etac (less_cprod4c RS conjunct1) 1); +qed "monofun_fst"; -qed_goalw "monofun_snd" thy [monofun] "monofun snd" - (fn prems => - [ - (strip_tac 1), - (res_inst_tac [("p","x")] PairE 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","y")] PairE 1), - (hyp_subst_tac 1), - (Asm_simp_tac 1), - (etac (less_cprod4c RS conjunct2) 1) - ]); +val prems = goalw thy [monofun] "monofun snd"; +by (strip_tac 1); +by (res_inst_tac [("p","x")] PairE 1); +by (hyp_subst_tac 1); +by (res_inst_tac [("p","y")] PairE 1); +by (hyp_subst_tac 1); +by (Asm_simp_tac 1); +by (etac (less_cprod4c RS conjunct2) 1); +qed "monofun_snd"; (* ------------------------------------------------------------------------ *) (* the type 'a * 'b is a cpo *) (* ------------------------------------------------------------------------ *) -qed_goal "lub_cprod" thy -"chain S ==> range S<<|(lub(range(%i. fst(S i))),lub(range(%i. snd(S i))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (conjI RS is_lubI) 1), - (rtac (allI RS ub_rangeI) 1), - (res_inst_tac [("t","S i")] (surjective_pairing RS ssubst) 1), - (rtac monofun_pair 1), - (rtac is_ub_thelub 1), - (etac (monofun_fst RS ch2ch_monofun) 1), - (rtac is_ub_thelub 1), - (etac (monofun_snd RS ch2ch_monofun) 1), - (strip_tac 1), - (res_inst_tac [("t","u")] (surjective_pairing RS ssubst) 1), - (rtac monofun_pair 1), - (rtac is_lub_thelub 1), - (etac (monofun_fst RS ch2ch_monofun) 1), - (etac (monofun_fst RS ub2ub_monofun) 1), - (rtac is_lub_thelub 1), - (etac (monofun_snd RS ch2ch_monofun) 1), - (etac (monofun_snd RS ub2ub_monofun) 1) - ]); +val prems = goal thy +"chain S ==> range S<<|(lub(range(%i. fst(S i))),lub(range(%i. snd(S i))))"; +by (cut_facts_tac prems 1); +by (rtac (conjI RS is_lubI) 1); +by (rtac (allI RS ub_rangeI) 1); +by (res_inst_tac [("t","S i")] (surjective_pairing RS ssubst) 1); +by (rtac monofun_pair 1); +by (rtac is_ub_thelub 1); +by (etac (monofun_fst RS ch2ch_monofun) 1); +by (rtac is_ub_thelub 1); +by (etac (monofun_snd RS ch2ch_monofun) 1); +by (strip_tac 1); +by (res_inst_tac [("t","u")] (surjective_pairing RS ssubst) 1); +by (rtac monofun_pair 1); +by (rtac is_lub_thelub 1); +by (etac (monofun_fst RS ch2ch_monofun) 1); +by (etac (monofun_fst RS ub2ub_monofun) 1); +by (rtac is_lub_thelub 1); +by (etac (monofun_snd RS ch2ch_monofun) 1); +by (etac (monofun_snd RS ub2ub_monofun) 1); +qed "lub_cprod"; bind_thm ("thelub_cprod", lub_cprod RS thelubI); (* @@ -147,12 +124,10 @@ *) -qed_goal "cpo_cprod" thy "chain(S::nat=>'a::cpo*'b::cpo)==>? x. range S<<| x" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac exI 1), - (etac lub_cprod 1) - ]); +val prems = goal thy "chain(S::nat=>'a::cpo*'b::cpo)==>? x. range S<<| x"; +by (cut_facts_tac prems 1); +by (rtac exI 1); +by (etac lub_cprod 1); +qed "cpo_cprod"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Cprod3.ML --- a/src/HOLCF/Cprod3.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Cprod3.ML Tue Jul 04 15:58:11 2000 +0200 @@ -1,143 +1,119 @@ -(* Title: HOLCF/cprod3.ML +(* Title: HOLCF/Cprod3 ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for Cprod3.thy +Class instance of * for class pcpo and cpo. *) -open Cprod3; - (* for compatibility with old HOLCF-Version *) -qed_goal "inst_cprod_pcpo" thy "UU = (UU,UU)" - (fn prems => - [ - (simp_tac (HOL_ss addsimps [UU_def,UU_cprod_def]) 1) - ]); +val prems = goal thy "UU = (UU,UU)"; +by (simp_tac (HOL_ss addsimps [UU_def,UU_cprod_def]) 1); +qed "inst_cprod_pcpo"; (* ------------------------------------------------------------------------ *) (* continuity of (_,_) , fst, snd *) (* ------------------------------------------------------------------------ *) -qed_goal "Cprod3_lemma1" Cprod3.thy +val prems = goal Cprod3.thy "chain(Y::(nat=>'a::cpo)) ==>\ \ (lub(range(Y)),(x::'b::cpo)) =\ -\ (lub(range(%i. fst(Y i,x))),lub(range(%i. snd(Y i,x))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("f1","Pair")] (arg_cong RS cong) 1), - (rtac lub_equal 1), - (atac 1), - (rtac (monofun_fst RS ch2ch_monofun) 1), - (rtac ch2ch_fun 1), - (rtac (monofun_pair1 RS ch2ch_monofun) 1), - (atac 1), - (rtac allI 1), - (Simp_tac 1), - (rtac sym 1), - (Simp_tac 1), - (rtac (lub_const RS thelubI) 1) - ]); +\ (lub(range(%i. fst(Y i,x))),lub(range(%i. snd(Y i,x))))"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("f1","Pair")] (arg_cong RS cong) 1); +by (rtac lub_equal 1); +by (atac 1); +by (rtac (monofun_fst RS ch2ch_monofun) 1); +by (rtac ch2ch_fun 1); +by (rtac (monofun_pair1 RS ch2ch_monofun) 1); +by (atac 1); +by (rtac allI 1); +by (Simp_tac 1); +by (rtac sym 1); +by (Simp_tac 1); +by (rtac (lub_const RS thelubI) 1); +qed "Cprod3_lemma1"; -qed_goal "contlub_pair1" Cprod3.thy "contlub(Pair)" - (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (expand_fun_eq RS iffD2) 1), - (strip_tac 1), - (stac (lub_fun RS thelubI) 1), - (etac (monofun_pair1 RS ch2ch_monofun) 1), - (rtac trans 1), - (rtac (thelub_cprod RS sym) 2), - (rtac ch2ch_fun 2), - (etac (monofun_pair1 RS ch2ch_monofun) 2), - (etac Cprod3_lemma1 1) - ]); +val prems = goal Cprod3.thy "contlub(Pair)"; +by (rtac contlubI 1); +by (strip_tac 1); +by (rtac (expand_fun_eq RS iffD2) 1); +by (strip_tac 1); +by (stac (lub_fun RS thelubI) 1); +by (etac (monofun_pair1 RS ch2ch_monofun) 1); +by (rtac trans 1); +by (rtac (thelub_cprod RS sym) 2); +by (rtac ch2ch_fun 2); +by (etac (monofun_pair1 RS ch2ch_monofun) 2); +by (etac Cprod3_lemma1 1); +qed "contlub_pair1"; -qed_goal "Cprod3_lemma2" Cprod3.thy +val prems = goal Cprod3.thy "chain(Y::(nat=>'a::cpo)) ==>\ \ ((x::'b::cpo),lub(range Y)) =\ -\ (lub(range(%i. fst(x,Y i))),lub(range(%i. snd(x, Y i))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("f1","Pair")] (arg_cong RS cong) 1), - (rtac sym 1), - (Simp_tac 1), - (rtac (lub_const RS thelubI) 1), - (rtac lub_equal 1), - (atac 1), - (rtac (monofun_snd RS ch2ch_monofun) 1), - (rtac (monofun_pair2 RS ch2ch_monofun) 1), - (atac 1), - (rtac allI 1), - (Simp_tac 1) - ]); +\ (lub(range(%i. fst(x,Y i))),lub(range(%i. snd(x, Y i))))"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("f1","Pair")] (arg_cong RS cong) 1); +by (rtac sym 1); +by (Simp_tac 1); +by (rtac (lub_const RS thelubI) 1); +by (rtac lub_equal 1); +by (atac 1); +by (rtac (monofun_snd RS ch2ch_monofun) 1); +by (rtac (monofun_pair2 RS ch2ch_monofun) 1); +by (atac 1); +by (rtac allI 1); +by (Simp_tac 1); +qed "Cprod3_lemma2"; -qed_goal "contlub_pair2" Cprod3.thy "contlub(Pair(x))" - (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac trans 1), - (rtac (thelub_cprod RS sym) 2), - (etac (monofun_pair2 RS ch2ch_monofun) 2), - (etac Cprod3_lemma2 1) - ]); +val prems = goal Cprod3.thy "contlub(Pair(x))"; +by (rtac contlubI 1); +by (strip_tac 1); +by (rtac trans 1); +by (rtac (thelub_cprod RS sym) 2); +by (etac (monofun_pair2 RS ch2ch_monofun) 2); +by (etac Cprod3_lemma2 1); +qed "contlub_pair2"; -qed_goal "cont_pair1" Cprod3.thy "cont(Pair)" -(fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_pair1 1), - (rtac contlub_pair1 1) - ]); +val prems = goal Cprod3.thy "cont(Pair)"; +by (rtac monocontlub2cont 1); +by (rtac monofun_pair1 1); +by (rtac contlub_pair1 1); +qed "cont_pair1"; -qed_goal "cont_pair2" Cprod3.thy "cont(Pair(x))" -(fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_pair2 1), - (rtac contlub_pair2 1) - ]); +val prems = goal Cprod3.thy "cont(Pair(x))"; +by (rtac monocontlub2cont 1); +by (rtac monofun_pair2 1); +by (rtac contlub_pair2 1); +qed "cont_pair2"; -qed_goal "contlub_fst" Cprod3.thy "contlub(fst)" - (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (stac (lub_cprod RS thelubI) 1), - (atac 1), - (Simp_tac 1) - ]); +val prems = goal Cprod3.thy "contlub(fst)"; +by (rtac contlubI 1); +by (strip_tac 1); +by (stac (lub_cprod RS thelubI) 1); +by (atac 1); +by (Simp_tac 1); +qed "contlub_fst"; -qed_goal "contlub_snd" Cprod3.thy "contlub(snd)" - (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (stac (lub_cprod RS thelubI) 1), - (atac 1), - (Simp_tac 1) - ]); +val prems = goal Cprod3.thy "contlub(snd)"; +by (rtac contlubI 1); +by (strip_tac 1); +by (stac (lub_cprod RS thelubI) 1); +by (atac 1); +by (Simp_tac 1); +qed "contlub_snd"; -qed_goal "cont_fst" Cprod3.thy "cont(fst)" -(fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_fst 1), - (rtac contlub_fst 1) - ]); +val prems = goal Cprod3.thy "cont(fst)"; +by (rtac monocontlub2cont 1); +by (rtac monofun_fst 1); +by (rtac contlub_fst 1); +qed "cont_fst"; -qed_goal "cont_snd" Cprod3.thy "cont(snd)" -(fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_snd 1), - (rtac contlub_snd 1) - ]); +val prems = goal Cprod3.thy "cont(snd)"; +by (rtac monocontlub2cont 1); +by (rtac monofun_snd 1); +by (rtac contlub_snd 1); +qed "cont_snd"; (* -------------------------------------------------------------------------- @@ -150,130 +126,109 @@ (* convert all lemmas to the continuous versions *) (* ------------------------------------------------------------------------ *) -qed_goalw "beta_cfun_cprod" Cprod3.thy [cpair_def] - "(LAM x y.(x,y))`a`b = (a,b)" - (fn prems => - [ - (stac beta_cfun 1), - (simp_tac (simpset() addsimps [cont_pair1,cont_pair2,cont2cont_CF1L]) 1), - (stac beta_cfun 1), - (rtac cont_pair2 1), - (rtac refl 1) - ]); +val prems = goalw Cprod3.thy [cpair_def] + "(LAM x y.(x,y))`a`b = (a,b)"; +by (stac beta_cfun 1); +by (simp_tac (simpset() addsimps [cont_pair1,cont_pair2,cont2cont_CF1L]) 1); +by (stac beta_cfun 1); +by (rtac cont_pair2 1); +by (rtac refl 1); +qed "beta_cfun_cprod"; -qed_goalw "inject_cpair" Cprod3.thy [cpair_def] - " = ==> a=aa & b=ba" - (fn prems => - [ - (cut_facts_tac prems 1), - (dtac (beta_cfun_cprod RS subst) 1), - (dtac (beta_cfun_cprod RS subst) 1), - (etac Pair_inject 1), - (fast_tac HOL_cs 1) - ]); +val prems = goalw Cprod3.thy [cpair_def] + " = ==> a=aa & b=ba"; +by (cut_facts_tac prems 1); +by (dtac (beta_cfun_cprod RS subst) 1); +by (dtac (beta_cfun_cprod RS subst) 1); +by (etac Pair_inject 1); +by (fast_tac HOL_cs 1); +qed "inject_cpair"; -qed_goalw "inst_cprod_pcpo2" Cprod3.thy [cpair_def] "UU = " - (fn prems => - [ - (rtac sym 1), - (rtac trans 1), - (rtac beta_cfun_cprod 1), - (rtac sym 1), - (rtac inst_cprod_pcpo 1) - ]); +val prems = goalw Cprod3.thy [cpair_def] "UU = "; +by (rtac sym 1); +by (rtac trans 1); +by (rtac beta_cfun_cprod 1); +by (rtac sym 1); +by (rtac inst_cprod_pcpo 1); +qed "inst_cprod_pcpo2"; -qed_goal "defined_cpair_rev" Cprod3.thy - " = UU ==> a = UU & b = UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (dtac (inst_cprod_pcpo2 RS subst) 1), - (etac inject_cpair 1) - ]); +val prems = goal Cprod3.thy + " = UU ==> a = UU & b = UU"; +by (cut_facts_tac prems 1); +by (dtac (inst_cprod_pcpo2 RS subst) 1); +by (etac inject_cpair 1); +qed "defined_cpair_rev"; -qed_goalw "Exh_Cprod2" Cprod3.thy [cpair_def] - "? a b. z=" - (fn prems => - [ - (rtac PairE 1), - (rtac exI 1), - (rtac exI 1), - (etac (beta_cfun_cprod RS ssubst) 1) - ]); +val prems = goalw Cprod3.thy [cpair_def] + "? a b. z="; +by (rtac PairE 1); +by (rtac exI 1); +by (rtac exI 1); +by (etac (beta_cfun_cprod RS ssubst) 1); +qed "Exh_Cprod2"; -qed_goalw "cprodE" Cprod3.thy [cpair_def] -"[|!!x y. [|p= |] ==> Q|] ==> Q" - (fn prems => - [ - (rtac PairE 1), - (resolve_tac prems 1), - (etac (beta_cfun_cprod RS ssubst) 1) - ]); +val prems = goalw Cprod3.thy [cpair_def] +"[|!!x y. [|p= |] ==> Q|] ==> Q"; +by (rtac PairE 1); +by (resolve_tac prems 1); +by (etac (beta_cfun_cprod RS ssubst) 1); +qed "cprodE"; -qed_goalw "cfst2" Cprod3.thy [cfst_def,cpair_def] - "cfst`=x" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun_cprod 1), - (stac beta_cfun 1), - (rtac cont_fst 1), - (Simp_tac 1) - ]); +val prems = goalw Cprod3.thy [cfst_def,cpair_def] + "cfst`=x"; +by (cut_facts_tac prems 1); +by (stac beta_cfun_cprod 1); +by (stac beta_cfun 1); +by (rtac cont_fst 1); +by (Simp_tac 1); +qed "cfst2"; -qed_goalw "csnd2" Cprod3.thy [csnd_def,cpair_def] - "csnd`=y" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun_cprod 1), - (stac beta_cfun 1), - (rtac cont_snd 1), - (Simp_tac 1) - ]); +val prems = goalw Cprod3.thy [csnd_def,cpair_def] + "csnd`=y"; +by (cut_facts_tac prems 1); +by (stac beta_cfun_cprod 1); +by (stac beta_cfun 1); +by (rtac cont_snd 1); +by (Simp_tac 1); +qed "csnd2"; -qed_goal "cfst_strict" Cprod3.thy "cfst`UU = UU" (fn _ => [ - (simp_tac (HOL_ss addsimps [inst_cprod_pcpo2,cfst2]) 1)]); -qed_goal "csnd_strict" Cprod3.thy "csnd`UU = UU" (fn _ => [ - (simp_tac (HOL_ss addsimps [inst_cprod_pcpo2,csnd2]) 1)]); +val _ = goal Cprod3.thy "cfst`UU = UU"; +by (simp_tac (HOL_ss addsimps [inst_cprod_pcpo2,cfst2]) 1); +qed "cfst_strict"; +val _ = goal Cprod3.thy "csnd`UU = UU"; +by (simp_tac (HOL_ss addsimps [inst_cprod_pcpo2,csnd2]) 1); +qed "csnd_strict"; -qed_goalw "surjective_pairing_Cprod2" Cprod3.thy - [cfst_def,csnd_def,cpair_def] " = p" - (fn prems => - [ - (stac beta_cfun_cprod 1), - (stac beta_cfun 1), - (rtac cont_snd 1), - (stac beta_cfun 1), - (rtac cont_fst 1), - (rtac (surjective_pairing RS sym) 1) - ]); +val prems = goalw Cprod3.thy [cfst_def,csnd_def,cpair_def] " = p"; +by (stac beta_cfun_cprod 1); +by (stac beta_cfun 1); +by (rtac cont_snd 1); +by (stac beta_cfun 1); +by (rtac cont_fst 1); +by (rtac (surjective_pairing RS sym) 1); +qed "surjective_pairing_Cprod2"; -qed_goalw "less_cprod5c" Cprod3.thy [cfst_def,csnd_def,cpair_def] - " << ==> xa< - [ - (cut_facts_tac prems 1), - (rtac less_cprod4c 1), - (dtac (beta_cfun_cprod RS subst) 1), - (dtac (beta_cfun_cprod RS subst) 1), - (atac 1) - ]); +val prems = goalw Cprod3.thy [cfst_def,csnd_def,cpair_def] + " << ==> xa< range(S) <<| \ -\ <(lub(range(%i. cfst`(S i)))) , lub(range(%i. csnd`(S i)))>" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun_cprod 1), - (stac (beta_cfun RS ext) 1), - (rtac cont_snd 1), - (stac (beta_cfun RS ext) 1), - (rtac cont_fst 1), - (rtac lub_cprod 1), - (atac 1) - ]); +\ <(lub(range(%i. cfst`(S i)))) , lub(range(%i. csnd`(S i)))>"; +by (cut_facts_tac prems 1); +by (stac beta_cfun_cprod 1); +by (stac (beta_cfun RS ext) 1); +by (rtac cont_snd 1); +by (stac (beta_cfun RS ext) 1); +by (rtac cont_fst 1); +by (rtac lub_cprod 1); +by (atac 1); +qed "lub_cprod2"; bind_thm ("thelub_cprod2", lub_cprod2 RS thelubI); (* @@ -281,23 +236,19 @@ lub (range ?S1) = " *) -qed_goalw "csplit2" Cprod3.thy [csplit_def] - "csplit`f` = f`x`y" - (fn prems => - [ - (stac beta_cfun 1), - (Simp_tac 1), - (simp_tac (simpset() addsimps [cfst2,csnd2]) 1) - ]); +val prems = goalw Cprod3.thy [csplit_def] + "csplit`f` = f`x`y"; +by (stac beta_cfun 1); +by (Simp_tac 1); +by (simp_tac (simpset() addsimps [cfst2,csnd2]) 1); +qed "csplit2"; -qed_goalw "csplit3" Cprod3.thy [csplit_def] - "csplit`cpair`z=z" - (fn prems => - [ - (stac beta_cfun 1), - (Simp_tac 1), - (simp_tac (simpset() addsimps [surjective_pairing_Cprod2]) 1) - ]); +val prems = goalw Cprod3.thy [csplit_def] + "csplit`cpair`z=z"; +by (stac beta_cfun 1); +by (Simp_tac 1); +by (simp_tac (simpset() addsimps [surjective_pairing_Cprod2]) 1); +qed "csplit3"; (* ------------------------------------------------------------------------ *) (* install simplifier for Cprod *) diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Fix.ML --- a/src/HOLCF/Fix.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Fix.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,52 +3,44 @@ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for Fix.thy +fixed point operator and admissibility *) -open Fix; - (* ------------------------------------------------------------------------ *) (* derive inductive properties of iterate from primitive recursion *) (* ------------------------------------------------------------------------ *) -qed_goal "iterate_Suc2" thy "iterate (Suc n) F x = iterate n F (F`x)" - (fn prems => - [ - (induct_tac "n" 1), - (Simp_tac 1), - (stac iterate_Suc 1), - (stac iterate_Suc 1), - (etac ssubst 1), - (rtac refl 1) - ]); +val prems = goal thy "iterate (Suc n) F x = iterate n F (F`x)"; +by (induct_tac "n" 1); +by (Simp_tac 1); +by (stac iterate_Suc 1); +by (stac iterate_Suc 1); +by (etac ssubst 1); +by (rtac refl 1); +qed "iterate_Suc2"; (* ------------------------------------------------------------------------ *) (* the sequence of function itertaions is a chain *) (* This property is essential since monotonicity of iterate makes no sense *) (* ------------------------------------------------------------------------ *) -qed_goalw "chain_iterate2" thy [chain] - " x << F`x ==> chain (%i. iterate i F x)" - (fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (Simp_tac 1), - (induct_tac "i" 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (etac monofun_cfun_arg 1) - ]); +val prems = goalw thy [chain] + " x << F`x ==> chain (%i. iterate i F x)"; +by (cut_facts_tac prems 1); +by (strip_tac 1); +by (Simp_tac 1); +by (induct_tac "i" 1); +by (Asm_simp_tac 1); +by (Asm_simp_tac 1); +by (etac monofun_cfun_arg 1); +qed "chain_iterate2"; -qed_goal "chain_iterate" thy - "chain (%i. iterate i F UU)" - (fn prems => - [ - (rtac chain_iterate2 1), - (rtac minimal 1) - ]); +val prems = goal thy + "chain (%i. iterate i F UU)"; +by (rtac chain_iterate2 1); +by (rtac minimal 1); +qed "chain_iterate"; (* ------------------------------------------------------------------------ *) @@ -57,65 +49,53 @@ (* ------------------------------------------------------------------------ *) -qed_goalw "Ifix_eq" thy [Ifix_def] "Ifix F =F`(Ifix F)" - (fn prems => - [ - (stac contlub_cfun_arg 1), - (rtac chain_iterate 1), - (rtac antisym_less 1), - (rtac lub_mono 1), - (rtac chain_iterate 1), - (rtac ch2ch_Rep_CFunR 1), - (rtac chain_iterate 1), - (rtac allI 1), - (rtac (iterate_Suc RS subst) 1), - (rtac (chain_iterate RS chainE RS spec) 1), - (rtac is_lub_thelub 1), - (rtac ch2ch_Rep_CFunR 1), - (rtac chain_iterate 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (rtac (iterate_Suc RS subst) 1), - (rtac is_ub_thelub 1), - (rtac chain_iterate 1) - ]); +val prems = goalw thy [Ifix_def] "Ifix F =F`(Ifix F)"; +by (stac contlub_cfun_arg 1); +by (rtac chain_iterate 1); +by (rtac antisym_less 1); +by (rtac lub_mono 1); +by (rtac chain_iterate 1); +by (rtac ch2ch_Rep_CFunR 1); +by (rtac chain_iterate 1); +by (rtac allI 1); +by (rtac (iterate_Suc RS subst) 1); +by (rtac (chain_iterate RS chainE RS spec) 1); +by (rtac is_lub_thelub 1); +by (rtac ch2ch_Rep_CFunR 1); +by (rtac chain_iterate 1); +by (rtac ub_rangeI 1); +by (rtac allI 1); +by (rtac (iterate_Suc RS subst) 1); +by (rtac is_ub_thelub 1); +by (rtac chain_iterate 1); +qed "Ifix_eq"; -qed_goalw "Ifix_least" thy [Ifix_def] "F`x=x ==> Ifix(F) << x" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_lub_thelub 1), - (rtac chain_iterate 1), - (rtac ub_rangeI 1), - (strip_tac 1), - (induct_tac "i" 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (res_inst_tac [("t","x")] subst 1), - (atac 1), - (etac monofun_cfun_arg 1) - ]); +val prems = goalw thy [Ifix_def] "F`x=x ==> Ifix(F) << x"; +by (cut_facts_tac prems 1); +by (rtac is_lub_thelub 1); +by (rtac chain_iterate 1); +by (rtac ub_rangeI 1); +by (strip_tac 1); +by (induct_tac "i" 1); +by (Asm_simp_tac 1); +by (Asm_simp_tac 1); +by (res_inst_tac [("t","x")] subst 1); +by (atac 1); +by (etac monofun_cfun_arg 1); +qed "Ifix_least"; (* ------------------------------------------------------------------------ *) (* monotonicity and continuity of iterate *) (* ------------------------------------------------------------------------ *) -qed_goalw "monofun_iterate" thy [monofun] "monofun(iterate(i))" - (fn prems => - [ - (strip_tac 1), - (induct_tac "i" 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (rtac (less_fun RS iffD2) 1), - (rtac allI 1), - (rtac monofun_cfun 1), - (atac 1), - (rtac (less_fun RS iffD1 RS spec) 1), - (atac 1) - ]); +Goalw [monofun] "monofun(iterate(i))"; +by (strip_tac 1); +by (induct_tac "i" 1); +by (Asm_simp_tac 1); +by (asm_full_simp_tac (simpset() addsimps [less_fun, monofun_cfun]) 1); +qed "monofun_iterate"; (* ------------------------------------------------------------------------ *) (* the following lemma uses contlub_cfun which itself is based on a *) @@ -123,118 +103,103 @@ (* In this special case it is the application function Rep_CFun *) (* ------------------------------------------------------------------------ *) -qed_goalw "contlub_iterate" thy [contlub] "contlub(iterate(i))" - (fn prems => - [ - (strip_tac 1), - (induct_tac "i" 1), - (Asm_simp_tac 1), - (rtac (lub_const RS thelubI RS sym) 1), - (asm_simp_tac (simpset() delsimps [range_composition]) 1), - (rtac ext 1), - (stac thelub_fun 1), - (rtac chainI 1), - (rtac allI 1), - (rtac (less_fun RS iffD2) 1), - (rtac allI 1), - (rtac (chainE RS spec) 1), - (rtac (monofun_Rep_CFun1 RS ch2ch_MF2LR) 1), - (rtac allI 1), - (rtac monofun_Rep_CFun2 1), - (atac 1), - (rtac ch2ch_fun 1), - (rtac (monofun_iterate RS ch2ch_monofun) 1), - (atac 1), - (stac thelub_fun 1), - (rtac (monofun_iterate RS ch2ch_monofun) 1), - (atac 1), - (rtac contlub_cfun 1), - (atac 1), - (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1) - ]); +val prems = goalw thy [contlub] "contlub(iterate(i))"; +by (strip_tac 1); +by (induct_tac "i" 1); +by (Asm_simp_tac 1); +by (rtac (lub_const RS thelubI RS sym) 1); +by (asm_simp_tac (simpset() delsimps [range_composition]) 1); +by (rtac ext 1); +by (stac thelub_fun 1); +by (rtac chainI 1); +by (rtac allI 1); +by (rtac (less_fun RS iffD2) 1); +by (rtac allI 1); +by (rtac (chainE RS spec) 1); +by (rtac (monofun_Rep_CFun1 RS ch2ch_MF2LR) 1); +by (rtac allI 1); +by (rtac monofun_Rep_CFun2 1); +by (atac 1); +by (rtac ch2ch_fun 1); +by (rtac (monofun_iterate RS ch2ch_monofun) 1); +by (atac 1); +by (stac thelub_fun 1); +by (rtac (monofun_iterate RS ch2ch_monofun) 1); +by (atac 1); +by (rtac contlub_cfun 1); +by (atac 1); +by (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1); +qed "contlub_iterate"; -qed_goal "cont_iterate" thy "cont(iterate(i))" - (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_iterate 1), - (rtac contlub_iterate 1) - ]); +val prems = goal thy "cont(iterate(i))"; +by (rtac monocontlub2cont 1); +by (rtac monofun_iterate 1); +by (rtac contlub_iterate 1); +qed "cont_iterate"; (* ------------------------------------------------------------------------ *) (* a lemma about continuity of iterate in its third argument *) (* ------------------------------------------------------------------------ *) -qed_goal "monofun_iterate2" thy "monofun(iterate n F)" - (fn prems => - [ - (rtac monofunI 1), - (strip_tac 1), - (induct_tac "n" 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (etac monofun_cfun_arg 1) - ]); +val prems = goal thy "monofun(iterate n F)"; +by (rtac monofunI 1); +by (strip_tac 1); +by (induct_tac "n" 1); +by (Asm_simp_tac 1); +by (Asm_simp_tac 1); +by (etac monofun_cfun_arg 1); +qed "monofun_iterate2"; -qed_goal "contlub_iterate2" thy "contlub(iterate n F)" - (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (induct_tac "n" 1), - (Simp_tac 1), - (Simp_tac 1), - (res_inst_tac [("t","iterate n F (lub(range(%u. Y u)))"), - ("s","lub(range(%i. iterate n F (Y i)))")] ssubst 1), - (atac 1), - (rtac contlub_cfun_arg 1), - (etac (monofun_iterate2 RS ch2ch_monofun) 1) - ]); +Goal "contlub(iterate n F)"; +by (rtac contlubI 1); +by (strip_tac 1); +by (induct_tac "n" 1); +by (Simp_tac 1); +by (Simp_tac 1); +by (res_inst_tac [("t","iterate n F (lub(range(%u. Y u)))"), + ("s","lub(range(%i. iterate n F (Y i)))")] ssubst 1); +by (atac 1); +by (rtac contlub_cfun_arg 1); +by (etac (monofun_iterate2 RS ch2ch_monofun) 1); +qed "contlub_iterate2"; -qed_goal "cont_iterate2" thy "cont (iterate n F)" - (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_iterate2 1), - (rtac contlub_iterate2 1) - ]); +val prems = goal thy "cont (iterate n F)"; +by (rtac monocontlub2cont 1); +by (rtac monofun_iterate2 1); +by (rtac contlub_iterate2 1); +qed "cont_iterate2"; (* ------------------------------------------------------------------------ *) (* monotonicity and continuity of Ifix *) (* ------------------------------------------------------------------------ *) -qed_goalw "monofun_Ifix" thy [monofun,Ifix_def] "monofun(Ifix)" - (fn prems => - [ - (strip_tac 1), - (rtac lub_mono 1), - (rtac chain_iterate 1), - (rtac chain_iterate 1), - (rtac allI 1), - (rtac (less_fun RS iffD1 RS spec) 1), - (etac (monofun_iterate RS monofunE RS spec RS spec RS mp) 1) - ]); +Goalw [monofun,Ifix_def] "monofun(Ifix)"; +by (strip_tac 1); +by (rtac lub_mono 1); +by (rtac chain_iterate 1); +by (rtac chain_iterate 1); +by (rtac allI 1); +by (rtac (less_fun RS iffD1 RS spec) 1 THEN + etac (monofun_iterate RS monofunE RS spec RS spec RS mp) 1); +qed "monofun_Ifix"; (* ------------------------------------------------------------------------ *) (* since iterate is not monotone in its first argument, special lemmas must *) (* be derived for lubs in this argument *) (* ------------------------------------------------------------------------ *) -qed_goal "chain_iterate_lub" thy -"chain(Y) ==> chain(%i. lub(range(%ia. iterate ia (Y i) UU)))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac chainI 1), - (strip_tac 1), - (rtac lub_mono 1), - (rtac chain_iterate 1), - (rtac chain_iterate 1), - (strip_tac 1), - (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun RS chainE - RS spec) 1) - ]); +val prems = goal thy +"chain(Y) ==> chain(%i. lub(range(%ia. iterate ia (Y i) UU)))"; +by (cut_facts_tac prems 1); +by (rtac chainI 1); +by (strip_tac 1); +by (rtac lub_mono 1); +by (rtac chain_iterate 1); +by (rtac chain_iterate 1); +by (strip_tac 1); +by (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun RS chainE RS spec) 1); +qed "chain_iterate_lub"; (* ------------------------------------------------------------------------ *) (* this exchange lemma is analog to the one for monotone functions *) @@ -242,138 +207,111 @@ (* chains is the essential argument which is usually derived from monot. *) (* ------------------------------------------------------------------------ *) -qed_goal "contlub_Ifix_lemma1" thy -"chain(Y) ==>iterate n (lub(range Y)) y = lub(range(%i. iterate n (Y i) y))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (thelub_fun RS subst) 1), - (rtac (monofun_iterate RS ch2ch_monofun) 1), - (atac 1), - (rtac fun_cong 1), - (stac (contlub_iterate RS contlubE RS spec RS mp) 1), - (atac 1), - (rtac refl 1) - ]); +Goal + "chain(Y) ==>iterate n (lub(range Y)) y = lub(range(%i. iterate n (Y i) y))"; +by (rtac (thelub_fun RS subst) 1); +by (etac (monofun_iterate RS ch2ch_monofun) 1); +by (asm_simp_tac (simpset() addsimps [contlub_iterate RS contlubE]) 1); +qed "contlub_Ifix_lemma1"; -qed_goal "ex_lub_iterate" thy "chain(Y) ==>\ +val prems = goal thy "chain(Y) ==>\ \ lub(range(%i. lub(range(%ia. iterate i (Y ia) UU)))) =\ -\ lub(range(%i. lub(range(%ia. iterate ia (Y i) UU))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac antisym_less 1), - (rtac is_lub_thelub 1), - (rtac (contlub_Ifix_lemma1 RS ext RS subst) 1), - (atac 1), - (rtac chain_iterate 1), - (rtac ub_rangeI 1), - (strip_tac 1), - (rtac lub_mono 1), - (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1), - (etac chain_iterate_lub 1), - (strip_tac 1), - (rtac is_ub_thelub 1), - (rtac chain_iterate 1), - (rtac is_lub_thelub 1), - (etac chain_iterate_lub 1), - (rtac ub_rangeI 1), - (strip_tac 1), - (rtac lub_mono 1), - (rtac chain_iterate 1), - (rtac (contlub_Ifix_lemma1 RS ext RS subst) 1), - (atac 1), - (rtac chain_iterate 1), - (strip_tac 1), - (rtac is_ub_thelub 1), - (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1) - ]); +\ lub(range(%i. lub(range(%ia. iterate ia (Y i) UU))))"; +by (cut_facts_tac prems 1); +by (rtac antisym_less 1); +by (rtac is_lub_thelub 1); +by (rtac (contlub_Ifix_lemma1 RS ext RS subst) 1); +by (atac 1); +by (rtac chain_iterate 1); +by (rtac ub_rangeI 1); +by (strip_tac 1); +by (rtac lub_mono 1); +by (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1); +by (etac chain_iterate_lub 1); +by (strip_tac 1); +by (rtac is_ub_thelub 1); +by (rtac chain_iterate 1); +by (rtac is_lub_thelub 1); +by (etac chain_iterate_lub 1); +by (rtac ub_rangeI 1); +by (strip_tac 1); +by (rtac lub_mono 1); +by (rtac chain_iterate 1); +by (rtac (contlub_Ifix_lemma1 RS ext RS subst) 1); +by (atac 1); +by (rtac chain_iterate 1); +by (strip_tac 1); +by (rtac is_ub_thelub 1); +by (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1); +qed "ex_lub_iterate"; -qed_goalw "contlub_Ifix" thy [contlub,Ifix_def] "contlub(Ifix)" - (fn prems => - [ - (strip_tac 1), - (stac (contlub_Ifix_lemma1 RS ext) 1), - (atac 1), - (etac ex_lub_iterate 1) - ]); +val prems = goalw thy [contlub,Ifix_def] "contlub(Ifix)"; +by (strip_tac 1); +by (stac (contlub_Ifix_lemma1 RS ext) 1); +by (atac 1); +by (etac ex_lub_iterate 1); +qed "contlub_Ifix"; -qed_goal "cont_Ifix" thy "cont(Ifix)" - (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Ifix 1), - (rtac contlub_Ifix 1) - ]); +val prems = goal thy "cont(Ifix)"; +by (rtac monocontlub2cont 1); +by (rtac monofun_Ifix 1); +by (rtac contlub_Ifix 1); +qed "cont_Ifix"; (* ------------------------------------------------------------------------ *) (* propagate properties of Ifix to its continuous counterpart *) (* ------------------------------------------------------------------------ *) -qed_goalw "fix_eq" thy [fix_def] "fix`F = F`(fix`F)" - (fn prems => - [ - (asm_simp_tac (simpset() addsimps [cont_Ifix]) 1), - (rtac Ifix_eq 1) - ]); +val prems = goalw thy [fix_def] "fix`F = F`(fix`F)"; +by (asm_simp_tac (simpset() addsimps [cont_Ifix]) 1); +by (rtac Ifix_eq 1); +qed "fix_eq"; -qed_goalw "fix_least" thy [fix_def] "F`x = x ==> fix`F << x" - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (simpset() addsimps [cont_Ifix]) 1), - (etac Ifix_least 1) - ]); +val prems = goalw thy [fix_def] "F`x = x ==> fix`F << x"; +by (cut_facts_tac prems 1); +by (asm_simp_tac (simpset() addsimps [cont_Ifix]) 1); +by (etac Ifix_least 1); +qed "fix_least"; -qed_goal "fix_eqI" thy -"[| F`x = x; !z. F`z = z --> x << z |] ==> x = fix`F" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac antisym_less 1), - (etac allE 1), - (etac mp 1), - (rtac (fix_eq RS sym) 1), - (etac fix_least 1) - ]); +val prems = goal thy +"[| F`x = x; !z. F`z = z --> x << z |] ==> x = fix`F"; +by (cut_facts_tac prems 1); +by (rtac antisym_less 1); +by (etac allE 1); +by (etac mp 1); +by (rtac (fix_eq RS sym) 1); +by (etac fix_least 1); +qed "fix_eqI"; -qed_goal "fix_eq2" thy "f == fix`F ==> f = F`f" - (fn prems => - [ - (rewrite_goals_tac prems), - (rtac fix_eq 1) - ]); +val prems = goal thy "f == fix`F ==> f = F`f"; +by (rewrite_goals_tac prems); +by (rtac fix_eq 1); +qed "fix_eq2"; -qed_goal "fix_eq3" thy "f == fix`F ==> f`x = F`f`x" - (fn prems => - [ - (rtac trans 1), - (rtac ((hd prems) RS fix_eq2 RS cfun_fun_cong) 1), - (rtac refl 1) - ]); +val prems = goal thy "f == fix`F ==> f`x = F`f`x"; +by (rtac trans 1); +by (rtac ((hd prems) RS fix_eq2 RS cfun_fun_cong) 1); +by (rtac refl 1); +qed "fix_eq3"; fun fix_tac3 thm i = ((rtac trans i) THEN (rtac (thm RS fix_eq3) i)); -qed_goal "fix_eq4" thy "f = fix`F ==> f = F`f" - (fn prems => - [ - (cut_facts_tac prems 1), - (hyp_subst_tac 1), - (rtac fix_eq 1) - ]); +val prems = goal thy "f = fix`F ==> f = F`f"; +by (cut_facts_tac prems 1); +by (hyp_subst_tac 1); +by (rtac fix_eq 1); +qed "fix_eq4"; -qed_goal "fix_eq5" thy "f = fix`F ==> f`x = F`f`x" - (fn prems => - [ - (rtac trans 1), - (rtac ((hd prems) RS fix_eq4 RS cfun_fun_cong) 1), - (rtac refl 1) - ]); +val prems = goal thy "f = fix`F ==> f`x = F`f`x"; +by (rtac trans 1); +by (rtac ((hd prems) RS fix_eq4 RS cfun_fun_cong) 1); +by (rtac refl 1); +qed "fix_eq5"; fun fix_tac5 thm i = ((rtac trans i) THEN (rtac (thm RS fix_eq5) i)); @@ -408,25 +346,21 @@ (* ------------------------------------------------------------------------ *) -qed_goal "Ifix_def2" thy "Ifix=(%x. lub(range(%i. iterate i x UU)))" - (fn prems => - [ - (rtac ext 1), - (rewtac Ifix_def), - (rtac refl 1) - ]); +val prems = goal thy "Ifix=(%x. lub(range(%i. iterate i x UU)))"; +by (rtac ext 1); +by (rewtac Ifix_def); +by (rtac refl 1); +qed "Ifix_def2"; (* ------------------------------------------------------------------------ *) (* direct connection between fix and iteration without Ifix *) (* ------------------------------------------------------------------------ *) -qed_goalw "fix_def2" thy [fix_def] - "fix`F = lub(range(%i. iterate i F UU))" - (fn prems => - [ - (fold_goals_tac [Ifix_def]), - (asm_simp_tac (simpset() addsimps [cont_Ifix]) 1) - ]); +val prems = goalw thy [fix_def] + "fix`F = lub(range(%i. iterate i F UU))"; +by (fold_goals_tac [Ifix_def]); +by (asm_simp_tac (simpset() addsimps [cont_Ifix]) 1); +qed "fix_def2"; (* ------------------------------------------------------------------------ *) @@ -437,106 +371,99 @@ (* access to definitions *) (* ------------------------------------------------------------------------ *) -qed_goalw "admI" thy [adm_def] - "(!!Y. [| chain(Y); !i. P(Y(i)) |] ==> P(lub(range(Y)))) ==> adm(P)" - (fn prems => [fast_tac (HOL_cs addIs prems) 1]); - -qed_goalw "admD" thy [adm_def] - "!!P. [| adm(P); chain(Y); !i. P(Y(i)) |] ==> P(lub(range(Y)))" - (fn prems => [fast_tac HOL_cs 1]); +val prems = goalw thy [adm_def] + "(!!Y. [| chain(Y); !i. P(Y(i)) |] ==> P(lub(range(Y)))) ==> adm(P)"; +by (fast_tac (HOL_cs addIs prems) 1); +qed "admI"; -qed_goalw "admw_def2" thy [admw_def] +Goalw [adm_def] "[| adm(P); chain(Y); !i. P(Y(i)) |] ==> P(lub(range(Y)))"; +by (Blast_tac 1); +qed "admD"; + +val prems = goalw thy [admw_def] "admw(P) = (!F.(!n. P(iterate n F UU)) -->\ -\ P (lub(range(%i. iterate i F UU))))" - (fn prems => - [ - (rtac refl 1) - ]); +\ P (lub(range(%i. iterate i F UU))))"; +by (rtac refl 1); +qed "admw_def2"; (* ------------------------------------------------------------------------ *) (* an admissible formula is also weak admissible *) (* ------------------------------------------------------------------------ *) -qed_goalw "adm_impl_admw" thy [admw_def] "!!P. adm(P)==>admw(P)" - (fn prems => - [ - (strip_tac 1), - (etac admD 1), - (rtac chain_iterate 1), - (atac 1) - ]); +val prems = goalw thy [admw_def] "!!P. adm(P)==>admw(P)"; +by (strip_tac 1); +by (etac admD 1); +by (rtac chain_iterate 1); +by (atac 1); +qed "adm_impl_admw"; (* ------------------------------------------------------------------------ *) (* fixed point induction *) (* ------------------------------------------------------------------------ *) -qed_goal "fix_ind" thy -"[| adm(P);P(UU);!!x. P(x) ==> P(F`x)|] ==> P(fix`F)" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac fix_def2 1), - (etac admD 1), - (rtac chain_iterate 1), - (rtac allI 1), - (induct_tac "i" 1), - (stac iterate_0 1), - (atac 1), - (stac iterate_Suc 1), - (resolve_tac prems 1), - (atac 1) - ]); +val prems = goal thy +"[| adm(P);P(UU);!!x. P(x) ==> P(F`x)|] ==> P(fix`F)"; +by (cut_facts_tac prems 1); +by (stac fix_def2 1); +by (etac admD 1); +by (rtac chain_iterate 1); +by (rtac allI 1); +by (induct_tac "i" 1); +by (stac iterate_0 1); +by (atac 1); +by (stac iterate_Suc 1); +by (resolve_tac prems 1); +by (atac 1); +qed "fix_ind"; -qed_goal "def_fix_ind" thy "[| f == fix`F; adm(P); \ -\ P(UU);!!x. P(x) ==> P(F`x)|] ==> P f" (fn prems => [ - (cut_facts_tac prems 1), - (asm_simp_tac HOL_ss 1), - (etac fix_ind 1), - (atac 1), - (eresolve_tac prems 1)]); +val prems = goal thy "[| f == fix`F; adm(P); \ +\ P(UU);!!x. P(x) ==> P(F`x)|] ==> P f"; +by (cut_facts_tac prems 1); +by (asm_simp_tac HOL_ss 1); +by (etac fix_ind 1); +by (atac 1); +by (eresolve_tac prems 1); +qed "def_fix_ind"; (* ------------------------------------------------------------------------ *) (* computational induction for weak admissible formulae *) (* ------------------------------------------------------------------------ *) -qed_goal "wfix_ind" thy -"[| admw(P); !n. P(iterate n F UU)|] ==> P(fix`F)" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac fix_def2 1), - (rtac (admw_def2 RS iffD1 RS spec RS mp) 1), - (atac 1), - (rtac allI 1), - (etac spec 1) - ]); +val prems = goal thy +"[| admw(P); !n. P(iterate n F UU)|] ==> P(fix`F)"; +by (cut_facts_tac prems 1); +by (stac fix_def2 1); +by (rtac (admw_def2 RS iffD1 RS spec RS mp) 1); +by (atac 1); +by (rtac allI 1); +by (etac spec 1); +qed "wfix_ind"; -qed_goal "def_wfix_ind" thy "[| f == fix`F; admw(P); \ -\ !n. P(iterate n F UU) |] ==> P f" (fn prems => [ - (cut_facts_tac prems 1), - (asm_simp_tac HOL_ss 1), - (etac wfix_ind 1), - (atac 1)]); +val prems = goal thy "[| f == fix`F; admw(P); \ +\ !n. P(iterate n F UU) |] ==> P f"; +by (cut_facts_tac prems 1); +by (asm_simp_tac HOL_ss 1); +by (etac wfix_ind 1); +by (atac 1); +qed "def_wfix_ind"; (* ------------------------------------------------------------------------ *) (* for chain-finite (easy) types every formula is admissible *) (* ------------------------------------------------------------------------ *) -qed_goalw "adm_max_in_chain" thy [adm_def] -"!Y. chain(Y::nat=>'a) --> (? n. max_in_chain n Y) ==> adm(P::'a=>bool)" - (fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (rtac exE 1), - (rtac mp 1), - (etac spec 1), - (atac 1), - (stac (lub_finch1 RS thelubI) 1), - (atac 1), - (atac 1), - (etac spec 1) - ]); +val prems = goalw thy [adm_def] +"!Y. chain(Y::nat=>'a) --> (? n. max_in_chain n Y) ==> adm(P::'a=>bool)"; +by (cut_facts_tac prems 1); +by (strip_tac 1); +by (rtac exE 1); +by (rtac mp 1); +by (etac spec 1); +by (atac 1); +by (stac (lub_finch1 RS thelubI) 1); +by (atac 1); +by (atac 1); +by (etac spec 1); +qed "adm_max_in_chain"; bind_thm ("adm_chfin" ,chfin RS adm_max_in_chain); @@ -544,12 +471,12 @@ (* some lemmata for functions with flat/chfin domain/range types *) (* ------------------------------------------------------------------------ *) -qed_goalw "adm_chfindom" thy [adm_def] "adm (%(u::'a::cpo->'b::chfin). P(u`s))" - (fn _ => [ - strip_tac 1, - dtac chfin_Rep_CFunR 1, - eres_inst_tac [("x","s")] allE 1, - fast_tac (HOL_cs addss (simpset() addsimps [chfin])) 1]); +val _ = goalw thy [adm_def] "adm (%(u::'a::cpo->'b::chfin). P(u`s))"; +by (strip_tac 1); +by (dtac chfin_Rep_CFunR 1); +by (eres_inst_tac [("x","s")] allE 1); +by (fast_tac (HOL_cs addss (simpset() addsimps [chfin])) 1); +qed "adm_chfindom"; (* adm_flat not needed any more, since it is a special case of adm_chfindom *) @@ -557,106 +484,103 @@ (* improved admisibility introduction *) (* ------------------------------------------------------------------------ *) -qed_goalw "admI2" thy [adm_def] +val prems = goalw thy [adm_def] "(!!Y. [| chain Y; !i. P (Y i); !i. ? j. i < j & Y i ~= Y j & Y i << Y j |]\ -\ ==> P(lub (range Y))) ==> adm P" - (fn prems => [ - strip_tac 1, - etac increasing_chain_adm_lemma 1, atac 1, - eresolve_tac prems 1, atac 1, atac 1]); +\ ==> P(lub (range Y))) ==> adm P"; +by (strip_tac 1); +by (etac increasing_chain_adm_lemma 1); +by (atac 1); +by (eresolve_tac prems 1); +by (atac 1); +by (atac 1); +qed "admI2"; (* ------------------------------------------------------------------------ *) (* admissibility of special formulae and propagation *) (* ------------------------------------------------------------------------ *) -qed_goalw "adm_less" thy [adm_def] - "[|cont u;cont v|]==> adm(%x. u x << v x)" - (fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (etac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1), - (atac 1), - (etac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1), - (atac 1), - (rtac lub_mono 1), - (cut_facts_tac prems 1), - (etac (cont2mono RS ch2ch_monofun) 1), - (atac 1), - (cut_facts_tac prems 1), - (etac (cont2mono RS ch2ch_monofun) 1), - (atac 1), - (atac 1) - ]); +val prems = goalw thy [adm_def] + "[|cont u;cont v|]==> adm(%x. u x << v x)"; +by (cut_facts_tac prems 1); +by (strip_tac 1); +by (etac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1); +by (atac 1); +by (etac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1); +by (atac 1); +by (rtac lub_mono 1); +by (cut_facts_tac prems 1); +by (etac (cont2mono RS ch2ch_monofun) 1); +by (atac 1); +by (cut_facts_tac prems 1); +by (etac (cont2mono RS ch2ch_monofun) 1); +by (atac 1); +by (atac 1); +qed "adm_less"; Addsimps [adm_less]; -qed_goal "adm_conj" thy - "!!P. [| adm P; adm Q |] ==> adm(%x. P x & Q x)" - (fn prems => [fast_tac (HOL_cs addEs [admD] addIs [admI]) 1]); +val prems = goal thy + "!!P. [| adm P; adm Q |] ==> adm(%x. P x & Q x)"; +by (fast_tac (HOL_cs addEs [admD] addIs [admI]) 1); +qed "adm_conj"; Addsimps [adm_conj]; -qed_goalw "adm_not_free" thy [adm_def] "adm(%x. t)" - (fn prems => [fast_tac HOL_cs 1]); +val prems = goalw thy [adm_def] "adm(%x. t)"; +by (fast_tac HOL_cs 1); +qed "adm_not_free"; Addsimps [adm_not_free]; -qed_goalw "adm_not_less" thy [adm_def] - "!!t. cont t ==> adm(%x.~ (t x) << u)" - (fn prems => - [ - (strip_tac 1), - (rtac contrapos 1), - (etac spec 1), - (rtac trans_less 1), - (atac 2), - (etac (cont2mono RS monofun_fun_arg) 1), - (rtac is_ub_thelub 1), - (atac 1) - ]); +val prems = goalw thy [adm_def] + "!!t. cont t ==> adm(%x.~ (t x) << u)"; +by (strip_tac 1); +by (rtac contrapos 1); +by (etac spec 1); +by (rtac trans_less 1); +by (atac 2); +by (etac (cont2mono RS monofun_fun_arg) 1); +by (rtac is_ub_thelub 1); +by (atac 1); +qed "adm_not_less"; -qed_goal "adm_all" thy - "!!P. !y. adm(P y) ==> adm(%x.!y. P y x)" - (fn prems => [fast_tac (HOL_cs addIs [admI] addEs [admD]) 1]); +val prems = goal thy + "!!P. !y. adm(P y) ==> adm(%x.!y. P y x)"; +by (fast_tac (HOL_cs addIs [admI] addEs [admD]) 1); +qed "adm_all"; bind_thm ("adm_all2", allI RS adm_all); -qed_goal "adm_subst" thy - "!!P. [|cont t; adm P|] ==> adm(%x. P (t x))" - (fn prems => - [ - (rtac admI 1), - (stac (cont2contlub RS contlubE RS spec RS mp) 1), - (atac 1), - (atac 1), - (etac admD 1), - (etac (cont2mono RS ch2ch_monofun) 1), - (atac 1), - (atac 1) - ]); +val prems = goal thy + "!!P. [|cont t; adm P|] ==> adm(%x. P (t x))"; +by (rtac admI 1); +by (stac (cont2contlub RS contlubE RS spec RS mp) 1); +by (atac 1); +by (atac 1); +by (etac admD 1); +by (etac (cont2mono RS ch2ch_monofun) 1); +by (atac 1); +by (atac 1); +qed "adm_subst"; -qed_goal "adm_UU_not_less" thy "adm(%x.~ UU << t(x))" - (fn prems => [Simp_tac 1]); +val prems = goal thy "adm(%x.~ UU << t(x))"; +by (Simp_tac 1); +qed "adm_UU_not_less"; -qed_goalw "adm_not_UU" thy [adm_def] - "!!t. cont(t)==> adm(%x.~ (t x) = UU)" - (fn prems => - [ - (strip_tac 1), - (rtac contrapos 1), - (etac spec 1), - (rtac (chain_UU_I RS spec) 1), - (rtac (cont2mono RS ch2ch_monofun) 1), - (atac 1), - (atac 1), - (rtac (cont2contlub RS contlubE RS spec RS mp RS subst) 1), - (atac 1), - (atac 1), - (atac 1) - ]); -qed_goal "adm_eq" thy - "!!u. [|cont u ; cont v|]==> adm(%x. u x = v x)" - (fn prems => [asm_simp_tac (simpset() addsimps [po_eq_conv]) 1]); +Goalw [adm_def] "cont(t)==> adm(%x.~ (t x) = UU)"; +by (strip_tac 1); +by (rtac contrapos 1); +by (etac spec 1); +by (rtac (chain_UU_I RS spec) 1); +by (etac (cont2mono RS ch2ch_monofun) 1); +by (atac 1); +by (etac (cont2contlub RS contlubE RS spec RS mp RS subst) 1); +by (atac 1); +by (atac 1); +qed "adm_not_UU"; + +Goal "[|cont u ; cont v|]==> adm(%x. u x = v x)"; +by (asm_simp_tac (simpset() addsimps [po_eq_conv]) 1); +qed "adm_eq"; @@ -664,202 +588,178 @@ (* admissibility for disjunction is hard to prove. It takes 10 Lemmas *) (* ------------------------------------------------------------------------ *) -local - val adm_disj_lemma1 = prove_goal HOL.thy - "!n. P(Y n)|Q(Y n) ==> (? i.!j. R i j --> Q(Y(j))) | (!i.? j. R i j & P(Y(j)))" - (fn prems => - [ - (cut_facts_tac prems 1), - (fast_tac HOL_cs 1) - ]); +val prems = goal HOL.thy + "!n. P(Y n)|Q(Y n) ==> (? i.!j. R i j --> Q(Y(j))) | (!i.? j. R i j & P(Y(j)))"; +by (cut_facts_tac prems 1); +by (fast_tac HOL_cs 1); +qed "adm_disj_lemma1"; - val adm_disj_lemma2 = prove_goal thy +val _ = goal thy "!!Q. [| adm(Q); ? X. chain(X) & (!n. Q(X(n))) &\ - \ lub(range(Y))=lub(range(X))|] ==> Q(lub(range(Y)))" - (fn _ => [fast_tac (claset() addEs [admD] addss simpset()) 1]); + \ lub(range(Y))=lub(range(X))|] ==> Q(lub(range(Y)))"; +by (fast_tac (claset() addEs [admD] addss simpset()) 1); +qed "adm_disj_lemma2"; - val adm_disj_lemma3 = prove_goalw thy [chain] - "!!Q. chain(Y) ==> chain(%m. if m < Suc i then Y(Suc i) else Y m)" - (fn _ => - [ - Asm_simp_tac 1, - safe_tac HOL_cs, - subgoal_tac "ia = i" 1, - ALLGOALS Asm_simp_tac - ]); +val _ = goalw thy [chain] + "!!Q. chain(Y) ==> chain(%m. if m < Suc i then Y(Suc i) else Y m)"; +by (Asm_simp_tac 1); +by (safe_tac HOL_cs); +by (subgoal_tac "ia = i" 1); +by (ALLGOALS Asm_simp_tac); +qed "adm_disj_lemma3"; - val adm_disj_lemma4 = prove_goal Arith.thy - "!!Q. !j. i < j --> Q(Y(j)) ==> !n. Q( if n < Suc i then Y(Suc i) else Y n)" - (fn _ => - [asm_simp_tac (simpset_of Arith.thy) 1]); +val _ = goal Arith.thy + "!!Q. !j. i < j --> Q(Y(j)) ==> !n. Q( if n < Suc i then Y(Suc i) else Y n)"; +by (asm_simp_tac (simpset_of Arith.thy) 1); +qed "adm_disj_lemma4"; - val adm_disj_lemma5 = prove_goal thy +val prems = goal thy "!!Y::nat=>'a::cpo. [| chain(Y); ! j. i < j --> Q(Y(j)) |] ==>\ - \ lub(range(Y)) = lub(range(%m. if m< Suc(i) then Y(Suc(i)) else Y m))" - (fn prems => - [ - safe_tac (HOL_cs addSIs [lub_equal2,adm_disj_lemma3]), - atac 2, - res_inst_tac [("x","i")] exI 1, - Asm_simp_tac 1 - ]); + \ lub(range(Y)) = lub(range(%m. if m< Suc(i) then Y(Suc(i)) else Y m))"; +by (safe_tac (HOL_cs addSIs [lub_equal2,adm_disj_lemma3])); +by (atac 2); +by (res_inst_tac [("x","i")] exI 1); +by (Asm_simp_tac 1); +qed "adm_disj_lemma5"; - val adm_disj_lemma6 = prove_goal thy +val prems = goal thy "[| chain(Y::nat=>'a::cpo); ? i. ! j. i < j --> Q(Y(j)) |] ==>\ - \ ? X. chain(X) & (! n. Q(X(n))) & lub(range(Y)) = lub(range(X))" - (fn prems => - [ - (cut_facts_tac prems 1), - (etac exE 1), - (res_inst_tac [("x","%m. if m'a::cpo); ! i. ? j. i < j & P(Y(j)) |] ==>\ - \ chain(%m. Y(Least(%j. m - [ - (cut_facts_tac prems 1), - (rtac chainI 1), - (rtac allI 1), - (rtac chain_mono3 1), - (atac 1), - (rtac Least_le 1), - (rtac conjI 1), - (rtac Suc_lessD 1), - (etac allE 1), - (etac exE 1), - (rtac (LeastI RS conjunct1) 1), - (atac 1), - (etac allE 1), - (etac exE 1), - (rtac (LeastI RS conjunct2) 1), - (atac 1) - ]); + \ chain(%m. Y(Least(%j. m ! m. P(Y(LEAST j::nat. m - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (etac allE 1), - (etac exE 1), - (etac (LeastI RS conjunct2) 1) - ]); +val prems = goal thy + "[| ! i. ? j. i < j & P(Y(j)) |] ==> ! m. P(Y(LEAST j::nat. m'a::cpo); ! i. ? j. i < j & P(Y(j)) |] ==>\ - \ lub(range(Y)) = lub(range(%m. Y(Least(%j. m - [ - (cut_facts_tac prems 1), - (rtac antisym_less 1), - (rtac lub_mono 1), - (atac 1), - (rtac adm_disj_lemma7 1), - (atac 1), - (atac 1), - (strip_tac 1), - (rtac (chain_mono RS mp) 1), - (atac 1), - (etac allE 1), - (etac exE 1), - (rtac (LeastI RS conjunct1) 1), - (atac 1), - (rtac lub_mono3 1), - (rtac adm_disj_lemma7 1), - (atac 1), - (atac 1), - (atac 1), - (strip_tac 1), - (rtac exI 1), - (rtac (chain_mono RS mp) 1), - (atac 1), - (rtac lessI 1) - ]); + \ lub(range(Y)) = lub(range(%m. Y(Least(%j. m'a::cpo); ! i. ? j. i < j & P(Y(j)) |] ==>\ - \ ? X. chain(X) & (! n. P(X(n))) & lub(range(Y)) = lub(range(X))" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("x","%m. Y(Least(%j. m P(Y(j))|]==>P(lub(range(Y)))" - (fn prems => - [ - (cut_facts_tac prems 1), - (etac adm_disj_lemma2 1), - (etac adm_disj_lemma6 1), - (atac 1) - ]); +val prems = goal thy + "[| adm(P); chain(Y);? i. ! j. i < j --> P(Y(j))|]==>P(lub(range(Y)))"; +by (cut_facts_tac prems 1); +by (etac adm_disj_lemma2 1); +by (etac adm_disj_lemma6 1); +by (atac 1); +qed "adm_disj_lemma12"; -in -val adm_lemma11 = prove_goal thy -"[| adm(P); chain(Y); ! i. ? j. i < j & P(Y(j)) |]==>P(lub(range(Y)))" - (fn prems => - [ - (cut_facts_tac prems 1), - (etac adm_disj_lemma2 1), - (etac adm_disj_lemma10 1), - (atac 1) - ]); +val prems = goal thy +"[| adm(P); chain(Y); ! i. ? j. i < j & P(Y(j)) |]==>P(lub(range(Y)))"; +by (cut_facts_tac prems 1); +by (etac adm_disj_lemma2 1); +by (etac adm_disj_lemma10 1); +by (atac 1); +qed "adm_lemma11"; -val adm_disj = prove_goal thy - "!!P. [| adm P; adm Q |] ==> adm(%x. P x | Q x)" - (fn prems => - [ - (rtac admI 1), - (rtac (adm_disj_lemma1 RS disjE) 1), - (atac 1), - (rtac disjI2 1), - (etac adm_disj_lemma12 1), - (atac 1), - (atac 1), - (rtac disjI1 1), - (etac adm_lemma11 1), - (atac 1), - (atac 1) - ]); +val prems = goal thy + "!!P. [| adm P; adm Q |] ==> adm(%x. P x | Q x)"; +by (rtac admI 1); +by (rtac (adm_disj_lemma1 RS disjE) 1); +by (atac 1); +by (rtac disjI2 1); +by (etac adm_disj_lemma12 1); +by (atac 1); +by (atac 1); +by (rtac disjI1 1); +by (etac adm_lemma11 1); +by (atac 1); +by (atac 1); +qed "adm_disj"; -end; bind_thm("adm_lemma11",adm_lemma11); bind_thm("adm_disj",adm_disj); -qed_goal "adm_imp" thy - "!!P. [| adm(%x.~(P x)); adm Q |] ==> adm(%x. P x --> Q x)" (K [ - (subgoal_tac "(%x. P x --> Q x) = (%x. ~P x | Q x)" 1), - (etac ssubst 1), - (etac adm_disj 1), - (atac 1), - (Simp_tac 1) - ]); +val prems = goal thy + "!!P. [| adm(%x.~(P x)); adm Q |] ==> adm(%x. P x --> Q x)"; +by (subgoal_tac "(%x. P x --> Q x) = (%x. ~P x | Q x)" 1); +by (etac ssubst 1); +by (etac adm_disj 1); +by (atac 1); +by (Simp_tac 1); +qed "adm_imp"; Goal "[| adm (%x. P x --> Q x); adm (%x. Q x --> P x) |] \ \ ==> adm (%x. P x = Q x)"; @@ -870,16 +770,16 @@ qed"adm_iff"; -qed_goal "adm_not_conj" thy -"[| adm (%x. ~ P x); adm (%x. ~ Q x) |] ==> adm (%x. ~ (P x & Q x))"(fn prems=>[ - cut_facts_tac prems 1, - subgoal_tac - "(%x. ~ (P x & Q x)) = (%x. ~ P x | ~ Q x)" 1, - rtac ext 2, - fast_tac HOL_cs 2, - etac ssubst 1, - etac adm_disj 1, - atac 1]); +val prems= goal thy +"[| adm (%x. ~ P x); adm (%x. ~ Q x) |] ==> adm (%x. ~ (P x & Q x))"; +by (cut_facts_tac prems 1); +by (subgoal_tac "(%x. ~ (P x & Q x)) = (%x. ~ P x | ~ Q x)" 1); +by (rtac ext 2); +by (fast_tac HOL_cs 2); +by (etac ssubst 1); +by (etac adm_disj 1); +by (atac 1); +qed "adm_not_conj"; bind_thms ("adm_lemmas", [adm_imp,adm_disj,adm_eq,adm_not_UU,adm_UU_not_less, adm_all2,adm_not_less,adm_not_conj,adm_iff]); diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Fun1.ML --- a/src/HOLCF/Fun1.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Fun1.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,38 +3,30 @@ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for fun1.thy +Definition of the partial ordering for the type of all functions => (fun) *) -open Fun1; - (* ------------------------------------------------------------------------ *) (* less_fun is a partial order on 'a => 'b *) (* ------------------------------------------------------------------------ *) -qed_goalw "refl_less_fun" thy [less_fun_def] "(f::'a::term =>'b::po) << f" -(fn prems => - [ - (fast_tac (HOL_cs addSIs [refl_less]) 1) - ]); +val prems = goalw thy [less_fun_def] "(f::'a::term =>'b::po) << f"; +by (fast_tac (HOL_cs addSIs [refl_less]) 1); +qed "refl_less_fun"; -qed_goalw "antisym_less_fun" Fun1.thy [less_fun_def] - "[|(f1::'a::term =>'b::po) << f2; f2 << f1|] ==> f1 = f2" -(fn prems => - [ - (cut_facts_tac prems 1), - (stac expand_fun_eq 1), - (fast_tac (HOL_cs addSIs [antisym_less]) 1) - ]); +val prems = goalw Fun1.thy [less_fun_def] + "[|(f1::'a::term =>'b::po) << f2; f2 << f1|] ==> f1 = f2"; +by (cut_facts_tac prems 1); +by (stac expand_fun_eq 1); +by (fast_tac (HOL_cs addSIs [antisym_less]) 1); +qed "antisym_less_fun"; -qed_goalw "trans_less_fun" Fun1.thy [less_fun_def] - "[|(f1::'a::term =>'b::po) << f2; f2 << f3 |] ==> f1 << f3" -(fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (rtac trans_less 1), - (etac allE 1), - (atac 1), - ((etac allE 1) THEN (atac 1)) - ]); +val prems = goalw Fun1.thy [less_fun_def] + "[|(f1::'a::term =>'b::po) << f2; f2 << f3 |] ==> f1 << f3"; +by (cut_facts_tac prems 1); +by (strip_tac 1); +by (rtac trans_less 1); +by (etac allE 1); +by (atac 1); +by ((etac allE 1) THEN (atac 1)); +qed "trans_less_fun"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Fun2.ML --- a/src/HOLCF/Fun2.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Fun2.ML Tue Jul 04 15:58:11 2000 +0200 @@ -1,120 +1,94 @@ -(* Title: HOLCF/fun2.ML +(* Title: HOLCF/Fun2.ML ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for fun2.thy +Class Instance =>::(term,po)po *) -open Fun2; - (* for compatibility with old HOLCF-Version *) -qed_goal "inst_fun_po" thy "(op <<)=(%f g.!x. f x << g x)" - (fn prems => - [ - (fold_goals_tac [less_fun_def]), - (rtac refl 1) - ]); +val prems = goal thy "(op <<)=(%f g.!x. f x << g x)"; +by (fold_goals_tac [less_fun_def]); +by (rtac refl 1); +qed "inst_fun_po"; (* ------------------------------------------------------------------------ *) (* Type 'a::term => 'b::pcpo is pointed *) (* ------------------------------------------------------------------------ *) -qed_goal "minimal_fun" thy "(%z. UU) << x" -(fn prems => - [ - (simp_tac (simpset() addsimps [inst_fun_po,minimal]) 1) - ]); +val prems = goal thy "(%z. UU) << x"; +by (simp_tac (simpset() addsimps [inst_fun_po,minimal]) 1); +qed "minimal_fun"; bind_thm ("UU_fun_def",minimal_fun RS minimal2UU RS sym); -qed_goal "least_fun" thy "? x::'a=>'b::pcpo.!y. x< - [ - (res_inst_tac [("x","(%z. UU)")] exI 1), - (rtac (minimal_fun RS allI) 1) - ]); +val prems = goal thy "? x::'a=>'b::pcpo.!y. x< - [ - (stac inst_fun_po 1), - (fold_goals_tac [less_fun_def]), - (rtac refl 1) - ]); +val prems = goal thy "(f1 << f2) = (! x. f1(x) << f2(x))"; +by (stac inst_fun_po 1); +by (rtac refl 1); +qed "less_fun"; (* ------------------------------------------------------------------------ *) (* chains of functions yield chains in the po range *) (* ------------------------------------------------------------------------ *) -qed_goal "ch2ch_fun" thy - "chain(S::nat=>('a=>'b::po)) ==> chain(% i. S(i)(x))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rewtac chain), - (rtac allI 1), - (rtac spec 1), - (rtac (less_fun RS subst) 1), - (etac allE 1), - (atac 1) - ]); +Goalw [chain] "chain(S::nat=>('a=>'b::po)) ==> chain(% i. S(i)(x))"; +by (asm_full_simp_tac (simpset() addsimps [less_fun]) 1); +qed "ch2ch_fun"; (* ------------------------------------------------------------------------ *) (* upper bounds of function chains yield upper bound in the po range *) (* ------------------------------------------------------------------------ *) -qed_goal "ub2ub_fun" Fun2.thy - " range(S::nat=>('a::term => 'b::po)) <| u ==> range(%i. S i x) <| u(x)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (rtac allE 1), - (rtac (less_fun RS subst) 1), - (etac (ub_rangeE RS spec) 1), - (atac 1) - ]); +val prems = goal Fun2.thy + " range(S::nat=>('a::term => 'b::po)) <| u ==> range(%i. S i x) <| u(x)"; +by (cut_facts_tac prems 1); +by (rtac ub_rangeI 1); +by (rtac allI 1); +by (rtac allE 1); +by (rtac (less_fun RS subst) 1); +by (etac (ub_rangeE RS spec) 1); +by (atac 1); +qed "ub2ub_fun"; (* ------------------------------------------------------------------------ *) (* Type 'a::term => 'b::pcpo is chain complete *) (* ------------------------------------------------------------------------ *) -qed_goal "lub_fun" Fun2.thy +val prems = goal Fun2.thy "chain(S::nat=>('a::term => 'b::cpo)) ==> \ -\ range(S) <<| (% x. lub(range(% i. S(i)(x))))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_lubI 1), - (rtac conjI 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (stac less_fun 1), - (rtac allI 1), - (rtac is_ub_thelub 1), - (etac ch2ch_fun 1), - (strip_tac 1), - (stac less_fun 1), - (rtac allI 1), - (rtac is_lub_thelub 1), - (etac ch2ch_fun 1), - (etac ub2ub_fun 1) - ]); +\ range(S) <<| (% x. lub(range(% i. S(i)(x))))"; +by (cut_facts_tac prems 1); +by (rtac is_lubI 1); +by (rtac conjI 1); +by (rtac ub_rangeI 1); +by (rtac allI 1); +by (stac less_fun 1); +by (rtac allI 1); +by (rtac is_ub_thelub 1); +by (etac ch2ch_fun 1); +by (strip_tac 1); +by (stac less_fun 1); +by (rtac allI 1); +by (rtac is_lub_thelub 1); +by (etac ch2ch_fun 1); +by (etac ub2ub_fun 1); +qed "lub_fun"; bind_thm ("thelub_fun", lub_fun RS thelubI); (* chain ?S1 ==> lub (range ?S1) = (%x. lub (range (%i. ?S1 i x))) *) -qed_goal "cpo_fun" Fun2.thy - "chain(S::nat=>('a::term => 'b::cpo)) ==> ? x. range(S) <<| x" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac exI 1), - (etac lub_fun 1) - ]); +val prems = goal Fun2.thy + "chain(S::nat=>('a::term => 'b::cpo)) ==> ? x. range(S) <<| x"; +by (cut_facts_tac prems 1); +by (rtac exI 1); +by (etac lub_fun 1); +qed "cpo_fun"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Fun3.ML --- a/src/HOLCF/Fun3.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Fun3.ML Tue Jul 04 15:58:11 2000 +0200 @@ -1,14 +1,10 @@ -(* Title: HOLCF/fun3.ML +(* Title: HOLCF/Fun3.ML ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen *) -open Fun3; - (* for compatibility with old HOLCF-Version *) -qed_goal "inst_fun_pcpo" thy "UU = (%x. UU)" - (fn prems => - [ - (simp_tac (HOL_ss addsimps [UU_def,UU_fun_def]) 1) - ]); +val prems = goal thy "UU = (%x. UU)"; +by (simp_tac (HOL_ss addsimps [UU_def,UU_fun_def]) 1); +qed "inst_fun_pcpo"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/HOLCF.ML --- a/src/HOLCF/HOLCF.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/HOLCF.ML Tue Jul 04 15:58:11 2000 +0200 @@ -4,8 +4,6 @@ Copyright 1993 Technische Universitaet Muenchen *) -open HOLCF; - use"adm.ML"; simpset_ref() := simpset() addSolver diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Lift2.ML --- a/src/HOLCF/Lift2.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Lift2.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,18 +3,14 @@ Author: Olaf Mueller Copyright 1996 Technische Universitaet Muenchen -Theorems for Lift2.thy +Class Instance lift::(term)po *) -open Lift2; - (* for compatibility with old HOLCF-Version *) -qed_goal "inst_lift_po" thy "(op <<)=(%x y. x=y|x=Undef)" - (fn prems => - [ - (fold_goals_tac [less_lift_def]), - (rtac refl 1) - ]); +val prems = goal thy "(op <<)=(%x y. x=y|x=Undef)"; +by (fold_goals_tac [less_lift_def]); +by (rtac refl 1); +qed "inst_lift_po"; (* -------------------------------------------------------------------------*) (* type ('a)lift is pointed *) @@ -26,12 +22,10 @@ bind_thm ("UU_lift_def",minimal_lift RS minimal2UU RS sym); -qed_goal "least_lift" thy "? x::'a lift.!y. x< - [ - (res_inst_tac [("x","Undef")] exI 1), - (rtac (minimal_lift RS allI) 1) - ]); +val prems = goal thy "? x::'a lift.!y. x<('a)lift)|] \ -\ ==> ? j.!i. j~Y(i)=Undef"; +Goal "[| ? j.~Y(j)=Undef; chain(Y::nat=>('a)lift) |] \ +\ ==> ? j.!i. j~Y(i)=Undef"; by Safe_tac; by (Step_tac 1); by (strip_tac 1); @@ -70,8 +63,7 @@ (* Tailoring flat_imp_chfin of Fix.ML to lift *) -Goal - "(! Y. chain(Y::nat=>('a)lift)-->(? n. max_in_chain n Y))"; +Goal "(! Y. chain(Y::nat=>('a)lift)-->(? n. max_in_chain n Y))"; by (rewtac max_in_chain_def); by (strip_tac 1); by (res_inst_tac [("P","!i. Y(i)=Undef")] case_split_thm 1); @@ -81,7 +73,6 @@ by (etac spec 1); by (rtac sym 1); by (etac spec 1); - by (subgoal_tac "!x y. x<<(y::('a)lift) --> x=Undef | x=y" 1); by (simp_tac (simpset() addsimps [inst_lift_po]) 2); by (rtac (chain_mono2_po RS exE) 1); @@ -109,8 +100,7 @@ (* Main Lemma: cpo_lift *) -Goal - "!!Y. chain(Y::nat=>('a)lift) ==> ? x. range(Y) <<|x"; +Goal "chain(Y::nat=>('a)lift) ==> ? x. range(Y) <<|x"; by (cut_inst_tac [] flat_imp_chfin_poo 1); by (Step_tac 1); by Safe_tac; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Lift3.ML --- a/src/HOLCF/Lift3.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Lift3.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,16 +3,14 @@ Author: Olaf Mueller Copyright 1996 Technische Universitaet Muenchen -Theorems for Lift3.thy +Class Instance lift::(term)pcpo *) (* for compatibility with old HOLCF-Version *) -qed_goal "inst_lift_pcpo" thy "UU = Undef" - (fn prems => - [ - (simp_tac (HOL_ss addsimps [UU_def,UU_lift_def]) 1) - ]); +val prems = goal thy "UU = Undef"; +by (simp_tac (HOL_ss addsimps [UU_def,UU_lift_def]) 1); +qed "inst_lift_pcpo"; (* ----------------------------------------------------------- *) (* In lift.simps Undef is replaced by UU *) @@ -105,10 +103,10 @@ bind_thm("Undef_eq_UU", inst_lift_pcpo RS sym); -val DefE = prove_goal thy "Def x = UU ==> R" - (fn prems => [ - cut_facts_tac prems 1, - asm_full_simp_tac (HOL_ss addsimps [Def_not_UU]) 1]); +val prems = goal thy "Def x = UU ==> R"; +by (cut_facts_tac prems 1); +by (asm_full_simp_tac (HOL_ss addsimps [Def_not_UU]) 1); +qed "DefE"; val prems = goal thy "[| x = Def s; x = UU |] ==> R"; by (cut_facts_tac prems 1); diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/One.ML --- a/src/HOLCF/One.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/One.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,31 +3,25 @@ Author: Oscar Slotosch Copyright 1997 Technische Universitaet Muenchen -Lemmas for One.thy +The unit domain *) -open One; - (* ------------------------------------------------------------------------ *) (* Exhaustion and Elimination for type one *) (* ------------------------------------------------------------------------ *) -qed_goalw "Exh_one" thy [ONE_def] "t=UU | t = ONE" - (fn prems => - [ - (lift.induct_tac "t" 1), - (Simp_tac 1), - (Simp_tac 1) - ]); +val prems = goalw thy [ONE_def] "t=UU | t = ONE"; +by (lift.induct_tac "t" 1); +by (Simp_tac 1); +by (Simp_tac 1); +qed "Exh_one"; -qed_goal "oneE" thy - "[| p=UU ==> Q; p = ONE ==>Q|] ==>Q" - (fn prems => - [ - (rtac (Exh_one RS disjE) 1), - (eresolve_tac prems 1), - (eresolve_tac prems 1) - ]); +val prems = goal thy + "[| p=UU ==> Q; p = ONE ==>Q|] ==>Q"; +by (rtac (Exh_one RS disjE) 1); +by (eresolve_tac prems 1); +by (eresolve_tac prems 1); +qed "oneE"; (* ------------------------------------------------------------------------ *) (* tactic for one-thms *) diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Porder.ML --- a/src/HOLCF/Porder.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Porder.ML Tue Jul 04 15:58:11 2000 +0200 @@ -1,49 +1,48 @@ -(* Title: HOLCF/Porder.thy +(* Title: HOLCF/Porder ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for theory Porder.thy +Conservative extension of theory Porder0 by constant definitions *) (* ------------------------------------------------------------------------ *) (* lubs are unique *) (* ------------------------------------------------------------------------ *) -qed_goalw "unique_lub" thy [is_lub, is_ub] - "[| S <<| x ; S <<| y |] ==> x=y" -( fn prems => - [ - (cut_facts_tac prems 1), - (etac conjE 1), - (etac conjE 1), - (rtac antisym_less 1), - (rtac mp 1),((etac allE 1) THEN (atac 1) THEN (atac 1)), - (rtac mp 1),((etac allE 1) THEN (atac 1) THEN (atac 1)) - ]); + +val prems = goalw thy [is_lub, is_ub] + "[| S <<| x ; S <<| y |] ==> x=y"; +by (cut_facts_tac prems 1); +by (etac conjE 1); +by (etac conjE 1); +by (rtac antisym_less 1); +by (rtac mp 1); +by ((etac allE 1) THEN (atac 1) THEN (atac 1)); +by (rtac mp 1); +by ((etac allE 1) THEN (atac 1) THEN (atac 1)); +qed "unique_lub"; (* ------------------------------------------------------------------------ *) (* chains are monotone functions *) (* ------------------------------------------------------------------------ *) -qed_goalw "chain_mono" thy [chain] "chain F ==> x F x< - [ - (cut_facts_tac prems 1), - (induct_tac "y" 1), - (rtac impI 1), - (etac less_zeroE 1), - (stac less_Suc_eq 1), - (strip_tac 1), - (etac disjE 1), - (rtac trans_less 1), - (etac allE 2), - (atac 2), - (fast_tac HOL_cs 1), - (hyp_subst_tac 1), - (etac allE 1), - (atac 1) - ]); +val prems = goalw thy [chain] "chain F ==> x F x< F x << F y"; by (rtac (le_imp_less_or_eq RS disjE) 1); @@ -56,16 +55,16 @@ (* ------------------------------------------------------------------------ *) -(* The range of a chain is a totaly ordered << *) +(* The range of a chain is a totally ordered << *) (* ------------------------------------------------------------------------ *) -qed_goalw "chain_tord" thy [tord] -"!!F. chain(F) ==> tord(range(F))" - (fn _ => - [ - Safe_tac, - (rtac nat_less_cases 1), - (ALLGOALS (fast_tac (claset() addIs [refl_less, chain_mono RS mp])))]); +val _ = goalw thy [tord] +"!!F. chain(F) ==> tord(range(F))"; +by (Safe_tac); +by (rtac nat_less_cases 1); +by (ALLGOALS (fast_tac (claset() addIs [refl_less, chain_mono RS mp]))); +qed "chain_tord"; + (* ------------------------------------------------------------------------ *) (* technical lemmas about lub and is_lub *) @@ -106,57 +105,47 @@ (* access to some definition as inference rule *) (* ------------------------------------------------------------------------ *) -qed_goalw "is_lubE" thy [is_lub] - "S <<| x ==> S <| x & (! u. S <| u --> x << u)" -(fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); +val prems = goalw thy [is_lub] + "S <<| x ==> S <| x & (! u. S <| u --> x << u)"; +by (cut_facts_tac prems 1); +by (atac 1); +qed "is_lubE"; -qed_goalw "is_lubI" thy [is_lub] - "S <| x & (! u. S <| u --> x << u) ==> S <<| x" -(fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); +val prems = goalw thy [is_lub] + "S <| x & (! u. S <| u --> x << u) ==> S <<| x"; +by (cut_facts_tac prems 1); +by (atac 1); +qed "is_lubI"; -qed_goalw "chainE" thy [chain] "chain F ==> !i. F(i) << F(Suc(i))" -(fn prems => - [ - (cut_facts_tac prems 1), - (atac 1)]); +val prems = goalw thy [chain] "chain F ==> !i. F(i) << F(Suc(i))"; +by (cut_facts_tac prems 1); +by (atac 1); +qed "chainE"; -qed_goalw "chainI" thy [chain] "!i. F i << F(Suc i) ==> chain F" -(fn prems => - [ - (cut_facts_tac prems 1), - (atac 1)]); +val prems = goalw thy [chain] "!i. F i << F(Suc i) ==> chain F"; +by (cut_facts_tac prems 1); +by (atac 1); +qed "chainI"; (* ------------------------------------------------------------------------ *) (* technical lemmas about (least) upper bounds of chains *) (* ------------------------------------------------------------------------ *) -qed_goalw "ub_rangeE" thy [is_ub] "range S <| x ==> !i. S(i) << x" -(fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (rtac mp 1), - (etac spec 1), - (rtac rangeI 1) - ]); +val prems = goalw thy [is_ub] "range S <| x ==> !i. S(i) << x"; +by (cut_facts_tac prems 1); +by (strip_tac 1); +by (rtac mp 1); +by (etac spec 1); +by (rtac rangeI 1); +qed "ub_rangeE"; -qed_goalw "ub_rangeI" thy [is_ub] "!i. S i << x ==> range S <| x" -(fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (etac rangeE 1), - (hyp_subst_tac 1), - (etac spec 1) - ]); +val prems = goalw thy [is_ub] "!i. S i << x ==> range S <| x"; +by (cut_facts_tac prems 1); +by (strip_tac 1); +by (etac rangeE 1); +by (hyp_subst_tac 1); +by (etac spec 1); +qed "ub_rangeI"; bind_thm ("is_ub_lub", is_lubE RS conjunct1 RS ub_rangeE RS spec); (* range(?S1) <<| ?x1 ==> ?S1(?x) << ?x1 *) @@ -168,38 +157,34 @@ (* results about finite chains *) (* ------------------------------------------------------------------------ *) -qed_goalw "lub_finch1" thy [max_in_chain_def] - "[| chain C; max_in_chain i C|] ==> range C <<| C i" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_lubI 1), - (rtac conjI 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (res_inst_tac [("m","i")] nat_less_cases 1), - (rtac (antisym_less_inverse RS conjunct2) 1), - (etac (disjI1 RS less_or_eq_imp_le RS rev_mp) 1), - (etac spec 1), - (rtac (antisym_less_inverse RS conjunct2) 1), - (etac (disjI2 RS less_or_eq_imp_le RS rev_mp) 1), - (etac spec 1), - (etac (chain_mono RS mp) 1), - (atac 1), - (strip_tac 1), - (etac (ub_rangeE RS spec) 1) - ]); +val prems = goalw thy [max_in_chain_def] + "[| chain C; max_in_chain i C|] ==> range C <<| C i"; +by (cut_facts_tac prems 1); +by (rtac is_lubI 1); +by (rtac conjI 1); +by (rtac ub_rangeI 1); +by (rtac allI 1); +by (res_inst_tac [("m","i")] nat_less_cases 1); +by (rtac (antisym_less_inverse RS conjunct2) 1); +by (etac (disjI1 RS less_or_eq_imp_le RS rev_mp) 1); +by (etac spec 1); +by (rtac (antisym_less_inverse RS conjunct2) 1); +by (etac (disjI2 RS less_or_eq_imp_le RS rev_mp) 1); +by (etac spec 1); +by (etac (chain_mono RS mp) 1); +by (atac 1); +by (strip_tac 1); +by (etac (ub_rangeE RS spec) 1); +qed "lub_finch1"; -qed_goalw "lub_finch2" thy [finite_chain_def] - "finite_chain(C) ==> range(C) <<| C(@ i. max_in_chain i C)" - (fn prems=> - [ - (cut_facts_tac prems 1), - (rtac lub_finch1 1), - (etac conjunct1 1), - (rtac (select_eq_Ex RS iffD2) 1), - (etac conjunct2 1) - ]); +val prems= goalw thy [finite_chain_def] + "finite_chain(C) ==> range(C) <<| C(@ i. max_in_chain i C)"; +by (cut_facts_tac prems 1); +by (rtac lub_finch1 1); +by (etac conjunct1 1); +by (rtac (select_eq_Ex RS iffD2) 1); +by (etac conjunct2 1); +qed "lub_finch2"; Goal "x< chain (%i. if i=0 then x else y)"; @@ -210,16 +195,14 @@ by (Asm_simp_tac 1); qed "bin_chain"; -qed_goalw "bin_chainmax" thy [max_in_chain_def,le_def] - "x< max_in_chain (Suc 0) (%i. if (i=0) then x else y)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac allI 1), - (induct_tac "j" 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1) - ]); +val prems = goalw thy [max_in_chain_def,le_def] + "x< max_in_chain (Suc 0) (%i. if (i=0) then x else y)"; +by (cut_facts_tac prems 1); +by (rtac allI 1); +by (induct_tac "j" 1); +by (Asm_simp_tac 1); +by (Asm_simp_tac 1); +qed "bin_chainmax"; Goal "x << y ==> range(%i::nat. if (i=0) then x else y) <<| y"; by (res_inst_tac [("s","if (Suc 0) = 0 then x else y")] subst 1 diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Porder0.ML --- a/src/HOLCF/Porder0.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Porder0.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,25 +3,21 @@ Author: Oscar Slotosch Copyright 1997 Technische Universitaet Muenchen - derive the characteristic axioms for the characteristic constants *) - -open Porder0; + derive the characteristic axioms for the characteristic constants +*) AddIffs [refl_less]; (* ------------------------------------------------------------------------ *) (* minimal fixes least element *) (* ------------------------------------------------------------------------ *) -bind_thm("minimal2UU",allI RS (prove_goal thy "!x::'a::po. uu<uu=(@u.!y. u< - [ - (cut_facts_tac prems 1), - (rtac antisym_less 1), - (etac spec 1), - (res_inst_tac [("a","uu")] selectI2 1), - (atac 1), - (etac spec 1) - ]))); +Goal "!x::'a::po. uu< uu=(@u.!y. u< - [ - (EVERY1 [rtac CollectI, rtac exI,rtac exI, rtac refl]) - ]); +val prems = goalw Sprod0.thy [Sprod_def] + "(Spair_Rep a b):Sprod"; +by (EVERY1 [rtac CollectI, rtac exI,rtac exI, rtac refl]); +qed "SprodI"; -qed_goal "inj_on_Abs_Sprod" Sprod0.thy - "inj_on Abs_Sprod Sprod" -(fn prems => - [ - (rtac inj_on_inverseI 1), - (etac Abs_Sprod_inverse 1) - ]); +val prems = goal Sprod0.thy + "inj_on Abs_Sprod Sprod"; +by (rtac inj_on_inverseI 1); +by (etac Abs_Sprod_inverse 1); +qed "inj_on_Abs_Sprod"; (* ------------------------------------------------------------------------ *) (* Strictness and definedness of Spair_Rep *) (* ------------------------------------------------------------------------ *) -qed_goalw "strict_Spair_Rep" Sprod0.thy [Spair_Rep_def] - "(a=UU | b=UU) ==> (Spair_Rep a b) = (Spair_Rep UU UU)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac ext 1), - (rtac ext 1), - (rtac iffI 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); +val prems = goalw Sprod0.thy [Spair_Rep_def] + "(a=UU | b=UU) ==> (Spair_Rep a b) = (Spair_Rep UU UU)"; +by (cut_facts_tac prems 1); +by (rtac ext 1); +by (rtac ext 1); +by (rtac iffI 1); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +qed "strict_Spair_Rep"; -qed_goalw "defined_Spair_Rep_rev" Sprod0.thy [Spair_Rep_def] - "(Spair_Rep a b) = (Spair_Rep UU UU) ==> (a=UU | b=UU)" - (fn prems => - [ - (case_tac "a=UU|b=UU" 1), - (atac 1), - (rtac disjI1 1), - (rtac ((hd prems) RS fun_cong RS fun_cong RS iffD2 RS mp RS - conjunct1 RS sym) 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); - +val prems = goalw Sprod0.thy [Spair_Rep_def] + "(Spair_Rep a b) = (Spair_Rep UU UU) ==> (a=UU | b=UU)"; +by (case_tac "a=UU|b=UU" 1); +by (atac 1); +by (rtac disjI1 1); +by (rtac ((hd prems) RS fun_cong RS fun_cong RS iffD2 RS mp RS conjunct1 RS sym) 1); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +qed "defined_Spair_Rep_rev"; (* ------------------------------------------------------------------------ *) (* injectivity of Spair_Rep and Ispair *) (* ------------------------------------------------------------------------ *) -qed_goalw "inject_Spair_Rep" Sprod0.thy [Spair_Rep_def] -"[|~aa=UU ; ~ba=UU ; Spair_Rep a b = Spair_Rep aa ba |] ==> a=aa & b=ba" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac ((nth_elem (2,prems)) RS fun_cong RS fun_cong - RS iffD1 RS mp) 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); +val prems = goalw Sprod0.thy [Spair_Rep_def] +"[|~aa=UU ; ~ba=UU ; Spair_Rep a b = Spair_Rep aa ba |] ==> a=aa & b=ba"; +by (cut_facts_tac prems 1); +by (rtac ((nth_elem (2,prems)) RS fun_cong RS fun_cong RS iffD1 RS mp) 1); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +qed "inject_Spair_Rep"; -qed_goalw "inject_Ispair" Sprod0.thy [Ispair_def] - "[|~aa=UU ; ~ba=UU ; Ispair a b = Ispair aa ba |] ==> a=aa & b=ba" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac inject_Spair_Rep 1), - (atac 1), - (etac (inj_on_Abs_Sprod RS inj_onD) 1), - (rtac SprodI 1), - (rtac SprodI 1) - ]); +val prems = goalw Sprod0.thy [Ispair_def] + "[|~aa=UU ; ~ba=UU ; Ispair a b = Ispair aa ba |] ==> a=aa & b=ba"; +by (cut_facts_tac prems 1); +by (etac inject_Spair_Rep 1); +by (atac 1); +by (etac (inj_on_Abs_Sprod RS inj_onD) 1); +by (rtac SprodI 1); +by (rtac SprodI 1); +qed "inject_Ispair"; (* ------------------------------------------------------------------------ *) (* strictness and definedness of Ispair *) (* ------------------------------------------------------------------------ *) -qed_goalw "strict_Ispair" Sprod0.thy [Ispair_def] - "(a=UU | b=UU) ==> Ispair a b = Ispair UU UU" -(fn prems => - [ - (cut_facts_tac prems 1), - (etac (strict_Spair_Rep RS arg_cong) 1) - ]); +val prems = goalw Sprod0.thy [Ispair_def] + "(a=UU | b=UU) ==> Ispair a b = Ispair UU UU"; +by (cut_facts_tac prems 1); +by (etac (strict_Spair_Rep RS arg_cong) 1); +qed "strict_Ispair"; -qed_goalw "strict_Ispair1" Sprod0.thy [Ispair_def] - "Ispair UU b = Ispair UU UU" -(fn prems => - [ - (rtac (strict_Spair_Rep RS arg_cong) 1), - (rtac disjI1 1), - (rtac refl 1) - ]); +val prems = goalw Sprod0.thy [Ispair_def] + "Ispair UU b = Ispair UU UU"; +by (rtac (strict_Spair_Rep RS arg_cong) 1); +by (rtac disjI1 1); +by (rtac refl 1); +qed "strict_Ispair1"; -qed_goalw "strict_Ispair2" Sprod0.thy [Ispair_def] - "Ispair a UU = Ispair UU UU" -(fn prems => - [ - (rtac (strict_Spair_Rep RS arg_cong) 1), - (rtac disjI2 1), - (rtac refl 1) - ]); +val prems = goalw Sprod0.thy [Ispair_def] + "Ispair a UU = Ispair UU UU"; +by (rtac (strict_Spair_Rep RS arg_cong) 1); +by (rtac disjI2 1); +by (rtac refl 1); +qed "strict_Ispair2"; -qed_goal "strict_Ispair_rev" Sprod0.thy - "~Ispair x y = Ispair UU UU ==> ~x=UU & ~y=UU" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac (de_Morgan_disj RS subst) 1), - (etac contrapos 1), - (etac strict_Ispair 1) - ]); +val prems = goal Sprod0.thy + "~Ispair x y = Ispair UU UU ==> ~x=UU & ~y=UU"; +by (cut_facts_tac prems 1); +by (rtac (de_Morgan_disj RS subst) 1); +by (etac contrapos 1); +by (etac strict_Ispair 1); +qed "strict_Ispair_rev"; -qed_goalw "defined_Ispair_rev" Sprod0.thy [Ispair_def] - "Ispair a b = Ispair UU UU ==> (a = UU | b = UU)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac defined_Spair_Rep_rev 1), - (rtac (inj_on_Abs_Sprod RS inj_onD) 1), - (atac 1), - (rtac SprodI 1), - (rtac SprodI 1) - ]); +val prems = goalw Sprod0.thy [Ispair_def] + "Ispair a b = Ispair UU UU ==> (a = UU | b = UU)"; +by (cut_facts_tac prems 1); +by (rtac defined_Spair_Rep_rev 1); +by (rtac (inj_on_Abs_Sprod RS inj_onD) 1); +by (atac 1); +by (rtac SprodI 1); +by (rtac SprodI 1); +qed "defined_Ispair_rev"; -qed_goal "defined_Ispair" Sprod0.thy -"[|a~=UU; b~=UU|] ==> (Ispair a b) ~= (Ispair UU UU)" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac contrapos 1), - (etac defined_Ispair_rev 2), - (rtac (de_Morgan_disj RS iffD2) 1), - (etac conjI 1), - (atac 1) - ]); +val prems = goal Sprod0.thy +"[|a~=UU; b~=UU|] ==> (Ispair a b) ~= (Ispair UU UU)"; +by (cut_facts_tac prems 1); +by (rtac contrapos 1); +by (etac defined_Ispair_rev 2); +by (rtac (de_Morgan_disj RS iffD2) 1); +by (etac conjI 1); +by (atac 1); +qed "defined_Ispair"; (* ------------------------------------------------------------------------ *) (* Exhaustion of the strict product ** *) (* ------------------------------------------------------------------------ *) -qed_goalw "Exh_Sprod" Sprod0.thy [Ispair_def] - "z=Ispair UU UU | (? a b. z=Ispair a b & a~=UU & b~=UU)" -(fn prems => - [ - (rtac (rewrite_rule [Sprod_def] Rep_Sprod RS CollectE) 1), - (etac exE 1), - (etac exE 1), - (rtac (excluded_middle RS disjE) 1), - (rtac disjI2 1), - (rtac exI 1), - (rtac exI 1), - (rtac conjI 1), - (rtac (Rep_Sprod_inverse RS sym RS trans) 1), - (etac arg_cong 1), - (rtac (de_Morgan_disj RS subst) 1), - (atac 1), - (rtac disjI1 1), - (rtac (Rep_Sprod_inverse RS sym RS trans) 1), - (res_inst_tac [("f","Abs_Sprod")] arg_cong 1), - (etac trans 1), - (etac strict_Spair_Rep 1) - ]); +val prems = goalw Sprod0.thy [Ispair_def] + "z=Ispair UU UU | (? a b. z=Ispair a b & a~=UU & b~=UU)"; +by (rtac (rewrite_rule [Sprod_def] Rep_Sprod RS CollectE) 1); +by (etac exE 1); +by (etac exE 1); +by (rtac (excluded_middle RS disjE) 1); +by (rtac disjI2 1); +by (rtac exI 1); +by (rtac exI 1); +by (rtac conjI 1); +by (rtac (Rep_Sprod_inverse RS sym RS trans) 1); +by (etac arg_cong 1); +by (rtac (de_Morgan_disj RS subst) 1); +by (atac 1); +by (rtac disjI1 1); +by (rtac (Rep_Sprod_inverse RS sym RS trans) 1); +by (res_inst_tac [("f","Abs_Sprod")] arg_cong 1); +by (etac trans 1); +by (etac strict_Spair_Rep 1); +qed "Exh_Sprod"; (* ------------------------------------------------------------------------ *) (* general elimination rule for strict product *) (* ------------------------------------------------------------------------ *) -qed_goal "IsprodE" Sprod0.thy -"[|p=Ispair UU UU ==> Q ;!!x y. [|p=Ispair x y; x~=UU ; y~=UU|] ==> Q|] ==> Q" -(fn prems => - [ - (rtac (Exh_Sprod RS disjE) 1), - (etac (hd prems) 1), - (etac exE 1), - (etac exE 1), - (etac conjE 1), - (etac conjE 1), - (etac (hd (tl prems)) 1), - (atac 1), - (atac 1) - ]); +val prems = goal Sprod0.thy +"[|p=Ispair UU UU ==> Q ;!!x y. [|p=Ispair x y; x~=UU ; y~=UU|] ==> Q|] ==> Q"; +by (rtac (Exh_Sprod RS disjE) 1); +by (etac (hd prems) 1); +by (etac exE 1); +by (etac exE 1); +by (etac conjE 1); +by (etac conjE 1); +by (etac (hd (tl prems)) 1); +by (atac 1); +by (atac 1); +qed "IsprodE"; (* ------------------------------------------------------------------------ *) (* some results about the selectors Isfst, Issnd *) (* ------------------------------------------------------------------------ *) -qed_goalw "strict_Isfst" Sprod0.thy [Isfst_def] - "p=Ispair UU UU ==> Isfst p = UU" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac select_equality 1), - (rtac conjI 1), - (fast_tac HOL_cs 1), - (strip_tac 1), - (res_inst_tac [("P","Ispair UU UU = Ispair a b")] notE 1), - (rtac not_sym 1), - (rtac defined_Ispair 1), - (REPEAT (fast_tac HOL_cs 1)) - ]); +val prems = goalw Sprod0.thy [Isfst_def] + "p=Ispair UU UU ==> Isfst p = UU"; +by (cut_facts_tac prems 1); +by (rtac select_equality 1); +by (rtac conjI 1); +by (fast_tac HOL_cs 1); +by (strip_tac 1); +by (res_inst_tac [("P","Ispair UU UU = Ispair a b")] notE 1); +by (rtac not_sym 1); +by (rtac defined_Ispair 1); +by (REPEAT (fast_tac HOL_cs 1)); +qed "strict_Isfst"; -qed_goal "strict_Isfst1" Sprod0.thy - "Isfst(Ispair UU y) = UU" -(fn prems => - [ - (stac strict_Ispair1 1), - (rtac strict_Isfst 1), - (rtac refl 1) - ]); +val prems = goal Sprod0.thy + "Isfst(Ispair UU y) = UU"; +by (stac strict_Ispair1 1); +by (rtac strict_Isfst 1); +by (rtac refl 1); +qed "strict_Isfst1"; -qed_goal "strict_Isfst2" Sprod0.thy - "Isfst(Ispair x UU) = UU" -(fn prems => - [ - (stac strict_Ispair2 1), - (rtac strict_Isfst 1), - (rtac refl 1) - ]); +val prems = goal Sprod0.thy + "Isfst(Ispair x UU) = UU"; +by (stac strict_Ispair2 1); +by (rtac strict_Isfst 1); +by (rtac refl 1); +qed "strict_Isfst2"; -qed_goalw "strict_Issnd" Sprod0.thy [Issnd_def] - "p=Ispair UU UU ==>Issnd p=UU" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac select_equality 1), - (rtac conjI 1), - (fast_tac HOL_cs 1), - (strip_tac 1), - (res_inst_tac [("P","Ispair UU UU = Ispair a b")] notE 1), - (rtac not_sym 1), - (rtac defined_Ispair 1), - (REPEAT (fast_tac HOL_cs 1)) - ]); +val prems = goalw Sprod0.thy [Issnd_def] + "p=Ispair UU UU ==>Issnd p=UU"; +by (cut_facts_tac prems 1); +by (rtac select_equality 1); +by (rtac conjI 1); +by (fast_tac HOL_cs 1); +by (strip_tac 1); +by (res_inst_tac [("P","Ispair UU UU = Ispair a b")] notE 1); +by (rtac not_sym 1); +by (rtac defined_Ispair 1); +by (REPEAT (fast_tac HOL_cs 1)); +qed "strict_Issnd"; -qed_goal "strict_Issnd1" Sprod0.thy - "Issnd(Ispair UU y) = UU" -(fn prems => - [ - (stac strict_Ispair1 1), - (rtac strict_Issnd 1), - (rtac refl 1) - ]); +val prems = goal Sprod0.thy + "Issnd(Ispair UU y) = UU"; +by (stac strict_Ispair1 1); +by (rtac strict_Issnd 1); +by (rtac refl 1); +qed "strict_Issnd1"; -qed_goal "strict_Issnd2" Sprod0.thy - "Issnd(Ispair x UU) = UU" -(fn prems => - [ - (stac strict_Ispair2 1), - (rtac strict_Issnd 1), - (rtac refl 1) - ]); +val prems = goal Sprod0.thy + "Issnd(Ispair x UU) = UU"; +by (stac strict_Ispair2 1); +by (rtac strict_Issnd 1); +by (rtac refl 1); +qed "strict_Issnd2"; -qed_goalw "Isfst" Sprod0.thy [Isfst_def] - "[|x~=UU ;y~=UU |] ==> Isfst(Ispair x y) = x" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac select_equality 1), - (rtac conjI 1), - (strip_tac 1), - (res_inst_tac [("P","Ispair x y = Ispair UU UU")] notE 1), - (etac defined_Ispair 1), - (atac 1), - (atac 1), - (strip_tac 1), - (rtac (inject_Ispair RS conjunct1) 1), - (fast_tac HOL_cs 3), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); +val prems = goalw Sprod0.thy [Isfst_def] + "[|x~=UU ;y~=UU |] ==> Isfst(Ispair x y) = x"; +by (cut_facts_tac prems 1); +by (rtac select_equality 1); +by (rtac conjI 1); +by (strip_tac 1); +by (res_inst_tac [("P","Ispair x y = Ispair UU UU")] notE 1); +by (etac defined_Ispair 1); +by (atac 1); +by (atac 1); +by (strip_tac 1); +by (rtac (inject_Ispair RS conjunct1) 1); +by (fast_tac HOL_cs 3); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +qed "Isfst"; -qed_goalw "Issnd" Sprod0.thy [Issnd_def] - "[|x~=UU ;y~=UU |] ==> Issnd(Ispair x y) = y" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac select_equality 1), - (rtac conjI 1), - (strip_tac 1), - (res_inst_tac [("P","Ispair x y = Ispair UU UU")] notE 1), - (etac defined_Ispair 1), - (atac 1), - (atac 1), - (strip_tac 1), - (rtac (inject_Ispair RS conjunct2) 1), - (fast_tac HOL_cs 3), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); +val prems = goalw Sprod0.thy [Issnd_def] + "[|x~=UU ;y~=UU |] ==> Issnd(Ispair x y) = y"; +by (cut_facts_tac prems 1); +by (rtac select_equality 1); +by (rtac conjI 1); +by (strip_tac 1); +by (res_inst_tac [("P","Ispair x y = Ispair UU UU")] notE 1); +by (etac defined_Ispair 1); +by (atac 1); +by (atac 1); +by (strip_tac 1); +by (rtac (inject_Ispair RS conjunct2) 1); +by (fast_tac HOL_cs 3); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +qed "Issnd"; -qed_goal "Isfst2" Sprod0.thy "y~=UU ==>Isfst(Ispair x y)=x" -(fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), - (etac Isfst 1), - (atac 1), - (hyp_subst_tac 1), - (rtac strict_Isfst1 1) - ]); +val prems = goal Sprod0.thy "y~=UU ==>Isfst(Ispair x y)=x"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1); +by (etac Isfst 1); +by (atac 1); +by (hyp_subst_tac 1); +by (rtac strict_Isfst1 1); +qed "Isfst2"; -qed_goal "Issnd2" Sprod0.thy "~x=UU ==>Issnd(Ispair x y)=y" -(fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("Q","y=UU")] (excluded_middle RS disjE) 1), - (etac Issnd 1), - (atac 1), - (hyp_subst_tac 1), - (rtac strict_Issnd2 1) - ]); +val prems = goal Sprod0.thy "~x=UU ==>Issnd(Ispair x y)=y"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("Q","y=UU")] (excluded_middle RS disjE) 1); +by (etac Issnd 1); +by (atac 1); +by (hyp_subst_tac 1); +by (rtac strict_Issnd2 1); +qed "Issnd2"; (* ------------------------------------------------------------------------ *) @@ -342,42 +289,36 @@ addsimps [strict_Isfst1,strict_Isfst2,strict_Issnd1,strict_Issnd2, Isfst2,Issnd2]; -qed_goal "defined_IsfstIssnd" Sprod0.thy - "p~=Ispair UU UU ==> Isfst p ~= UU & Issnd p ~= UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p")] IsprodE 1), - (contr_tac 1), - (hyp_subst_tac 1), - (rtac conjI 1), - (asm_simp_tac Sprod0_ss 1), - (asm_simp_tac Sprod0_ss 1) - ]); +val prems = goal Sprod0.thy + "p~=Ispair UU UU ==> Isfst p ~= UU & Issnd p ~= UU"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("p","p")] IsprodE 1); +by (contr_tac 1); +by (hyp_subst_tac 1); +by (rtac conjI 1); +by (asm_simp_tac Sprod0_ss 1); +by (asm_simp_tac Sprod0_ss 1); +qed "defined_IsfstIssnd"; (* ------------------------------------------------------------------------ *) (* Surjective pairing: equivalent to Exh_Sprod *) (* ------------------------------------------------------------------------ *) -qed_goal "surjective_pairing_Sprod" Sprod0.thy - "z = Ispair(Isfst z)(Issnd z)" -(fn prems => - [ - (res_inst_tac [("z1","z")] (Exh_Sprod RS disjE) 1), - (asm_simp_tac Sprod0_ss 1), - (etac exE 1), - (etac exE 1), - (asm_simp_tac Sprod0_ss 1) - ]); +val prems = goal Sprod0.thy + "z = Ispair(Isfst z)(Issnd z)"; +by (res_inst_tac [("z1","z")] (Exh_Sprod RS disjE) 1); +by (asm_simp_tac Sprod0_ss 1); +by (etac exE 1); +by (etac exE 1); +by (asm_simp_tac Sprod0_ss 1); +qed "surjective_pairing_Sprod"; -qed_goal "Sel_injective_Sprod" thy - "[|Isfst x = Isfst y; Issnd x = Issnd y|] ==> x = y" -(fn prems => - [ - (cut_facts_tac prems 1), - (subgoal_tac "Ispair(Isfst x)(Issnd x)=Ispair(Isfst y)(Issnd y)" 1), - (rotate_tac ~1 1), - (asm_full_simp_tac(HOL_ss addsimps[surjective_pairing_Sprod RS sym])1), - (Asm_simp_tac 1) - ]); +val prems = goal thy + "[|Isfst x = Isfst y; Issnd x = Issnd y|] ==> x = y"; +by (cut_facts_tac prems 1); +by (subgoal_tac "Ispair(Isfst x)(Issnd x)=Ispair(Isfst y)(Issnd y)" 1); +by (rotate_tac ~1 1); +by (asm_full_simp_tac(HOL_ss addsimps[surjective_pairing_Sprod RS sym])1); +by (Asm_simp_tac 1); +qed "Sel_injective_Sprod"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Sprod1.ML --- a/src/HOLCF/Sprod1.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Sprod1.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,34 +3,28 @@ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for theory Sprod1.thy *) -open Sprod1; - (* ------------------------------------------------------------------------ *) (* less_sprod is a partial order on Sprod *) (* ------------------------------------------------------------------------ *) -qed_goalw "refl_less_sprod" thy [less_sprod_def]"(p::'a ** 'b) << p" -(fn prems => [(fast_tac (HOL_cs addIs [refl_less]) 1)]); +val prems = goalw thy [less_sprod_def]"(p::'a ** 'b) << p"; +by (fast_tac (HOL_cs addIs [refl_less]) 1); +qed "refl_less_sprod"; -qed_goalw "antisym_less_sprod" thy [less_sprod_def] - "[|(p1::'a ** 'b) << p2;p2 << p1|] ==> p1=p2" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac Sel_injective_Sprod 1), - (fast_tac (HOL_cs addIs [antisym_less]) 1), - (fast_tac (HOL_cs addIs [antisym_less]) 1) - ]); +val prems = goalw thy [less_sprod_def] + "[|(p1::'a ** 'b) << p2;p2 << p1|] ==> p1=p2"; +by (cut_facts_tac prems 1); +by (rtac Sel_injective_Sprod 1); +by (fast_tac (HOL_cs addIs [antisym_less]) 1); +by (fast_tac (HOL_cs addIs [antisym_less]) 1); +qed "antisym_less_sprod"; -qed_goalw "trans_less_sprod" thy [less_sprod_def] - "[|(p1::'a**'b) << p2;p2 << p3|] ==> p1 << p3" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - (fast_tac (HOL_cs addIs [trans_less]) 1), - (fast_tac (HOL_cs addIs [trans_less]) 1) - ]); +val prems = goalw thy [less_sprod_def] + "[|(p1::'a**'b) << p2;p2 << p3|] ==> p1 << p3"; +by (cut_facts_tac prems 1); +by (rtac conjI 1); +by (fast_tac (HOL_cs addIs [trans_less]) 1); +by (fast_tac (HOL_cs addIs [trans_less]) 1); +qed "trans_less_sprod"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Sprod2.ML --- a/src/HOLCF/Sprod2.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Sprod2.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,138 +3,108 @@ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for Sprod2.thy +Class Instance **::(pcpo,pcpo)po *) -open Sprod2; - (* for compatibility with old HOLCF-Version *) -qed_goal "inst_sprod_po" thy "(op <<)=(%x y. Isfst x< - [ - (fold_goals_tac [less_sprod_def]), - (rtac refl 1) - ]); +val prems = goal thy "(op <<)=(%x y. Isfst x< - [ - (simp_tac(Sprod0_ss addsimps[inst_sprod_po,minimal])1) - ]); +val prems = goal thy "Ispair UU UU << p"; +by (simp_tac(Sprod0_ss addsimps[inst_sprod_po,minimal])1); +qed "minimal_sprod"; bind_thm ("UU_sprod_def",minimal_sprod RS minimal2UU RS sym); -qed_goal "least_sprod" thy "? x::'a**'b.!y. x< - [ - (res_inst_tac [("x","Ispair UU UU")] exI 1), - (rtac (minimal_sprod RS allI) 1) - ]); +val prems = goal thy "? x::'a**'b.!y. x< - [ - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (res_inst_tac [("Q","xa=UU")] (excluded_middle RS disjE) 1), - (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), - (ftac notUU_I 1), - (atac 1), - (REPEAT(asm_simp_tac(Sprod0_ss - addsimps[inst_sprod_po,refl_less,minimal]) 1)) - ]); +val prems = goalw Sprod2.thy [monofun] "monofun(Ispair)"; +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 (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1); +by (ftac notUU_I 1); +by (atac 1); +by (REPEAT(asm_simp_tac(Sprod0_ss addsimps[inst_sprod_po,refl_less,minimal]) 1)); +qed "monofun_Ispair1"; -qed_goalw "monofun_Ispair2" Sprod2.thy [monofun] "monofun(Ispair(x))" -(fn prems => - [ - (strip_tac 1), - (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), - (res_inst_tac [("Q","xa=UU")] (excluded_middle RS disjE) 1), - (ftac notUU_I 1), - (atac 1), - (REPEAT(asm_simp_tac(Sprod0_ss - addsimps[inst_sprod_po,refl_less,minimal]) 1)) - ]); - +val prems = goalw Sprod2.thy [monofun] "monofun(Ispair(x))"; +by (strip_tac 1); +by (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1); +by (res_inst_tac [("Q","xa=UU")] (excluded_middle RS disjE) 1); +by (ftac notUU_I 1); +by (atac 1); +by (REPEAT(asm_simp_tac(Sprod0_ss addsimps[inst_sprod_po,refl_less,minimal]) 1)); +qed "monofun_Ispair2"; -qed_goal "monofun_Ispair" Sprod2.thy - "[|x1< Ispair x1 y1 << Ispair x2 y2" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans_less 1), - (rtac (monofun_Ispair1 RS monofunE RS spec RS spec RS mp RS - (less_fun RS iffD1 RS spec)) 1), - (rtac (monofun_Ispair2 RS monofunE RS spec RS spec RS mp) 2), - (atac 1), - (atac 1) - ]); - +Goal "[|x1< Ispair x1 y1 << Ispair x2 y2"; +by (rtac trans_less 1); +by (rtac (monofun_Ispair1 RS monofunE RS spec RS spec RS mp RS + (less_fun RS iffD1 RS spec)) 1); +by (rtac (monofun_Ispair2 RS monofunE RS spec RS spec RS mp) 2); +by (atac 1); +by (atac 1); +qed "monofun_Ispair"; (* ------------------------------------------------------------------------ *) (* Isfst and Issnd are monotone *) (* ------------------------------------------------------------------------ *) -qed_goalw "monofun_Isfst" Sprod2.thy [monofun] "monofun(Isfst)" -(fn prems => [(simp_tac (HOL_ss addsimps [inst_sprod_po]) 1)]); +val prems = goalw Sprod2.thy [monofun] "monofun(Isfst)"; +by (simp_tac (HOL_ss addsimps [inst_sprod_po]) 1); +qed "monofun_Isfst"; -qed_goalw "monofun_Issnd" Sprod2.thy [monofun] "monofun(Issnd)" -(fn prems => [(simp_tac (HOL_ss addsimps [inst_sprod_po]) 1)]); +val prems = goalw Sprod2.thy [monofun] "monofun(Issnd)"; +by (simp_tac (HOL_ss addsimps [inst_sprod_po]) 1); +qed "monofun_Issnd"; (* ------------------------------------------------------------------------ *) (* the type 'a ** 'b is a cpo *) (* ------------------------------------------------------------------------ *) -qed_goal "lub_sprod" Sprod2.thy +val prems = goal Sprod2.thy "[|chain(S)|] ==> range(S) <<| \ -\ Ispair (lub(range(%i. Isfst(S i)))) (lub(range(%i. Issnd(S i))))" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac (conjI RS is_lubI) 1), - (rtac (allI RS ub_rangeI) 1), - (res_inst_tac [("t","S(i)")] (surjective_pairing_Sprod RS ssubst) 1), - (rtac monofun_Ispair 1), - (rtac is_ub_thelub 1), - (etac (monofun_Isfst RS ch2ch_monofun) 1), - (rtac is_ub_thelub 1), - (etac (monofun_Issnd RS ch2ch_monofun) 1), - (strip_tac 1), - (res_inst_tac [("t","u")] (surjective_pairing_Sprod RS ssubst) 1), - (rtac monofun_Ispair 1), - (rtac is_lub_thelub 1), - (etac (monofun_Isfst RS ch2ch_monofun) 1), - (etac (monofun_Isfst RS ub2ub_monofun) 1), - (rtac is_lub_thelub 1), - (etac (monofun_Issnd RS ch2ch_monofun) 1), - (etac (monofun_Issnd RS ub2ub_monofun) 1) - ]); +\ Ispair (lub(range(%i. Isfst(S i)))) (lub(range(%i. Issnd(S i))))"; +by (cut_facts_tac prems 1); +by (rtac (conjI RS is_lubI) 1); +by (rtac (allI RS ub_rangeI) 1); +by (res_inst_tac [("t","S(i)")] (surjective_pairing_Sprod RS ssubst) 1); +by (rtac monofun_Ispair 1); +by (rtac is_ub_thelub 1); +by (etac (monofun_Isfst RS ch2ch_monofun) 1); +by (rtac is_ub_thelub 1); +by (etac (monofun_Issnd RS ch2ch_monofun) 1); +by (strip_tac 1); +by (res_inst_tac [("t","u")] (surjective_pairing_Sprod RS ssubst) 1); +by (rtac monofun_Ispair 1); +by (rtac is_lub_thelub 1); +by (etac (monofun_Isfst RS ch2ch_monofun) 1); +by (etac (monofun_Isfst RS ub2ub_monofun) 1); +by (rtac is_lub_thelub 1); +by (etac (monofun_Issnd RS ch2ch_monofun) 1); +by (etac (monofun_Issnd RS ub2ub_monofun) 1); +qed "lub_sprod"; bind_thm ("thelub_sprod", lub_sprod RS thelubI); -qed_goal "cpo_sprod" Sprod2.thy - "chain(S::nat=>'a**'b)==>? x. range(S)<<| x" -(fn prems => - [ - (cut_facts_tac prems 1), - (rtac exI 1), - (etac lub_sprod 1) - ]); - - - - - - - - +val prems = goal Sprod2.thy + "chain(S::nat=>'a**'b)==>? x. range(S)<<| x"; +by (cut_facts_tac prems 1); +by (rtac exI 1); +by (etac lub_sprod 1); +qed "cpo_sprod"; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Sprod3.ML --- a/src/HOLCF/Sprod3.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Sprod3.ML Tue Jul 04 15:58:11 2000 +0200 @@ -1,581 +1,488 @@ -(* Title: HOLCF/Sprod3.thy +(* Title: HOLCF/Sprod3 ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for Sprod.thy +Class instance of ** for class pcpo *) -open Sprod3; - (* for compatibility with old HOLCF-Version *) -qed_goal "inst_sprod_pcpo" thy "UU = Ispair UU UU" - (fn prems => - [ - (simp_tac (HOL_ss addsimps [UU_def,UU_sprod_def]) 1) - ]); +val prems = goal thy "UU = Ispair UU UU"; +by (simp_tac (HOL_ss addsimps [UU_def,UU_sprod_def]) 1); +qed "inst_sprod_pcpo"; + +Addsimps [inst_sprod_pcpo RS sym]; + (* ------------------------------------------------------------------------ *) (* continuity of Ispair, Isfst, Issnd *) (* ------------------------------------------------------------------------ *) -qed_goal "sprod3_lemma1" thy +val prems = goal thy "[| chain(Y); x~= UU; lub(range(Y))~= UU |] ==>\ \ Ispair (lub(range Y)) x =\ \ Ispair (lub(range(%i. Isfst(Ispair(Y i) x)))) \ -\ (lub(range(%i. Issnd(Ispair(Y i) x))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("f1","Ispair")] (arg_cong RS cong) 1), - (rtac lub_equal 1), - (atac 1), - (rtac (monofun_Isfst RS ch2ch_monofun) 1), - (rtac ch2ch_fun 1), - (rtac (monofun_Ispair1 RS ch2ch_monofun) 1), - (atac 1), - (rtac allI 1), - (asm_simp_tac Sprod0_ss 1), - (rtac sym 1), - (rtac lub_chain_maxelem 1), - (res_inst_tac [("P","%j.~Y(j)=UU")] exE 1), - (rtac (not_all RS iffD1) 1), - (res_inst_tac [("Q","lub(range(Y)) = UU")] contrapos 1), - (atac 1), - (rtac chain_UU_I_inverse 1), - (atac 1), - (rtac exI 1), - (etac Issnd2 1), - (rtac allI 1), - (res_inst_tac [("Q","Y(i)=UU")] (excluded_middle RS disjE) 1), - (asm_simp_tac Sprod0_ss 1), - (rtac refl_less 1), - (res_inst_tac [("s","UU"),("t","Y(i)")] subst 1), - (etac sym 1), - (asm_simp_tac Sprod0_ss 1), - (rtac minimal 1) - ]); +\ (lub(range(%i. Issnd(Ispair(Y i) x))))"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("f1","Ispair")] (arg_cong RS cong) 1); +by (rtac lub_equal 1); +by (atac 1); +by (rtac (monofun_Isfst RS ch2ch_monofun) 1); +by (rtac ch2ch_fun 1); +by (rtac (monofun_Ispair1 RS ch2ch_monofun) 1); +by (atac 1); +by (rtac allI 1); +by (asm_simp_tac Sprod0_ss 1); +by (rtac sym 1); +by (rtac lub_chain_maxelem 1); +by (res_inst_tac [("P","%j.~Y(j)=UU")] exE 1); +by (rtac (not_all RS iffD1) 1); +by (res_inst_tac [("Q","lub(range(Y)) = UU")] contrapos 1); +by (atac 1); +by (rtac chain_UU_I_inverse 1); +by (atac 1); +by (rtac exI 1); +by (etac Issnd2 1); +by (rtac allI 1); +by (res_inst_tac [("Q","Y(i)=UU")] (excluded_middle RS disjE) 1); +by (asm_simp_tac Sprod0_ss 1); +by (rtac refl_less 1); +by (res_inst_tac [("s","UU"),("t","Y(i)")] subst 1); +by (etac sym 1); +by (asm_simp_tac Sprod0_ss 1); +by (rtac minimal 1); +qed "sprod3_lemma1"; -qed_goal "sprod3_lemma2" thy +val prems = goal thy "[| chain(Y); x ~= UU; lub(range(Y)) = UU |] ==>\ \ Ispair (lub(range Y)) x =\ \ Ispair (lub(range(%i. Isfst(Ispair(Y i) x))))\ -\ (lub(range(%i. Issnd(Ispair(Y i) x))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1), - (atac 1), - (rtac trans 1), - (rtac strict_Ispair1 1), - (rtac (strict_Ispair RS sym) 1), - (rtac disjI1 1), - (rtac chain_UU_I_inverse 1), - (rtac allI 1), - (asm_simp_tac Sprod0_ss 1), - (etac (chain_UU_I RS spec) 1), - (atac 1) - ]); +\ (lub(range(%i. Issnd(Ispair(Y i) x))))"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1); +by (atac 1); +by (rtac trans 1); +by (rtac strict_Ispair1 1); +by (rtac (strict_Ispair RS sym) 1); +by (rtac disjI1 1); +by (rtac chain_UU_I_inverse 1); +by (rtac allI 1); +by (asm_simp_tac Sprod0_ss 1); +by (etac (chain_UU_I RS spec) 1); +by (atac 1); +qed "sprod3_lemma2"; -qed_goal "sprod3_lemma3" thy +val prems = goal thy "[| chain(Y); x = UU |] ==>\ \ Ispair (lub(range Y)) x =\ \ Ispair (lub(range(%i. Isfst(Ispair (Y i) x))))\ -\ (lub(range(%i. Issnd(Ispair (Y i) x))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("s","UU"),("t","x")] ssubst 1), - (atac 1), - (rtac trans 1), - (rtac strict_Ispair2 1), - (rtac (strict_Ispair RS sym) 1), - (rtac disjI1 1), - (rtac chain_UU_I_inverse 1), - (rtac allI 1), - (simp_tac Sprod0_ss 1) - ]); - +\ (lub(range(%i. Issnd(Ispair (Y i) x))))"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("s","UU"),("t","x")] ssubst 1); +by (atac 1); +by (rtac trans 1); +by (rtac strict_Ispair2 1); +by (rtac (strict_Ispair RS sym) 1); +by (rtac disjI1 1); +by (rtac chain_UU_I_inverse 1); +by (rtac allI 1); +by (simp_tac Sprod0_ss 1); +qed "sprod3_lemma3"; -qed_goal "contlub_Ispair1" thy "contlub(Ispair)" -(fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (expand_fun_eq RS iffD2) 1), - (strip_tac 1), - (stac (lub_fun RS thelubI) 1), - (etac (monofun_Ispair1 RS ch2ch_monofun) 1), - (rtac trans 1), - (rtac (thelub_sprod RS sym) 2), - (rtac ch2ch_fun 2), - (etac (monofun_Ispair1 RS ch2ch_monofun) 2), - (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), - (res_inst_tac - [("Q","lub(range(Y))=UU")] (excluded_middle RS disjE) 1), - (etac sprod3_lemma1 1), - (atac 1), - (atac 1), - (etac sprod3_lemma2 1), - (atac 1), - (atac 1), - (etac sprod3_lemma3 1), - (atac 1) - ]); +Goal "contlub(Ispair)"; +by (rtac contlubI 1); +by (strip_tac 1); +by (rtac (expand_fun_eq RS iffD2) 1); +by (strip_tac 1); +by (stac (lub_fun RS thelubI) 1); +by (etac (monofun_Ispair1 RS ch2ch_monofun) 1); +by (rtac trans 1); +by (rtac (thelub_sprod RS sym) 2); +by (rtac ch2ch_fun 2); +by (etac (monofun_Ispair1 RS ch2ch_monofun) 2); +by (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1); +by (res_inst_tac [("Q","lub(range(Y))=UU")] (excluded_middle RS disjE) 1); +by (etac sprod3_lemma1 1); +by (atac 1); +by (atac 1); +by (etac sprod3_lemma2 1); +by (atac 1); +by (atac 1); +by (etac sprod3_lemma3 1); +by (atac 1); +qed "contlub_Ispair1"; -qed_goal "sprod3_lemma4" thy +val prems = goal thy "[| chain(Y); x ~= UU; lub(range(Y)) ~= UU |] ==>\ \ Ispair x (lub(range Y)) =\ \ Ispair (lub(range(%i. Isfst (Ispair x (Y i)))))\ -\ (lub(range(%i. Issnd (Ispair x (Y i)))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("f1","Ispair")] (arg_cong RS cong) 1), - (rtac sym 1), - (rtac lub_chain_maxelem 1), - (res_inst_tac [("P","%j. Y(j)~=UU")] exE 1), - (rtac (not_all RS iffD1) 1), - (res_inst_tac [("Q","lub(range(Y)) = UU")] contrapos 1), - (atac 1), - (rtac chain_UU_I_inverse 1), - (atac 1), - (rtac exI 1), - (etac Isfst2 1), - (rtac allI 1), - (res_inst_tac [("Q","Y(i)=UU")] (excluded_middle RS disjE) 1), - (asm_simp_tac Sprod0_ss 1), - (rtac refl_less 1), - (res_inst_tac [("s","UU"),("t","Y(i)")] subst 1), - (etac sym 1), - (asm_simp_tac Sprod0_ss 1), - (rtac minimal 1), - (rtac lub_equal 1), - (atac 1), - (rtac (monofun_Issnd RS ch2ch_monofun) 1), - (rtac (monofun_Ispair2 RS ch2ch_monofun) 1), - (atac 1), - (rtac allI 1), - (asm_simp_tac Sprod0_ss 1) - ]); +\ (lub(range(%i. Issnd (Ispair x (Y i)))))"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("f1","Ispair")] (arg_cong RS cong) 1); +by (rtac sym 1); +by (rtac lub_chain_maxelem 1); +by (res_inst_tac [("P","%j. Y(j)~=UU")] exE 1); +by (rtac (not_all RS iffD1) 1); +by (res_inst_tac [("Q","lub(range(Y)) = UU")] contrapos 1); +by (atac 1); +by (rtac chain_UU_I_inverse 1); +by (atac 1); +by (rtac exI 1); +by (etac Isfst2 1); +by (rtac allI 1); +by (res_inst_tac [("Q","Y(i)=UU")] (excluded_middle RS disjE) 1); +by (asm_simp_tac Sprod0_ss 1); +by (rtac refl_less 1); +by (res_inst_tac [("s","UU"),("t","Y(i)")] subst 1); +by (etac sym 1); +by (asm_simp_tac Sprod0_ss 1); +by (rtac minimal 1); +by (rtac lub_equal 1); +by (atac 1); +by (rtac (monofun_Issnd RS ch2ch_monofun) 1); +by (rtac (monofun_Ispair2 RS ch2ch_monofun) 1); +by (atac 1); +by (rtac allI 1); +by (asm_simp_tac Sprod0_ss 1); +qed "sprod3_lemma4"; -qed_goal "sprod3_lemma5" thy +val prems = goal thy "[| chain(Y); x ~= UU; lub(range(Y)) = UU |] ==>\ \ Ispair x (lub(range Y)) =\ \ Ispair (lub(range(%i. Isfst(Ispair x (Y i)))))\ -\ (lub(range(%i. Issnd(Ispair x (Y i)))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1), - (atac 1), - (rtac trans 1), - (rtac strict_Ispair2 1), - (rtac (strict_Ispair RS sym) 1), - (rtac disjI2 1), - (rtac chain_UU_I_inverse 1), - (rtac allI 1), - (asm_simp_tac Sprod0_ss 1), - (etac (chain_UU_I RS spec) 1), - (atac 1) - ]); +\ (lub(range(%i. Issnd(Ispair x (Y i)))))"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1); +by (atac 1); +by (rtac trans 1); +by (rtac strict_Ispair2 1); +by (rtac (strict_Ispair RS sym) 1); +by (rtac disjI2 1); +by (rtac chain_UU_I_inverse 1); +by (rtac allI 1); +by (asm_simp_tac Sprod0_ss 1); +by (etac (chain_UU_I RS spec) 1); +by (atac 1); +qed "sprod3_lemma5"; -qed_goal "sprod3_lemma6" thy +val prems = goal thy "[| chain(Y); x = UU |] ==>\ \ Ispair x (lub(range Y)) =\ \ Ispair (lub(range(%i. Isfst (Ispair x (Y i)))))\ -\ (lub(range(%i. Issnd (Ispair x (Y i)))))" -(fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("s","UU"),("t","x")] ssubst 1), - (atac 1), - (rtac trans 1), - (rtac strict_Ispair1 1), - (rtac (strict_Ispair RS sym) 1), - (rtac disjI1 1), - (rtac chain_UU_I_inverse 1), - (rtac allI 1), - (simp_tac Sprod0_ss 1) - ]); +\ (lub(range(%i. Issnd (Ispair x (Y i)))))"; +by (cut_facts_tac prems 1); +by (res_inst_tac [("s","UU"),("t","x")] ssubst 1); +by (atac 1); +by (rtac trans 1); +by (rtac strict_Ispair1 1); +by (rtac (strict_Ispair RS sym) 1); +by (rtac disjI1 1); +by (rtac chain_UU_I_inverse 1); +by (rtac allI 1); +by (simp_tac Sprod0_ss 1); +qed "sprod3_lemma6"; -qed_goal "contlub_Ispair2" thy "contlub(Ispair(x))" -(fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac trans 1), - (rtac (thelub_sprod RS sym) 2), - (etac (monofun_Ispair2 RS ch2ch_monofun) 2), - (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), - (res_inst_tac [("Q","lub(range(Y))=UU")] - (excluded_middle RS disjE) 1), - (etac sprod3_lemma4 1), - (atac 1), - (atac 1), - (etac sprod3_lemma5 1), - (atac 1), - (atac 1), - (etac sprod3_lemma6 1), - (atac 1) - ]); +val prems = goal thy "contlub(Ispair(x))"; +by (rtac contlubI 1); +by (strip_tac 1); +by (rtac trans 1); +by (rtac (thelub_sprod RS sym) 2); +by (etac (monofun_Ispair2 RS ch2ch_monofun) 2); +by (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1); +by (res_inst_tac [("Q","lub(range(Y))=UU")] (excluded_middle RS disjE) 1); +by (etac sprod3_lemma4 1); +by (atac 1); +by (atac 1); +by (etac sprod3_lemma5 1); +by (atac 1); +by (atac 1); +by (etac sprod3_lemma6 1); +by (atac 1); +qed "contlub_Ispair2"; - -qed_goal "cont_Ispair1" thy "cont(Ispair)" -(fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Ispair1 1), - (rtac contlub_Ispair1 1) - ]); +val prems = goal thy "cont(Ispair)"; +by (rtac monocontlub2cont 1); +by (rtac monofun_Ispair1 1); +by (rtac contlub_Ispair1 1); +qed "cont_Ispair1"; -qed_goal "cont_Ispair2" thy "cont(Ispair(x))" -(fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Ispair2 1), - (rtac contlub_Ispair2 1) - ]); +val prems = goal thy "cont(Ispair(x))"; +by (rtac monocontlub2cont 1); +by (rtac monofun_Ispair2 1); +by (rtac contlub_Ispair2 1); +qed "cont_Ispair2"; -qed_goal "contlub_Isfst" thy "contlub(Isfst)" - (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (stac (lub_sprod RS thelubI) 1), - (atac 1), - (res_inst_tac [("Q","lub(range(%i. Issnd(Y(i))))=UU")] - (excluded_middle RS disjE) 1), - (asm_simp_tac Sprod0_ss 1), - (res_inst_tac [("s","UU"),("t","lub(range(%i. Issnd(Y(i))))")] - ssubst 1), - (atac 1), - (rtac trans 1), - (asm_simp_tac Sprod0_ss 1), - (rtac sym 1), - (rtac chain_UU_I_inverse 1), - (rtac allI 1), - (rtac strict_Isfst 1), - (rtac swap 1), - (etac (defined_IsfstIssnd RS conjunct2) 2), - (fast_tac (HOL_cs addSDs [monofun_Issnd RS ch2ch_monofun RS - chain_UU_I RS spec]) 1) - ]); +val prems = goal thy "contlub(Isfst)"; +by (rtac contlubI 1); +by (strip_tac 1); +by (stac (lub_sprod RS thelubI) 1); +by (atac 1); +by (res_inst_tac [("Q","lub(range(%i. Issnd(Y(i))))=UU")] (excluded_middle RS disjE) 1); +by (asm_simp_tac Sprod0_ss 1); +by (res_inst_tac [("s","UU"),("t","lub(range(%i. Issnd(Y(i))))")] ssubst 1); +by (atac 1); +by (rtac trans 1); +by (asm_simp_tac Sprod0_ss 1); +by (rtac sym 1); +by (rtac chain_UU_I_inverse 1); +by (rtac allI 1); +by (rtac strict_Isfst 1); +by (rtac swap 1); +by (etac (defined_IsfstIssnd RS conjunct2) 2); +by (fast_tac (HOL_cs addSDs [monofun_Issnd RS ch2ch_monofun RS chain_UU_I RS spec]) 1); +qed "contlub_Isfst"; -qed_goal "contlub_Issnd" thy "contlub(Issnd)" -(fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (stac (lub_sprod RS thelubI) 1), - (atac 1), - (res_inst_tac [("Q","lub(range(%i. Isfst(Y(i))))=UU")] - (excluded_middle RS disjE) 1), - (asm_simp_tac Sprod0_ss 1), - (res_inst_tac [("s","UU"),("t","lub(range(%i. Isfst(Y(i))))")] - ssubst 1), - (atac 1), - (asm_simp_tac Sprod0_ss 1), - (rtac sym 1), - (rtac chain_UU_I_inverse 1), - (rtac allI 1), - (rtac strict_Issnd 1), - (rtac swap 1), - (etac (defined_IsfstIssnd RS conjunct1) 2), - (fast_tac (HOL_cs addSDs [monofun_Isfst RS ch2ch_monofun RS - chain_UU_I RS spec]) 1) - ]); +val prems = goal thy "contlub(Issnd)"; +by (rtac contlubI 1); +by (strip_tac 1); +by (stac (lub_sprod RS thelubI) 1); +by (atac 1); +by (res_inst_tac [("Q","lub(range(%i. Isfst(Y(i))))=UU")] (excluded_middle RS disjE) 1); +by (asm_simp_tac Sprod0_ss 1); +by (res_inst_tac [("s","UU"),("t","lub(range(%i. Isfst(Y(i))))")] ssubst 1); +by (atac 1); +by (asm_simp_tac Sprod0_ss 1); +by (rtac sym 1); +by (rtac chain_UU_I_inverse 1); +by (rtac allI 1); +by (rtac strict_Issnd 1); +by (rtac swap 1); +by (etac (defined_IsfstIssnd RS conjunct1) 2); +by (fast_tac (HOL_cs addSDs [monofun_Isfst RS ch2ch_monofun RS chain_UU_I RS spec]) 1); +qed "contlub_Issnd"; -qed_goal "cont_Isfst" thy "cont(Isfst)" -(fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Isfst 1), - (rtac contlub_Isfst 1) - ]); +val prems = goal thy "cont(Isfst)"; +by (rtac monocontlub2cont 1); +by (rtac monofun_Isfst 1); +by (rtac contlub_Isfst 1); +qed "cont_Isfst"; -qed_goal "cont_Issnd" thy "cont(Issnd)" -(fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Issnd 1), - (rtac contlub_Issnd 1) - ]); +val prems = goal thy "cont(Issnd)"; +by (rtac monocontlub2cont 1); +by (rtac monofun_Issnd 1); +by (rtac contlub_Issnd 1); +qed "cont_Issnd"; -qed_goal "spair_eq" thy "[|x1=x2;y1=y2|] ==> (:x1,y1:) = (:x2,y2:)" - (fn prems => - [ - (cut_facts_tac prems 1), - (fast_tac HOL_cs 1) - ]); +val prems = goal thy "[|x1=x2;y1=y2|] ==> (:x1,y1:) = (:x2,y2:)"; +by (cut_facts_tac prems 1); +by (fast_tac HOL_cs 1); +qed "spair_eq"; (* ------------------------------------------------------------------------ *) (* convert all lemmas to the continuous versions *) (* ------------------------------------------------------------------------ *) -qed_goalw "beta_cfun_sprod" thy [spair_def] - "(LAM x y. Ispair x y)`a`b = Ispair a b" - (fn prems => - [ - (stac beta_cfun 1), - (simp_tac (simpset() addsimps [cont_Ispair2, cont_Ispair1, - cont2cont_CF1L]) 1), - (stac beta_cfun 1), - (rtac cont_Ispair2 1), - (rtac refl 1) - ]); +val prems = goalw thy [spair_def] + "(LAM x y. Ispair x y)`a`b = Ispair a b"; +by (stac beta_cfun 1); +by (simp_tac (simpset() addsimps [cont_Ispair2, cont_Ispair1, cont2cont_CF1L]) 1); +by (stac beta_cfun 1); +by (rtac cont_Ispair2 1); +by (rtac refl 1); +qed "beta_cfun_sprod"; -qed_goalw "inject_spair" thy [spair_def] - "[| aa~=UU ; ba~=UU ; (:a,b:)=(:aa,ba:) |] ==> a=aa & b=ba" - (fn prems => - [ - (cut_facts_tac prems 1), - (etac inject_Ispair 1), - (atac 1), - (etac box_equals 1), - (rtac beta_cfun_sprod 1), - (rtac beta_cfun_sprod 1) - ]); +Addsimps [beta_cfun_sprod]; -qed_goalw "inst_sprod_pcpo2" thy [spair_def] "UU = (:UU,UU:)" - (fn prems => - [ - (rtac sym 1), - (rtac trans 1), - (rtac beta_cfun_sprod 1), - (rtac sym 1), - (rtac inst_sprod_pcpo 1) - ]); +val prems = goalw thy [spair_def] + "[| aa~=UU ; ba~=UU ; (:a,b:)=(:aa,ba:) |] ==> a=aa & b=ba"; +by (cut_facts_tac prems 1); +by (etac inject_Ispair 1); +by (atac 1); +by (etac box_equals 1); +by (rtac beta_cfun_sprod 1); +by (rtac beta_cfun_sprod 1); +qed "inject_spair"; -qed_goalw "strict_spair" thy [spair_def] - "(a=UU | b=UU) ==> (:a,b:)=UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans 1), - (rtac beta_cfun_sprod 1), - (rtac trans 1), - (rtac (inst_sprod_pcpo RS sym) 2), - (etac strict_Ispair 1) - ]); +val prems = goalw thy [spair_def] "UU = (:UU,UU:)"; +by (rtac sym 1); +by (rtac trans 1); +by (rtac beta_cfun_sprod 1); +by (rtac sym 1); +by (rtac inst_sprod_pcpo 1); +qed "inst_sprod_pcpo2"; -qed_goalw "strict_spair1" thy [spair_def] "(:UU,b:) = UU" - (fn prems => - [ - (stac beta_cfun_sprod 1), - (rtac trans 1), - (rtac (inst_sprod_pcpo RS sym) 2), - (rtac strict_Ispair1 1) - ]); +val prems = goalw thy [spair_def] + "(a=UU | b=UU) ==> (:a,b:)=UU"; +by (cut_facts_tac prems 1); +by (rtac trans 1); +by (rtac beta_cfun_sprod 1); +by (rtac trans 1); +by (rtac (inst_sprod_pcpo RS sym) 2); +by (etac strict_Ispair 1); +qed "strict_spair"; -qed_goalw "strict_spair2" thy [spair_def] "(:a,UU:) = UU" - (fn prems => - [ - (stac beta_cfun_sprod 1), - (rtac trans 1), - (rtac (inst_sprod_pcpo RS sym) 2), - (rtac strict_Ispair2 1) - ]); +val prems = goalw thy [spair_def] "(:UU,b:) = UU"; +by (stac beta_cfun_sprod 1); +by (rtac trans 1); +by (rtac (inst_sprod_pcpo RS sym) 2); +by (rtac strict_Ispair1 1); +qed "strict_spair1"; -qed_goalw "strict_spair_rev" thy [spair_def] - "(:x,y:)~=UU ==> ~x=UU & ~y=UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac strict_Ispair_rev 1), - (rtac (beta_cfun_sprod RS subst) 1), - (rtac (inst_sprod_pcpo RS subst) 1), - (atac 1) - ]); +val prems = goalw thy [spair_def] "(:a,UU:) = UU"; +by (stac beta_cfun_sprod 1); +by (rtac trans 1); +by (rtac (inst_sprod_pcpo RS sym) 2); +by (rtac strict_Ispair2 1); +qed "strict_spair2"; + +Addsimps [strict_spair1,strict_spair2]; + +Goalw [spair_def] "(:x,y:)~=UU ==> ~x=UU & ~y=UU"; +by (rtac strict_Ispair_rev 1); +by Auto_tac; +qed "strict_spair_rev"; + +Goalw [spair_def] "(:a,b:) = UU ==> (a = UU | b = UU)"; +by (rtac defined_Ispair_rev 1); +by Auto_tac; +qed "defined_spair_rev"; -qed_goalw "defined_spair_rev" thy [spair_def] - "(:a,b:) = UU ==> (a = UU | b = UU)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac defined_Ispair_rev 1), - (rtac (beta_cfun_sprod RS subst) 1), - (rtac (inst_sprod_pcpo RS subst) 1), - (atac 1) - ]); +val prems = goalw thy [spair_def] + "[|a~=UU; b~=UU|] ==> (:a,b:) ~= UU"; +by (cut_facts_tac prems 1); +by (stac beta_cfun_sprod 1); +by (stac inst_sprod_pcpo 1); +by (etac defined_Ispair 1); +by (atac 1); +qed "defined_spair"; -qed_goalw "defined_spair" thy [spair_def] - "[|a~=UU; b~=UU|] ==> (:a,b:) ~= UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun_sprod 1), - (stac inst_sprod_pcpo 1), - (etac defined_Ispair 1), - (atac 1) - ]); - -qed_goalw "Exh_Sprod2" thy [spair_def] - "z=UU | (? a b. z=(:a,b:) & a~=UU & b~=UU)" - (fn prems => - [ - (rtac (Exh_Sprod RS disjE) 1), - (rtac disjI1 1), - (stac inst_sprod_pcpo 1), - (atac 1), - (rtac disjI2 1), - (etac exE 1), - (etac exE 1), - (rtac exI 1), - (rtac exI 1), - (rtac conjI 1), - (stac beta_cfun_sprod 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); +val prems = goalw thy [spair_def] + "z=UU | (? a b. z=(:a,b:) & a~=UU & b~=UU)"; +by (rtac (Exh_Sprod RS disjE) 1); +by (rtac disjI1 1); +by (stac inst_sprod_pcpo 1); +by (atac 1); +by (rtac disjI2 1); +by (etac exE 1); +by (etac exE 1); +by (rtac exI 1); +by (rtac exI 1); +by (rtac conjI 1); +by (stac beta_cfun_sprod 1); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +qed "Exh_Sprod2"; -qed_goalw "sprodE" thy [spair_def] -"[|p=UU ==> Q;!!x y. [|p=(:x,y:);x~=UU ; y~=UU|] ==> Q|] ==> Q" -(fn prems => - [ - (rtac IsprodE 1), - (resolve_tac prems 1), - (stac inst_sprod_pcpo 1), - (atac 1), - (resolve_tac prems 1), - (atac 2), - (atac 2), - (stac beta_cfun_sprod 1), - (atac 1) - ]); +val [prem1,prem2] = Goalw [spair_def] + "[|p=UU ==> Q; !!x y. [| p=(:x,y:); x~=UU; y~=UU|] ==> Q|] ==> Q"; +by (rtac IsprodE 1); +by (rtac prem1 1); +by (stac inst_sprod_pcpo 1); +by (atac 1); +by (rtac prem2 1); +by (atac 2); +by (atac 2); +by (stac beta_cfun_sprod 1); +by (atac 1); +qed "sprodE"; -qed_goalw "strict_sfst" thy [sfst_def] - "p=UU==>sfst`p=UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun 1), - (rtac cont_Isfst 1), - (rtac strict_Isfst 1), - (rtac (inst_sprod_pcpo RS subst) 1), - (atac 1) - ]); +val prems = goalw thy [sfst_def] + "p=UU==>sfst`p=UU"; +by (cut_facts_tac prems 1); +by (stac beta_cfun 1); +by (rtac cont_Isfst 1); +by (rtac strict_Isfst 1); +by (rtac (inst_sprod_pcpo RS subst) 1); +by (atac 1); +qed "strict_sfst"; -qed_goalw "strict_sfst1" thy [sfst_def,spair_def] - "sfst`(:UU,y:) = UU" - (fn prems => - [ - (stac beta_cfun_sprod 1), - (stac beta_cfun 1), - (rtac cont_Isfst 1), - (rtac strict_Isfst1 1) - ]); +val prems = goalw thy [sfst_def,spair_def] + "sfst`(:UU,y:) = UU"; +by (stac beta_cfun_sprod 1); +by (stac beta_cfun 1); +by (rtac cont_Isfst 1); +by (rtac strict_Isfst1 1); +qed "strict_sfst1"; -qed_goalw "strict_sfst2" thy [sfst_def,spair_def] - "sfst`(:x,UU:) = UU" - (fn prems => - [ - (stac beta_cfun_sprod 1), - (stac beta_cfun 1), - (rtac cont_Isfst 1), - (rtac strict_Isfst2 1) - ]); +val prems = goalw thy [sfst_def,spair_def] + "sfst`(:x,UU:) = UU"; +by (stac beta_cfun_sprod 1); +by (stac beta_cfun 1); +by (rtac cont_Isfst 1); +by (rtac strict_Isfst2 1); +qed "strict_sfst2"; -qed_goalw "strict_ssnd" thy [ssnd_def] - "p=UU==>ssnd`p=UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun 1), - (rtac cont_Issnd 1), - (rtac strict_Issnd 1), - (rtac (inst_sprod_pcpo RS subst) 1), - (atac 1) - ]); +val prems = goalw thy [ssnd_def] + "p=UU==>ssnd`p=UU"; +by (cut_facts_tac prems 1); +by (stac beta_cfun 1); +by (rtac cont_Issnd 1); +by (rtac strict_Issnd 1); +by (rtac (inst_sprod_pcpo RS subst) 1); +by (atac 1); +qed "strict_ssnd"; -qed_goalw "strict_ssnd1" thy [ssnd_def,spair_def] - "ssnd`(:UU,y:) = UU" - (fn prems => - [ - (stac beta_cfun_sprod 1), - (stac beta_cfun 1), - (rtac cont_Issnd 1), - (rtac strict_Issnd1 1) - ]); +val prems = goalw thy [ssnd_def,spair_def] + "ssnd`(:UU,y:) = UU"; +by (stac beta_cfun_sprod 1); +by (stac beta_cfun 1); +by (rtac cont_Issnd 1); +by (rtac strict_Issnd1 1); +qed "strict_ssnd1"; -qed_goalw "strict_ssnd2" thy [ssnd_def,spair_def] - "ssnd`(:x,UU:) = UU" - (fn prems => - [ - (stac beta_cfun_sprod 1), - (stac beta_cfun 1), - (rtac cont_Issnd 1), - (rtac strict_Issnd2 1) - ]); +val prems = goalw thy [ssnd_def,spair_def] + "ssnd`(:x,UU:) = UU"; +by (stac beta_cfun_sprod 1); +by (stac beta_cfun 1); +by (rtac cont_Issnd 1); +by (rtac strict_Issnd2 1); +qed "strict_ssnd2"; -qed_goalw "sfst2" thy [sfst_def,spair_def] - "y~=UU ==>sfst`(:x,y:)=x" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun_sprod 1), - (stac beta_cfun 1), - (rtac cont_Isfst 1), - (etac Isfst2 1) - ]); +val prems = goalw thy [sfst_def,spair_def] + "y~=UU ==>sfst`(:x,y:)=x"; +by (cut_facts_tac prems 1); +by (stac beta_cfun_sprod 1); +by (stac beta_cfun 1); +by (rtac cont_Isfst 1); +by (etac Isfst2 1); +qed "sfst2"; -qed_goalw "ssnd2" thy [ssnd_def,spair_def] - "x~=UU ==>ssnd`(:x,y:)=y" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun_sprod 1), - (stac beta_cfun 1), - (rtac cont_Issnd 1), - (etac Issnd2 1) - ]); +val prems = goalw thy [ssnd_def,spair_def] + "x~=UU ==>ssnd`(:x,y:)=y"; +by (cut_facts_tac prems 1); +by (stac beta_cfun_sprod 1); +by (stac beta_cfun 1); +by (rtac cont_Issnd 1); +by (etac Issnd2 1); +qed "ssnd2"; -qed_goalw "defined_sfstssnd" thy [sfst_def,ssnd_def,spair_def] - "p~=UU ==> sfst`p ~=UU & ssnd`p ~=UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun 1), - (rtac cont_Issnd 1), - (stac beta_cfun 1), - (rtac cont_Isfst 1), - (rtac defined_IsfstIssnd 1), - (rtac (inst_sprod_pcpo RS subst) 1), - (atac 1) - ]); +val prems = goalw thy [sfst_def,ssnd_def,spair_def] + "p~=UU ==> sfst`p ~=UU & ssnd`p ~=UU"; +by (cut_facts_tac prems 1); +by (stac beta_cfun 1); +by (rtac cont_Issnd 1); +by (stac beta_cfun 1); +by (rtac cont_Isfst 1); +by (rtac defined_IsfstIssnd 1); +by (rtac (inst_sprod_pcpo RS subst) 1); +by (atac 1); +qed "defined_sfstssnd"; +val prems = goalw thy [sfst_def,ssnd_def,spair_def] "(:sfst`p , ssnd`p:) = p"; +by (stac beta_cfun_sprod 1); +by (stac beta_cfun 1); +by (rtac cont_Issnd 1); +by (stac beta_cfun 1); +by (rtac cont_Isfst 1); +by (rtac (surjective_pairing_Sprod RS sym) 1); +qed "surjective_pairing_Sprod2"; -qed_goalw "surjective_pairing_Sprod2" thy - [sfst_def,ssnd_def,spair_def] "(:sfst`p , ssnd`p:) = p" - (fn prems => - [ - (stac beta_cfun_sprod 1), - (stac beta_cfun 1), - (rtac cont_Issnd 1), - (stac beta_cfun 1), - (rtac cont_Isfst 1), - (rtac (surjective_pairing_Sprod RS sym) 1) - ]); - - -qed_goalw "lub_sprod2" thy [sfst_def,ssnd_def,spair_def] +val prems = goalw thy [sfst_def,ssnd_def,spair_def] "[|chain(S)|] ==> range(S) <<| \ -\ (: lub(range(%i. sfst`(S i))), lub(range(%i. ssnd`(S i))) :)" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun_sprod 1), - (stac (beta_cfun RS ext) 1), - (rtac cont_Issnd 1), - (stac (beta_cfun RS ext) 1), - (rtac cont_Isfst 1), - (rtac lub_sprod 1), - (resolve_tac prems 1) - ]); +\ (: lub(range(%i. sfst`(S i))), lub(range(%i. ssnd`(S i))) :)"; +by (cut_facts_tac prems 1); +by (stac beta_cfun_sprod 1); +by (stac (beta_cfun RS ext) 1); +by (rtac cont_Issnd 1); +by (stac (beta_cfun RS ext) 1); +by (rtac cont_Isfst 1); +by (rtac lub_sprod 1); +by (resolve_tac prems 1); +qed "lub_sprod2"; bind_thm ("thelub_sprod2", lub_sprod2 RS thelubI); @@ -585,58 +492,52 @@ (:lub (range (%i. sfst`(?S1 i))), lub (range (%i. ssnd`(?S1 i))):)" : thm *) -qed_goalw "ssplit1" thy [ssplit_def] - "ssplit`f`UU=UU" - (fn prems => - [ - (stac beta_cfun 1), - (Simp_tac 1), - (stac strictify1 1), - (rtac refl 1) - ]); +val prems = goalw thy [ssplit_def] + "ssplit`f`UU=UU"; +by (stac beta_cfun 1); +by (Simp_tac 1); +by (stac strictify1 1); +by (rtac refl 1); +qed "ssplit1"; -qed_goalw "ssplit2" thy [ssplit_def] - "[|x~=UU;y~=UU|] ==> ssplit`f`(:x,y:)= f`x`y" - (fn prems => - [ - (stac beta_cfun 1), - (Simp_tac 1), - (stac strictify2 1), - (rtac defined_spair 1), - (resolve_tac prems 1), - (resolve_tac prems 1), - (stac beta_cfun 1), - (Simp_tac 1), - (stac sfst2 1), - (resolve_tac prems 1), - (stac ssnd2 1), - (resolve_tac prems 1), - (rtac refl 1) - ]); +val prems = goalw thy [ssplit_def] + "[|x~=UU;y~=UU|] ==> ssplit`f`(:x,y:)= f`x`y"; +by (stac beta_cfun 1); +by (Simp_tac 1); +by (stac strictify2 1); +by (rtac defined_spair 1); +by (resolve_tac prems 1); +by (resolve_tac prems 1); +by (stac beta_cfun 1); +by (Simp_tac 1); +by (stac sfst2 1); +by (resolve_tac prems 1); +by (stac ssnd2 1); +by (resolve_tac prems 1); +by (rtac refl 1); +qed "ssplit2"; -qed_goalw "ssplit3" thy [ssplit_def] - "ssplit`spair`z=z" - (fn prems => - [ - (stac beta_cfun 1), - (Simp_tac 1), - (case_tac "z=UU" 1), - (hyp_subst_tac 1), - (rtac strictify1 1), - (rtac trans 1), - (rtac strictify2 1), - (atac 1), - (stac beta_cfun 1), - (Simp_tac 1), - (rtac surjective_pairing_Sprod2 1) - ]); +val prems = goalw thy [ssplit_def] + "ssplit`spair`z=z"; +by (stac beta_cfun 1); +by (Simp_tac 1); +by (case_tac "z=UU" 1); +by (hyp_subst_tac 1); +by (rtac strictify1 1); +by (rtac trans 1); +by (rtac strictify2 1); +by (atac 1); +by (stac beta_cfun 1); +by (Simp_tac 1); +by (rtac surjective_pairing_Sprod2 1); +qed "ssplit3"; (* ------------------------------------------------------------------------ *) (* install simplifier for Sprod *) (* ------------------------------------------------------------------------ *) -val Sprod_rews = [strict_spair1,strict_spair2,strict_sfst1,strict_sfst2, +val Sprod_rews = [strict_sfst1,strict_sfst2, strict_ssnd1,strict_ssnd2,sfst2,ssnd2,defined_spair, ssplit1,ssplit2]; Addsimps Sprod_rews; diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Ssum0.ML --- a/src/HOLCF/Ssum0.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Ssum0.ML Tue Jul 04 15:58:11 2000 +0200 @@ -10,23 +10,19 @@ (* A non-emptyness result for Sssum *) (* ------------------------------------------------------------------------ *) -qed_goalw "SsumIl" Ssum0.thy [Ssum_def] "Sinl_Rep(a):Ssum" - (fn prems => - [ - (rtac CollectI 1), - (rtac disjI1 1), - (rtac exI 1), - (rtac refl 1) - ]); +val prems = goalw Ssum0.thy [Ssum_def] "Sinl_Rep(a):Ssum"; +by (rtac CollectI 1); +by (rtac disjI1 1); +by (rtac exI 1); +by (rtac refl 1); +qed "SsumIl"; -qed_goalw "SsumIr" Ssum0.thy [Ssum_def] "Sinr_Rep(a):Ssum" - (fn prems => - [ - (rtac CollectI 1), - (rtac disjI2 1), - (rtac exI 1), - (rtac refl 1) - ]); +val prems = goalw Ssum0.thy [Ssum_def] "Sinr_Rep(a):Ssum"; +by (rtac CollectI 1); +by (rtac disjI2 1); +by (rtac exI 1); +by (rtac refl 1); +qed "SsumIr"; Goal "inj_on Abs_Ssum Ssum"; by (rtac inj_on_inverseI 1); @@ -37,58 +33,48 @@ (* Strictness of Sinr_Rep, Sinl_Rep and Isinl, Isinr *) (* ------------------------------------------------------------------------ *) -qed_goalw "strict_SinlSinr_Rep" Ssum0.thy [Sinr_Rep_def,Sinl_Rep_def] - "Sinl_Rep(UU) = Sinr_Rep(UU)" - (fn prems => - [ - (rtac ext 1), - (rtac ext 1), - (rtac ext 1), - (fast_tac HOL_cs 1) - ]); +val prems = goalw Ssum0.thy [Sinr_Rep_def,Sinl_Rep_def] + "Sinl_Rep(UU) = Sinr_Rep(UU)"; +by (rtac ext 1); +by (rtac ext 1); +by (rtac ext 1); +by (fast_tac HOL_cs 1); +qed "strict_SinlSinr_Rep"; -qed_goalw "strict_IsinlIsinr" Ssum0.thy [Isinl_def,Isinr_def] - "Isinl(UU) = Isinr(UU)" - (fn prems => - [ - (rtac (strict_SinlSinr_Rep RS arg_cong) 1) - ]); +val prems = goalw Ssum0.thy [Isinl_def,Isinr_def] + "Isinl(UU) = Isinr(UU)"; +by (rtac (strict_SinlSinr_Rep RS arg_cong) 1); +qed "strict_IsinlIsinr"; (* ------------------------------------------------------------------------ *) (* distinctness of Sinl_Rep, Sinr_Rep and Isinl, Isinr *) (* ------------------------------------------------------------------------ *) -qed_goalw "noteq_SinlSinr_Rep" Ssum0.thy [Sinl_Rep_def,Sinr_Rep_def] - "(Sinl_Rep(a) = Sinr_Rep(b)) ==> a=UU & b=UU" - (fn prems => - [ - (rtac conjI 1), - (case_tac "a=UU" 1), - (atac 1), - (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong RS iffD2 - RS mp RS conjunct1 RS sym) 1), - (fast_tac HOL_cs 1), - (atac 1), - (case_tac "b=UU" 1), - (atac 1), - (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong RS iffD1 - RS mp RS conjunct1 RS sym) 1), - (fast_tac HOL_cs 1), - (atac 1) - ]); +val prems = goalw Ssum0.thy [Sinl_Rep_def,Sinr_Rep_def] + "(Sinl_Rep(a) = Sinr_Rep(b)) ==> a=UU & b=UU"; +by (rtac conjI 1); +by (case_tac "a=UU" 1); +by (atac 1); +by (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong RS iffD2 RS mp RS conjunct1 RS sym) 1); +by (fast_tac HOL_cs 1); +by (atac 1); +by (case_tac "b=UU" 1); +by (atac 1); +by (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong RS iffD1 RS mp RS conjunct1 RS sym) 1); +by (fast_tac HOL_cs 1); +by (atac 1); +qed "noteq_SinlSinr_Rep"; -qed_goalw "noteq_IsinlIsinr" Ssum0.thy [Isinl_def,Isinr_def] - "Isinl(a)=Isinr(b) ==> a=UU & b=UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac noteq_SinlSinr_Rep 1), - (etac (inj_on_Abs_Ssum RS inj_onD) 1), - (rtac SsumIl 1), - (rtac SsumIr 1) - ]); +val prems = goalw Ssum0.thy [Isinl_def,Isinr_def] + "Isinl(a)=Isinr(b) ==> a=UU & b=UU"; +by (cut_facts_tac prems 1); +by (rtac noteq_SinlSinr_Rep 1); +by (etac (inj_on_Abs_Ssum RS inj_onD) 1); +by (rtac SsumIl 1); +by (rtac SsumIr 1); +qed "noteq_IsinlIsinr"; @@ -96,49 +82,37 @@ (* injectivity of Sinl_Rep, Sinr_Rep and Isinl, Isinr *) (* ------------------------------------------------------------------------ *) -qed_goalw "inject_Sinl_Rep1" Ssum0.thy [Sinl_Rep_def] - "(Sinl_Rep(a) = Sinl_Rep(UU)) ==> a=UU" - (fn prems => - [ - (case_tac "a=UU" 1), - (atac 1), - (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong - RS iffD2 RS mp RS conjunct1 RS sym) 1), - (fast_tac HOL_cs 1), - (atac 1) - ]); +val prems = goalw Ssum0.thy [Sinl_Rep_def] + "(Sinl_Rep(a) = Sinl_Rep(UU)) ==> a=UU"; +by (case_tac "a=UU" 1); +by (atac 1); +by (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong RS iffD2 RS mp RS conjunct1 RS sym) 1); +by (fast_tac HOL_cs 1); +by (atac 1); +qed "inject_Sinl_Rep1"; -qed_goalw "inject_Sinr_Rep1" Ssum0.thy [Sinr_Rep_def] - "(Sinr_Rep(b) = Sinr_Rep(UU)) ==> b=UU" - (fn prems => - [ - (case_tac "b=UU" 1), - (atac 1), - (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong - RS iffD2 RS mp RS conjunct1 RS sym) 1), - (fast_tac HOL_cs 1), - (atac 1) - ]); +val prems = goalw Ssum0.thy [Sinr_Rep_def] + "(Sinr_Rep(b) = Sinr_Rep(UU)) ==> b=UU"; +by (case_tac "b=UU" 1); +by (atac 1); +by (rtac ((hd prems) RS fun_cong RS fun_cong RS fun_cong RS iffD2 RS mp RS conjunct1 RS sym) 1); +by (fast_tac HOL_cs 1); +by (atac 1); +qed "inject_Sinr_Rep1"; -qed_goalw "inject_Sinl_Rep2" Ssum0.thy [Sinl_Rep_def] -"[| a1~=UU ; a2~=UU ; Sinl_Rep(a1)=Sinl_Rep(a2) |] ==> a1=a2" - (fn prems => - [ - (rtac ((nth_elem (2,prems)) RS fun_cong RS fun_cong RS fun_cong - RS iffD1 RS mp RS conjunct1) 1), - (fast_tac HOL_cs 1), - (resolve_tac prems 1) - ]); +val prems = goalw Ssum0.thy [Sinl_Rep_def] +"[| a1~=UU ; a2~=UU ; Sinl_Rep(a1)=Sinl_Rep(a2) |] ==> a1=a2"; +by (rtac ((nth_elem (2,prems)) RS fun_cong RS fun_cong RS fun_cong RS iffD1 RS mp RS conjunct1) 1); +by (fast_tac HOL_cs 1); +by (resolve_tac prems 1); +qed "inject_Sinl_Rep2"; -qed_goalw "inject_Sinr_Rep2" Ssum0.thy [Sinr_Rep_def] -"[|b1~=UU ; b2~=UU ; Sinr_Rep(b1)=Sinr_Rep(b2) |] ==> b1=b2" - (fn prems => - [ - (rtac ((nth_elem (2,prems)) RS fun_cong RS fun_cong RS fun_cong - RS iffD1 RS mp RS conjunct1) 1), - (fast_tac HOL_cs 1), - (resolve_tac prems 1) - ]); +val prems = goalw Ssum0.thy [Sinr_Rep_def] +"[|b1~=UU ; b2~=UU ; Sinr_Rep(b1)=Sinr_Rep(b2) |] ==> b1=b2"; +by (rtac ((nth_elem (2,prems)) RS fun_cong RS fun_cong RS fun_cong RS iffD1 RS mp RS conjunct1) 1); +by (fast_tac HOL_cs 1); +by (resolve_tac prems 1); +qed "inject_Sinr_Rep2"; Goal "Sinl_Rep(a1)=Sinl_Rep(a2) ==> a1=a2"; by (case_tac "a1=UU" 1); @@ -166,27 +140,23 @@ by (atac 1); qed "inject_Sinr_Rep"; -qed_goalw "inject_Isinl" Ssum0.thy [Isinl_def] -"Isinl(a1)=Isinl(a2)==> a1=a2" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac inject_Sinl_Rep 1), - (etac (inj_on_Abs_Ssum RS inj_onD) 1), - (rtac SsumIl 1), - (rtac SsumIl 1) - ]); +val prems = goalw Ssum0.thy [Isinl_def] +"Isinl(a1)=Isinl(a2)==> a1=a2"; +by (cut_facts_tac prems 1); +by (rtac inject_Sinl_Rep 1); +by (etac (inj_on_Abs_Ssum RS inj_onD) 1); +by (rtac SsumIl 1); +by (rtac SsumIl 1); +qed "inject_Isinl"; -qed_goalw "inject_Isinr" Ssum0.thy [Isinr_def] -"Isinr(b1)=Isinr(b2) ==> b1=b2" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac inject_Sinr_Rep 1), - (etac (inj_on_Abs_Ssum RS inj_onD) 1), - (rtac SsumIr 1), - (rtac SsumIr 1) - ]); +val prems = goalw Ssum0.thy [Isinr_def] +"Isinr(b1)=Isinr(b2) ==> b1=b2"; +by (cut_facts_tac prems 1); +by (rtac inject_Sinr_Rep 1); +by (etac (inj_on_Abs_Ssum RS inj_onD) 1); +by (rtac SsumIr 1); +by (rtac SsumIr 1); +qed "inject_Isinr"; Goal "a1~=a2 ==> Isinl(a1) ~= Isinl(a2)"; by (rtac contrapos 1); @@ -205,46 +175,44 @@ (* choice of the bottom representation is arbitrary *) (* ------------------------------------------------------------------------ *) -qed_goalw "Exh_Ssum" Ssum0.thy [Isinl_def,Isinr_def] - "z=Isinl(UU) | (? a. z=Isinl(a) & a~=UU) | (? b. z=Isinr(b) & b~=UU)" - (fn prems => - [ - (rtac (rewrite_rule [Ssum_def] Rep_Ssum RS CollectE) 1), - (etac disjE 1), - (etac exE 1), - (case_tac "z= Abs_Ssum(Sinl_Rep(UU))" 1), - (etac disjI1 1), - (rtac disjI2 1), - (rtac disjI1 1), - (rtac exI 1), - (rtac conjI 1), - (rtac (Rep_Ssum_inverse RS sym RS trans) 1), - (etac arg_cong 1), - (res_inst_tac [("Q","Sinl_Rep(a)=Sinl_Rep(UU)")] contrapos 1), - (etac arg_cong 2), - (etac contrapos 1), - (rtac (Rep_Ssum_inverse RS sym RS trans) 1), - (rtac trans 1), - (etac arg_cong 1), - (etac arg_cong 1), - (etac exE 1), - (case_tac "z= Abs_Ssum(Sinl_Rep(UU))" 1), - (etac disjI1 1), - (rtac disjI2 1), - (rtac disjI2 1), - (rtac exI 1), - (rtac conjI 1), - (rtac (Rep_Ssum_inverse RS sym RS trans) 1), - (etac arg_cong 1), - (res_inst_tac [("Q","Sinr_Rep(b)=Sinl_Rep(UU)")] contrapos 1), - (hyp_subst_tac 2), - (rtac (strict_SinlSinr_Rep RS sym) 2), - (etac contrapos 1), - (rtac (Rep_Ssum_inverse RS sym RS trans) 1), - (rtac trans 1), - (etac arg_cong 1), - (etac arg_cong 1) - ]); +val prems = goalw Ssum0.thy [Isinl_def,Isinr_def] + "z=Isinl(UU) | (? a. z=Isinl(a) & a~=UU) | (? b. z=Isinr(b) & b~=UU)"; +by (rtac (rewrite_rule [Ssum_def] Rep_Ssum RS CollectE) 1); +by (etac disjE 1); +by (etac exE 1); +by (case_tac "z= Abs_Ssum(Sinl_Rep(UU))" 1); +by (etac disjI1 1); +by (rtac disjI2 1); +by (rtac disjI1 1); +by (rtac exI 1); +by (rtac conjI 1); +by (rtac (Rep_Ssum_inverse RS sym RS trans) 1); +by (etac arg_cong 1); +by (res_inst_tac [("Q","Sinl_Rep(a)=Sinl_Rep(UU)")] contrapos 1); +by (etac arg_cong 2); +by (etac contrapos 1); +by (rtac (Rep_Ssum_inverse RS sym RS trans) 1); +by (rtac trans 1); +by (etac arg_cong 1); +by (etac arg_cong 1); +by (etac exE 1); +by (case_tac "z= Abs_Ssum(Sinl_Rep(UU))" 1); +by (etac disjI1 1); +by (rtac disjI2 1); +by (rtac disjI2 1); +by (rtac exI 1); +by (rtac conjI 1); +by (rtac (Rep_Ssum_inverse RS sym RS trans) 1); +by (etac arg_cong 1); +by (res_inst_tac [("Q","Sinr_Rep(b)=Sinl_Rep(UU)")] contrapos 1); +by (hyp_subst_tac 2); +by (rtac (strict_SinlSinr_Rep RS sym) 2); +by (etac contrapos 1); +by (rtac (Rep_Ssum_inverse RS sym RS trans) 1); +by (rtac trans 1); +by (etac arg_cong 1); +by (etac arg_cong 1); +qed "Exh_Ssum"; (* ------------------------------------------------------------------------ *) (* elimination rules for the strict sum ++ *) @@ -282,83 +250,77 @@ (* rewrites for Iwhen *) (* ------------------------------------------------------------------------ *) -qed_goalw "Iwhen1" Ssum0.thy [Iwhen_def] - "Iwhen f g (Isinl UU) = UU" - (fn prems => - [ - (rtac select_equality 1), - (rtac conjI 1), - (fast_tac HOL_cs 1), - (rtac conjI 1), - (strip_tac 1), - (res_inst_tac [("P","a=UU")] notE 1), - (fast_tac HOL_cs 1), - (rtac inject_Isinl 1), - (rtac sym 1), - (fast_tac HOL_cs 1), - (strip_tac 1), - (res_inst_tac [("P","b=UU")] notE 1), - (fast_tac HOL_cs 1), - (rtac inject_Isinr 1), - (rtac sym 1), - (rtac (strict_IsinlIsinr RS subst) 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); +val prems = goalw Ssum0.thy [Iwhen_def] + "Iwhen f g (Isinl UU) = UU"; +by (rtac select_equality 1); +by (rtac conjI 1); +by (fast_tac HOL_cs 1); +by (rtac conjI 1); +by (strip_tac 1); +by (res_inst_tac [("P","a=UU")] notE 1); +by (fast_tac HOL_cs 1); +by (rtac inject_Isinl 1); +by (rtac sym 1); +by (fast_tac HOL_cs 1); +by (strip_tac 1); +by (res_inst_tac [("P","b=UU")] notE 1); +by (fast_tac HOL_cs 1); +by (rtac inject_Isinr 1); +by (rtac sym 1); +by (rtac (strict_IsinlIsinr RS subst) 1); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +qed "Iwhen1"; -qed_goalw "Iwhen2" Ssum0.thy [Iwhen_def] - "x~=UU ==> Iwhen f g (Isinl x) = f`x" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac select_equality 1), - (fast_tac HOL_cs 2), - (rtac conjI 1), - (strip_tac 1), - (res_inst_tac [("P","x=UU")] notE 1), - (atac 1), - (rtac inject_Isinl 1), - (atac 1), - (rtac conjI 1), - (strip_tac 1), - (rtac cfun_arg_cong 1), - (rtac inject_Isinl 1), - (fast_tac HOL_cs 1), - (strip_tac 1), - (res_inst_tac [("P","Isinl(x) = Isinr(b)")] notE 1), - (fast_tac HOL_cs 2), - (rtac contrapos 1), - (etac noteq_IsinlIsinr 2), - (fast_tac HOL_cs 1) - ]); +val prems = goalw Ssum0.thy [Iwhen_def] + "x~=UU ==> Iwhen f g (Isinl x) = f`x"; +by (cut_facts_tac prems 1); +by (rtac select_equality 1); +by (fast_tac HOL_cs 2); +by (rtac conjI 1); +by (strip_tac 1); +by (res_inst_tac [("P","x=UU")] notE 1); +by (atac 1); +by (rtac inject_Isinl 1); +by (atac 1); +by (rtac conjI 1); +by (strip_tac 1); +by (rtac cfun_arg_cong 1); +by (rtac inject_Isinl 1); +by (fast_tac HOL_cs 1); +by (strip_tac 1); +by (res_inst_tac [("P","Isinl(x) = Isinr(b)")] notE 1); +by (fast_tac HOL_cs 2); +by (rtac contrapos 1); +by (etac noteq_IsinlIsinr 2); +by (fast_tac HOL_cs 1); +qed "Iwhen2"; -qed_goalw "Iwhen3" Ssum0.thy [Iwhen_def] - "y~=UU ==> Iwhen f g (Isinr y) = g`y" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac select_equality 1), - (fast_tac HOL_cs 2), - (rtac conjI 1), - (strip_tac 1), - (res_inst_tac [("P","y=UU")] notE 1), - (atac 1), - (rtac inject_Isinr 1), - (rtac (strict_IsinlIsinr RS subst) 1), - (atac 1), - (rtac conjI 1), - (strip_tac 1), - (res_inst_tac [("P","Isinr(y) = Isinl(a)")] notE 1), - (fast_tac HOL_cs 2), - (rtac contrapos 1), - (etac (sym RS noteq_IsinlIsinr) 2), - (fast_tac HOL_cs 1), - (strip_tac 1), - (rtac cfun_arg_cong 1), - (rtac inject_Isinr 1), - (fast_tac HOL_cs 1) - ]); +val prems = goalw Ssum0.thy [Iwhen_def] + "y~=UU ==> Iwhen f g (Isinr y) = g`y"; +by (cut_facts_tac prems 1); +by (rtac select_equality 1); +by (fast_tac HOL_cs 2); +by (rtac conjI 1); +by (strip_tac 1); +by (res_inst_tac [("P","y=UU")] notE 1); +by (atac 1); +by (rtac inject_Isinr 1); +by (rtac (strict_IsinlIsinr RS subst) 1); +by (atac 1); +by (rtac conjI 1); +by (strip_tac 1); +by (res_inst_tac [("P","Isinr(y) = Isinl(a)")] notE 1); +by (fast_tac HOL_cs 2); +by (rtac contrapos 1); +by (etac (sym RS noteq_IsinlIsinr) 2); +by (fast_tac HOL_cs 1); +by (strip_tac 1); +by (rtac cfun_arg_cong 1); +by (rtac inject_Isinr 1); +by (fast_tac HOL_cs 1); +qed "Iwhen3"; (* ------------------------------------------------------------------------ *) (* instantiate the simplifier *) diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Ssum1.ML --- a/src/HOLCF/Ssum1.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Ssum1.ML Tue Jul 04 15:58:11 2000 +0200 @@ -6,8 +6,6 @@ Partial ordering for the strict sum ++ *) -local - fun eq_left s1 s2 = ( (res_inst_tac [("s",s1),("t",s2)] (inject_Isinl RS subst) 1) @@ -36,174 +34,163 @@ THEN (atac 2) THEN (etac sym 1)) -in +val prems = goalw thy [less_ssum_def] +"[|s1=Isinl(x::'a); s2=Isinl(y::'a)|] ==> s1 << s2 = (x << y)"; +by (cut_facts_tac prems 1); +by (rtac select_equality 1); +by (dtac conjunct1 2); +by (dtac spec 2); +by (dtac spec 2); +by (etac mp 2); +by (fast_tac HOL_cs 2); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (eq_left "x" "u"); +by (eq_left "y" "xa"); +by (rtac refl 1); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (UU_left "x"); +by (UU_right "v"); +by (Simp_tac 1); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (eq_left "x" "u"); +by (UU_left "y"); +by (rtac iffI 1); +by (etac UU_I 1); +by (res_inst_tac [("s","x"),("t","UU::'a")] subst 1); +by (atac 1); +by (rtac refl_less 1); +by (strip_tac 1); +by (etac conjE 1); +by (UU_left "x"); +by (UU_right "v"); +by (Simp_tac 1); +qed "less_ssum1a"; + -val less_ssum1a = prove_goalw thy [less_ssum_def] -"[|s1=Isinl(x::'a); s2=Isinl(y::'a)|] ==> s1 << s2 = (x << y)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac select_equality 1), - (dtac conjunct1 2), - (dtac spec 2), - (dtac spec 2), - (etac mp 2), - (fast_tac HOL_cs 2), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (eq_left "x" "u"), - (eq_left "y" "xa"), - (rtac refl 1), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (UU_left "x"), - (UU_right "v"), - (Simp_tac 1), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (eq_left "x" "u"), - (UU_left "y"), - (rtac iffI 1), - (etac UU_I 1), - (res_inst_tac [("s","x"),("t","UU::'a")] subst 1), - (atac 1), - (rtac refl_less 1), - (strip_tac 1), - (etac conjE 1), - (UU_left "x"), - (UU_right "v"), - (Simp_tac 1) - ]); +val prems = goalw thy [less_ssum_def] +"[|s1=Isinr(x::'b); s2=Isinr(y::'b)|] ==> s1 << s2 = (x << y)"; +by (cut_facts_tac prems 1); +by (rtac select_equality 1); +by (dtac conjunct2 2); +by (dtac conjunct1 2); +by (dtac spec 2); +by (dtac spec 2); +by (etac mp 2); +by (fast_tac HOL_cs 2); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (UU_right "x"); +by (UU_left "u"); +by (Simp_tac 1); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (eq_right "x" "v"); +by (eq_right "y" "ya"); +by (rtac refl 1); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (UU_right "x"); +by (UU_left "u"); +by (Simp_tac 1); +by (strip_tac 1); +by (etac conjE 1); +by (eq_right "x" "v"); +by (UU_right "y"); +by (rtac iffI 1); +by (etac UU_I 1); +by (res_inst_tac [("s","UU::'b"),("t","x")] subst 1); +by (etac sym 1); +by (rtac refl_less 1); +qed "less_ssum1b"; -val less_ssum1b = prove_goalw thy [less_ssum_def] -"[|s1=Isinr(x::'b); s2=Isinr(y::'b)|] ==> s1 << s2 = (x << y)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac select_equality 1), - (dtac conjunct2 2), - (dtac conjunct1 2), - (dtac spec 2), - (dtac spec 2), - (etac mp 2), - (fast_tac HOL_cs 2), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (UU_right "x"), - (UU_left "u"), - (Simp_tac 1), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (eq_right "x" "v"), - (eq_right "y" "ya"), - (rtac refl 1), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (UU_right "x"), - (UU_left "u"), - (Simp_tac 1), - (strip_tac 1), - (etac conjE 1), - (eq_right "x" "v"), - (UU_right "y"), - (rtac iffI 1), - (etac UU_I 1), - (res_inst_tac [("s","UU::'b"),("t","x")] subst 1), - (etac sym 1), - (rtac refl_less 1) - ]); +val prems = goalw thy [less_ssum_def] +"[|s1=Isinl(x::'a); s2=Isinr(y::'b)|] ==> s1 << s2 = ((x::'a) = UU)"; +by (cut_facts_tac prems 1); +by (rtac select_equality 1); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (eq_left "x" "u"); +by (UU_left "xa"); +by (rtac iffI 1); +by (res_inst_tac [("s","x"),("t","UU::'a")] subst 1); +by (atac 1); +by (rtac refl_less 1); +by (etac UU_I 1); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (UU_left "x"); +by (UU_right "v"); +by (Simp_tac 1); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (eq_left "x" "u"); +by (rtac refl 1); +by (strip_tac 1); +by (etac conjE 1); +by (UU_left "x"); +by (UU_right "v"); +by (Simp_tac 1); +by (dtac conjunct2 1); +by (dtac conjunct2 1); +by (dtac conjunct1 1); +by (dtac spec 1); +by (dtac spec 1); +by (etac mp 1); +by (fast_tac HOL_cs 1); +qed "less_ssum1c"; -val less_ssum1c = prove_goalw thy [less_ssum_def] -"[|s1=Isinl(x::'a); s2=Isinr(y::'b)|] ==> s1 << s2 = ((x::'a) = UU)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac select_equality 1), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (eq_left "x" "u"), - (UU_left "xa"), - (rtac iffI 1), - (res_inst_tac [("s","x"),("t","UU::'a")] subst 1), - (atac 1), - (rtac refl_less 1), - (etac UU_I 1), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (UU_left "x"), - (UU_right "v"), - (Simp_tac 1), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (eq_left "x" "u"), - (rtac refl 1), - (strip_tac 1), - (etac conjE 1), - (UU_left "x"), - (UU_right "v"), - (Simp_tac 1), - (dtac conjunct2 1), - (dtac conjunct2 1), - (dtac conjunct1 1), - (dtac spec 1), - (dtac spec 1), - (etac mp 1), - (fast_tac HOL_cs 1) - ]); - - -val less_ssum1d = prove_goalw thy [less_ssum_def] -"[|s1=Isinr(x); s2=Isinl(y)|] ==> s1 << s2 = (x = UU)" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac select_equality 1), - (dtac conjunct2 2), - (dtac conjunct2 2), - (dtac conjunct2 2), - (dtac spec 2), - (dtac spec 2), - (etac mp 2), - (fast_tac HOL_cs 2), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (UU_right "x"), - (UU_left "u"), - (Simp_tac 1), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (UU_right "ya"), - (eq_right "x" "v"), - (rtac iffI 1), - (etac UU_I 2), - (res_inst_tac [("s","UU"),("t","x")] subst 1), - (etac sym 1), - (rtac refl_less 1), - (rtac conjI 1), - (strip_tac 1), - (etac conjE 1), - (UU_right "x"), - (UU_left "u"), - (Simp_tac 1), - (strip_tac 1), - (etac conjE 1), - (eq_right "x" "v"), - (rtac refl 1) - ]) -end; +val prems = goalw thy [less_ssum_def] +"[|s1=Isinr(x); s2=Isinl(y)|] ==> s1 << s2 = (x = UU)"; +by (cut_facts_tac prems 1); +by (rtac select_equality 1); +by (dtac conjunct2 2); +by (dtac conjunct2 2); +by (dtac conjunct2 2); +by (dtac spec 2); +by (dtac spec 2); +by (etac mp 2); +by (fast_tac HOL_cs 2); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (UU_right "x"); +by (UU_left "u"); +by (Simp_tac 1); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (UU_right "ya"); +by (eq_right "x" "v"); +by (rtac iffI 1); +by (etac UU_I 2); +by (res_inst_tac [("s","UU"),("t","x")] subst 1); +by (etac sym 1); +by (rtac refl_less 1); +by (rtac conjI 1); +by (strip_tac 1); +by (etac conjE 1); +by (UU_right "x"); +by (UU_left "u"); +by (Simp_tac 1); +by (strip_tac 1); +by (etac conjE 1); +by (eq_right "x" "v"); +by (rtac refl 1); +qed "less_ssum1d"; (* ------------------------------------------------------------------------ *) diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Ssum2.ML --- a/src/HOLCF/Ssum2.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Ssum2.ML Tue Jul 04 15:58:11 2000 +0200 @@ -62,19 +62,15 @@ (* Isinl, Isinr are monotone *) (* ------------------------------------------------------------------------ *) -qed_goalw "monofun_Isinl" thy [monofun] "monofun(Isinl)" - (fn prems => - [ - (strip_tac 1), - (etac (less_ssum3a RS iffD2) 1) - ]); +val prems = goalw thy [monofun] "monofun(Isinl)"; +by (strip_tac 1); +by (etac (less_ssum3a RS iffD2) 1); +qed "monofun_Isinl"; -qed_goalw "monofun_Isinr" thy [monofun] "monofun(Isinr)" - (fn prems => - [ - (strip_tac 1), - (etac (less_ssum3b RS iffD2) 1) - ]); +val prems = goalw thy [monofun] "monofun(Isinr)"; +by (strip_tac 1); +by (etac (less_ssum3b RS iffD2) 1); +qed "monofun_Isinr"; (* ------------------------------------------------------------------------ *) @@ -82,75 +78,69 @@ (* ------------------------------------------------------------------------ *) -qed_goalw "monofun_Iwhen1" thy [monofun] "monofun(Iwhen)" - (fn prems => - [ - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (res_inst_tac [("p","xb")] IssumE 1), - (hyp_subst_tac 1), - (asm_simp_tac Ssum0_ss 1), - (asm_simp_tac Ssum0_ss 1), - (etac monofun_cfun_fun 1), - (asm_simp_tac Ssum0_ss 1) - ]); +val prems = goalw thy [monofun] "monofun(Iwhen)"; +by (strip_tac 1); +by (rtac (less_fun RS iffD2) 1); +by (strip_tac 1); +by (rtac (less_fun RS iffD2) 1); +by (strip_tac 1); +by (res_inst_tac [("p","xb")] IssumE 1); +by (hyp_subst_tac 1); +by (asm_simp_tac Ssum0_ss 1); +by (asm_simp_tac Ssum0_ss 1); +by (etac monofun_cfun_fun 1); +by (asm_simp_tac Ssum0_ss 1); +qed "monofun_Iwhen1"; -qed_goalw "monofun_Iwhen2" thy [monofun] "monofun(Iwhen(f))" - (fn prems => - [ - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (res_inst_tac [("p","xa")] IssumE 1), - (hyp_subst_tac 1), - (asm_simp_tac Ssum0_ss 1), - (asm_simp_tac Ssum0_ss 1), - (asm_simp_tac Ssum0_ss 1), - (etac monofun_cfun_fun 1) - ]); +val prems = goalw thy [monofun] "monofun(Iwhen(f))"; +by (strip_tac 1); +by (rtac (less_fun RS iffD2) 1); +by (strip_tac 1); +by (res_inst_tac [("p","xa")] IssumE 1); +by (hyp_subst_tac 1); +by (asm_simp_tac Ssum0_ss 1); +by (asm_simp_tac Ssum0_ss 1); +by (asm_simp_tac Ssum0_ss 1); +by (etac monofun_cfun_fun 1); +qed "monofun_Iwhen2"; -qed_goalw "monofun_Iwhen3" thy [monofun] "monofun(Iwhen(f)(g))" - (fn prems => - [ - (strip_tac 1), - (res_inst_tac [("p","x")] IssumE 1), - (hyp_subst_tac 1), - (asm_simp_tac Ssum0_ss 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","y")] IssumE 1), - (hyp_subst_tac 1), - (asm_simp_tac Ssum0_ss 1), - (res_inst_tac [("P","xa=UU")] notE 1), - (atac 1), - (rtac UU_I 1), - (rtac (less_ssum3a RS iffD1) 1), - (atac 1), - (hyp_subst_tac 1), - (asm_simp_tac Ssum0_ss 1), - (rtac monofun_cfun_arg 1), - (etac (less_ssum3a RS iffD1) 1), - (hyp_subst_tac 1), - (res_inst_tac [("s","UU"),("t","xa")] subst 1), - (etac (less_ssum3c RS iffD1 RS sym) 1), - (asm_simp_tac Ssum0_ss 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","y")] IssumE 1), - (hyp_subst_tac 1), - (res_inst_tac [("s","UU"),("t","ya")] subst 1), - (etac (less_ssum3d RS iffD1 RS sym) 1), - (asm_simp_tac Ssum0_ss 1), - (hyp_subst_tac 1), - (res_inst_tac [("s","UU"),("t","ya")] subst 1), - (etac (less_ssum3d RS iffD1 RS sym) 1), - (asm_simp_tac Ssum0_ss 1), - (hyp_subst_tac 1), - (asm_simp_tac Ssum0_ss 1), - (rtac monofun_cfun_arg 1), - (etac (less_ssum3b RS iffD1) 1) - ]); +val prems = goalw thy [monofun] "monofun(Iwhen(f)(g))"; +by (strip_tac 1); +by (res_inst_tac [("p","x")] IssumE 1); +by (hyp_subst_tac 1); +by (asm_simp_tac Ssum0_ss 1); +by (hyp_subst_tac 1); +by (res_inst_tac [("p","y")] IssumE 1); +by (hyp_subst_tac 1); +by (asm_simp_tac Ssum0_ss 1); +by (res_inst_tac [("P","xa=UU")] notE 1); +by (atac 1); +by (rtac UU_I 1); +by (rtac (less_ssum3a RS iffD1) 1); +by (atac 1); +by (hyp_subst_tac 1); +by (asm_simp_tac Ssum0_ss 1); +by (rtac monofun_cfun_arg 1); +by (etac (less_ssum3a RS iffD1) 1); +by (hyp_subst_tac 1); +by (res_inst_tac [("s","UU"),("t","xa")] subst 1); +by (etac (less_ssum3c RS iffD1 RS sym) 1); +by (asm_simp_tac Ssum0_ss 1); +by (hyp_subst_tac 1); +by (res_inst_tac [("p","y")] IssumE 1); +by (hyp_subst_tac 1); +by (res_inst_tac [("s","UU"),("t","ya")] subst 1); +by (etac (less_ssum3d RS iffD1 RS sym) 1); +by (asm_simp_tac Ssum0_ss 1); +by (hyp_subst_tac 1); +by (res_inst_tac [("s","UU"),("t","ya")] subst 1); +by (etac (less_ssum3d RS iffD1 RS sym) 1); +by (asm_simp_tac Ssum0_ss 1); +by (hyp_subst_tac 1); +by (asm_simp_tac Ssum0_ss 1); +by (rtac monofun_cfun_arg 1); +by (etac (less_ssum3b RS iffD1) 1); +qed "monofun_Iwhen3"; (* ------------------------------------------------------------------------ *) diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Ssum3.ML --- a/src/HOLCF/Ssum3.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Ssum3.ML Tue Jul 04 15:58:11 2000 +0200 @@ -86,6 +86,8 @@ by (rtac contlub_Isinr 1); qed "cont_Isinr"; +AddIffs [cont_Isinl, cont_Isinr]; + (* ------------------------------------------------------------------------ *) (* continuity for Iwhen in the firts two arguments *) @@ -332,52 +334,42 @@ (* continuous versions of lemmas for 'a ++ 'b *) (* ------------------------------------------------------------------------ *) -qed_goalw "strict_sinl" Ssum3.thy [sinl_def] "sinl`UU =UU" - (fn prems => - [ - (simp_tac (Ssum0_ss addsimps [cont_Isinl]) 1), - (rtac (inst_ssum_pcpo RS sym) 1) - ]); +val prems = goalw Ssum3.thy [sinl_def] "sinl`UU =UU"; +by (simp_tac (Ssum0_ss addsimps [cont_Isinl]) 1); +by (rtac (inst_ssum_pcpo RS sym) 1); +qed "strict_sinl"; -qed_goalw "strict_sinr" Ssum3.thy [sinr_def] "sinr`UU=UU" - (fn prems => - [ - (simp_tac (Ssum0_ss addsimps [cont_Isinr]) 1), - (rtac (inst_ssum_pcpo RS sym) 1) - ]); +val prems = goalw Ssum3.thy [sinr_def] "sinr`UU=UU"; +by (simp_tac (Ssum0_ss addsimps [cont_Isinr]) 1); +by (rtac (inst_ssum_pcpo RS sym) 1); +qed "strict_sinr"; -qed_goalw "noteq_sinlsinr" Ssum3.thy [sinl_def,sinr_def] - "sinl`a=sinr`b ==> a=UU & b=UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac noteq_IsinlIsinr 1), - (etac box_equals 1), - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1), - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1) - ]); +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "sinl`a=sinr`b ==> a=UU & b=UU"; +by (cut_facts_tac prems 1); +by (rtac noteq_IsinlIsinr 1); +by (etac box_equals 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1); +qed "noteq_sinlsinr"; -qed_goalw "inject_sinl" Ssum3.thy [sinl_def,sinr_def] - "sinl`a1=sinl`a2==> a1=a2" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac inject_Isinl 1), - (etac box_equals 1), - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1), - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1) - ]); +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "sinl`a1=sinl`a2==> a1=a2"; +by (cut_facts_tac prems 1); +by (rtac inject_Isinl 1); +by (etac box_equals 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1); +qed "inject_sinl"; -qed_goalw "inject_sinr" Ssum3.thy [sinl_def,sinr_def] - "sinr`a1=sinr`a2==> a1=a2" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac inject_Isinr 1), - (etac box_equals 1), - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1), - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1) - ]); +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "sinr`a1=sinr`a2==> a1=a2"; +by (cut_facts_tac prems 1); +by (rtac inject_Isinr 1); +by (etac box_equals 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1); +qed "inject_sinr"; Goal "x~=UU ==> sinl`x ~= UU"; @@ -394,227 +386,192 @@ by (etac notnotD 1); qed "defined_sinr"; -qed_goalw "Exh_Ssum1" Ssum3.thy [sinl_def,sinr_def] - "z=UU | (? a. z=sinl`a & a~=UU) | (? b. z=sinr`b & b~=UU)" - (fn prems => - [ - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1), - (stac inst_ssum_pcpo 1), - (rtac Exh_Ssum 1) - ]); +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "z=UU | (? a. z=sinl`a & a~=UU) | (? b. z=sinr`b & b~=UU)"; +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1); +by (stac inst_ssum_pcpo 1); +by (rtac Exh_Ssum 1); +qed "Exh_Ssum1"; -qed_goalw "ssumE" Ssum3.thy [sinl_def,sinr_def] +val [major,prem2,prem3] = Goalw [sinl_def,sinr_def] "[|p=UU ==> Q ;\ \ !!x.[|p=sinl`x; x~=UU |] ==> Q;\ -\ !!y.[|p=sinr`y; y~=UU |] ==> Q|] ==> Q" - (fn prems => - [ - (rtac IssumE 1), - (resolve_tac prems 1), - (stac inst_ssum_pcpo 1), - (atac 1), - (resolve_tac prems 1), - (atac 2), - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1), - (resolve_tac prems 1), - (atac 2), - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1) - ]); +\ !!y.[|p=sinr`y; y~=UU |] ==> Q|] ==> Q"; +by (rtac (major RS IssumE) 1); +by (stac inst_ssum_pcpo 1); +by (atac 1); +by (rtac prem2 1); +by (atac 2); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1); +by (rtac prem3 1); +by (atac 2); +by (Asm_simp_tac 1); +qed "ssumE"; -qed_goalw "ssumE2" Ssum3.thy [sinl_def,sinr_def] +val [preml,premr] = Goalw [sinl_def,sinr_def] "[|!!x.[|p=sinl`x|] ==> Q;\ -\ !!y.[|p=sinr`y|] ==> Q|] ==> Q" - (fn prems => - [ - (rtac IssumE2 1), - (resolve_tac prems 1), - (stac beta_cfun 1), - (rtac cont_Isinl 1), - (atac 1), - (resolve_tac prems 1), - (stac beta_cfun 1), - (rtac cont_Isinr 1), - (atac 1) - ]); +\ !!y.[|p=sinr`y|] ==> Q|] ==> Q"; +by (rtac IssumE2 1); +by (rtac preml 1); +by (rtac premr 2); +by Auto_tac; +qed "ssumE2"; + +val tac = (REPEAT (resolve_tac (cont_lemmas1 @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont2cont_CF1L]) 1)); -qed_goalw "sscase1" Ssum3.thy [sscase_def,sinl_def,sinr_def] - "sscase`f`g`UU = UU" (fn _ => let -val tac = (REPEAT (resolve_tac (cont_lemmas1 @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont2cont_CF1L]) 1)) in - [ - (stac inst_ssum_pcpo 1), - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (simp_tac Ssum0_ss 1) - ] end); +val _ = goalw Ssum3.thy [sscase_def,sinl_def,sinr_def] + "sscase`f`g`UU = UU"; +by (stac inst_ssum_pcpo 1); +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (simp_tac Ssum0_ss 1); +qed "sscase1"; val tac = (REPEAT (resolve_tac (cont_lemmas1 @ [cont_Iwhen1,cont_Iwhen2, cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)); -qed_goalw "sscase2" Ssum3.thy [sscase_def,sinl_def,sinr_def] - "x~=UU==> sscase`f`g`(sinl`x) = f`x" (fn prems => [ - (cut_facts_tac prems 1), - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (asm_simp_tac Ssum0_ss 1) - ]); - -qed_goalw "sscase3" Ssum3.thy [sscase_def,sinl_def,sinr_def] - "x~=UU==> sscase`f`g`(sinr`x) = g`x" (fn prems => [ - (cut_facts_tac prems 1), - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (asm_simp_tac Ssum0_ss 1) - ]); - - -qed_goalw "less_ssum4a" Ssum3.thy [sinl_def,sinr_def] - "(sinl`x << sinl`y) = (x << y)" (fn prems => [ - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (rtac less_ssum3a 1) - ]); +val prems = goalw Ssum3.thy [sscase_def,sinl_def,sinr_def] + "x~=UU==> sscase`f`g`(sinl`x) = f`x"; +by (cut_facts_tac prems 1); +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (asm_simp_tac Ssum0_ss 1); +qed "sscase2"; -qed_goalw "less_ssum4b" Ssum3.thy [sinl_def,sinr_def] - "(sinr`x << sinr`y) = (x << y)" (fn prems => [ - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (rtac less_ssum3b 1) - ]); - -qed_goalw "less_ssum4c" Ssum3.thy [sinl_def,sinr_def] - "(sinl`x << sinr`y) = (x = UU)" (fn prems => - [ - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (rtac less_ssum3c 1) - ]); - -qed_goalw "less_ssum4d" Ssum3.thy [sinl_def,sinr_def] - "(sinr`x << sinl`y) = (x = UU)" - (fn prems => - [ - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (rtac less_ssum3d 1) - ]); - -qed_goalw "ssum_chainE" Ssum3.thy [sinl_def,sinr_def] - "chain(Y) ==> (!i.? x.(Y i)=sinl`x)|(!i.? y.(Y i)=sinr`y)" - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1), - (etac ssum_lemma4 1) - ]); +val prems = goalw Ssum3.thy [sscase_def,sinl_def,sinr_def] + "x~=UU==> sscase`f`g`(sinr`x) = g`x"; +by (cut_facts_tac prems 1); +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (asm_simp_tac Ssum0_ss 1); +qed "sscase3"; -qed_goalw "thelub_ssum2a" Ssum3.thy [sinl_def,sinr_def,sscase_def] +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "(sinl`x << sinl`y) = (x << y)"; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (rtac less_ssum3a 1); +qed "less_ssum4a"; + +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "(sinr`x << sinr`y) = (x << y)"; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (rtac less_ssum3b 1); +qed "less_ssum4b"; + +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "(sinl`x << sinr`y) = (x = UU)"; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (rtac less_ssum3c 1); +qed "less_ssum4c"; + +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "(sinr`x << sinl`y) = (x = UU)"; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (rtac less_ssum3d 1); +qed "less_ssum4d"; + +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "chain(Y) ==> (!i.? x.(Y i)=sinl`x)|(!i.? y.(Y i)=sinr`y)"; +by (cut_facts_tac prems 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1); +by (etac ssum_lemma4 1); +qed "ssum_chainE"; + + +val prems = goalw Ssum3.thy [sinl_def,sinr_def,sscase_def] "[| chain(Y); !i.? x. Y(i) = sinl`x |] ==>\ -\ lub(range(Y)) = sinl`(lub(range(%i. sscase`(LAM x. x)`(LAM y. UU)`(Y i))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (stac (beta_cfun RS ext) 1), - tac, - (rtac thelub_ssum1a 1), - (atac 1), - (rtac allI 1), - (etac allE 1), - (etac exE 1), - (rtac exI 1), - (etac box_equals 1), - (rtac refl 1), - (asm_simp_tac (Ssum0_ss addsimps [cont_Isinl]) 1) - ]); +\ lub(range(Y)) = sinl`(lub(range(%i. sscase`(LAM x. x)`(LAM y. UU)`(Y i))))"; +by (cut_facts_tac prems 1); +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (stac (beta_cfun RS ext) 1); +by tac; +by (rtac thelub_ssum1a 1); +by (atac 1); +by (rtac allI 1); +by (etac allE 1); +by (etac exE 1); +by (rtac exI 1); +by (etac box_equals 1); +by (rtac refl 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinl]) 1); +qed "thelub_ssum2a"; -qed_goalw "thelub_ssum2b" Ssum3.thy [sinl_def,sinr_def,sscase_def] +val prems = goalw Ssum3.thy [sinl_def,sinr_def,sscase_def] "[| chain(Y); !i.? x. Y(i) = sinr`x |] ==>\ -\ lub(range(Y)) = sinr`(lub(range(%i. sscase`(LAM y. UU)`(LAM x. x)`(Y i))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (stac (beta_cfun RS ext) 1), - tac, - (rtac thelub_ssum1b 1), - (atac 1), - (rtac allI 1), - (etac allE 1), - (etac exE 1), - (rtac exI 1), - (etac box_equals 1), - (rtac refl 1), - (asm_simp_tac (Ssum0_ss addsimps - [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3]) 1) - ]); +\ lub(range(Y)) = sinr`(lub(range(%i. sscase`(LAM y. UU)`(LAM x. x)`(Y i))))"; +by (cut_facts_tac prems 1); +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (stac (beta_cfun RS ext) 1); +by tac; +by (rtac thelub_ssum1b 1); +by (atac 1); +by (rtac allI 1); +by (etac allE 1); +by (etac exE 1); +by (rtac exI 1); +by (etac box_equals 1); +by (rtac refl 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, cont_Iwhen3]) 1); +qed "thelub_ssum2b"; -qed_goalw "thelub_ssum2a_rev" Ssum3.thy [sinl_def,sinr_def] - "[| chain(Y); lub(range(Y)) = sinl`x|] ==> !i.? x. Y(i)=sinl`x" - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (Ssum0_ss addsimps - [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3]) 1), - (etac ssum_lemma9 1), - (asm_simp_tac (Ssum0_ss addsimps - [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3]) 1) - ]); +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "[| chain(Y); lub(range(Y)) = sinl`x|] ==> !i.? x. Y(i)=sinl`x"; +by (cut_facts_tac prems 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, cont_Iwhen3]) 1); +by (etac ssum_lemma9 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, cont_Iwhen3]) 1); +qed "thelub_ssum2a_rev"; -qed_goalw "thelub_ssum2b_rev" Ssum3.thy [sinl_def,sinr_def] - "[| chain(Y); lub(range(Y)) = sinr`x|] ==> !i.? x. Y(i)=sinr`x" - (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (Ssum0_ss addsimps - [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3]) 1), - (etac ssum_lemma10 1), - (asm_simp_tac (Ssum0_ss addsimps - [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3]) 1) - ]); +val prems = goalw Ssum3.thy [sinl_def,sinr_def] + "[| chain(Y); lub(range(Y)) = sinr`x|] ==> !i.? x. Y(i)=sinr`x"; +by (cut_facts_tac prems 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, cont_Iwhen3]) 1); +by (etac ssum_lemma10 1); +by (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, cont_Iwhen3]) 1); +qed "thelub_ssum2b_rev"; Goal "chain(Y) ==>\ \ lub(range(Y)) = sinl`(lub(range(%i. sscase`(LAM x. x)`(LAM y. UU)`(Y i))))\ diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Tr.ML --- a/src/HOLCF/Tr.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Tr.ML Tue Jul 04 15:58:11 2000 +0200 @@ -3,21 +3,18 @@ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen -Lemmas for Tr.thy +Introduce infix if_then_else_fi and boolean connectives andalso, orelse *) -open Tr; - (* ------------------------------------------------------------------------ *) (* Exhaustion and Elimination for type one *) (* ------------------------------------------------------------------------ *) -qed_goalw "Exh_tr" thy [FF_def,TT_def] "t=UU | t = TT | t = FF" - (fn prems => - [ - (lift.induct_tac "t" 1), - (fast_tac HOL_cs 1), - (fast_tac (HOL_cs addss simpset()) 1) - ]); + +Goalw [FF_def,TT_def] "t=UU | t = TT | t = FF"; +by (lift.induct_tac "t" 1); +by (fast_tac HOL_cs 1); +by (fast_tac (HOL_cs addss simpset()) 1); +qed "Exh_tr"; val prems = Goal "[| p=UU ==> Q; p = TT ==>Q; p = FF ==>Q|] ==>Q"; by (rtac (Exh_tr RS disjE) 1); diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Up1.ML --- a/src/HOLCF/Up1.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Up1.ML Tue Jul 04 15:58:11 2000 +0200 @@ -2,27 +2,27 @@ ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen + +Lifting *) Goal "Rep_Up (Abs_Up y) = y"; by (simp_tac (simpset() addsimps [Up_def,Abs_Up_inverse]) 1); qed "Abs_Up_inverse2"; -qed_goalw "Exh_Up" thy [Iup_def ] - "z = Abs_Up(Inl ()) | (? x. z = Iup x)" - (fn prems => - [ - (rtac (Rep_Up_inverse RS subst) 1), - (res_inst_tac [("s","Rep_Up z")] sumE 1), - (rtac disjI1 1), - (res_inst_tac [("f","Abs_Up")] arg_cong 1), - (rtac (unit_eq RS subst) 1), - (atac 1), - (rtac disjI2 1), - (rtac exI 1), - (res_inst_tac [("f","Abs_Up")] arg_cong 1), - (atac 1) - ]); +val prems = Goalw [Iup_def ] + "z = Abs_Up(Inl ()) | (? x. z = Iup x)"; +by (rtac (Rep_Up_inverse RS subst) 1); +by (res_inst_tac [("s","Rep_Up z")] sumE 1); +by (rtac disjI1 1); +by (res_inst_tac [("f","Abs_Up")] arg_cong 1); +by (rtac (unit_eq RS subst) 1); +by (atac 1); +by (rtac disjI2 1); +by (rtac exI 1); +by (res_inst_tac [("f","Abs_Up")] arg_cong 1); +by (atac 1); +qed "Exh_Up"; Goal "inj(Abs_Up)"; by (rtac inj_inverseI 1); @@ -34,26 +34,22 @@ by (rtac Rep_Up_inverse 1); qed "inj_Rep_Up"; -qed_goalw "inject_Iup" thy [Iup_def] "Iup x=Iup y ==> x=y" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (inj_Inr RS injD) 1), - (rtac (inj_Abs_Up RS injD) 1), - (atac 1) - ]); +val prems = goalw thy [Iup_def] "Iup x=Iup y ==> x=y"; +by (cut_facts_tac prems 1); +by (rtac (inj_Inr RS injD) 1); +by (rtac (inj_Abs_Up RS injD) 1); +by (atac 1); +qed "inject_Iup"; AddSDs [inject_Iup]; -qed_goalw "defined_Iup" thy [Iup_def] "Iup x~=Abs_Up(Inl ())" - (fn prems => - [ - (rtac notI 1), - (rtac notE 1), - (rtac Inl_not_Inr 1), - (rtac sym 1), - (etac (inj_Abs_Up RS injD) 1) - ]); +val prems = goalw thy [Iup_def] "Iup x~=Abs_Up(Inl ())"; +by (rtac notI 1); +by (rtac notE 1); +by (rtac Inl_not_Inr 1); +by (rtac sym 1); +by (etac (inj_Abs_Up RS injD) 1); +qed "defined_Iup"; val prems = Goal "[| p=Abs_Up(Inl ()) ==> Q; !!x. p=Iup(x)==>Q|] ==>Q"; @@ -63,62 +59,52 @@ by (eresolve_tac prems 1); qed "upE"; -qed_goalw "Ifup1" thy [Ifup_def] - "Ifup(f)(Abs_Up(Inl ()))=UU" - (fn prems => - [ - (stac Abs_Up_inverse2 1), - (stac sum_case_Inl 1), - (rtac refl 1) - ]); +val prems = goalw thy [Ifup_def] + "Ifup(f)(Abs_Up(Inl ()))=UU"; +by (stac Abs_Up_inverse2 1); +by (stac sum_case_Inl 1); +by (rtac refl 1); +qed "Ifup1"; -qed_goalw "Ifup2" thy [Ifup_def,Iup_def] - "Ifup(f)(Iup(x))=f`x" - (fn prems => - [ - (stac Abs_Up_inverse2 1), - (stac sum_case_Inr 1), - (rtac refl 1) - ]); +val prems = goalw thy [Ifup_def,Iup_def] + "Ifup(f)(Iup(x))=f`x"; +by (stac Abs_Up_inverse2 1); +by (stac sum_case_Inr 1); +by (rtac refl 1); +qed "Ifup2"; val Up0_ss = (simpset_of Cfun3.thy) delsimps [range_composition] addsimps [Ifup1,Ifup2]; Addsimps [Ifup1,Ifup2]; -qed_goalw "less_up1a" thy [less_up_def] - "Abs_Up(Inl ())<< z" - (fn prems => - [ - (stac Abs_Up_inverse2 1), - (stac sum_case_Inl 1), - (rtac TrueI 1) - ]); +val prems = goalw thy [less_up_def] + "Abs_Up(Inl ())<< z"; +by (stac Abs_Up_inverse2 1); +by (stac sum_case_Inl 1); +by (rtac TrueI 1); +qed "less_up1a"; -qed_goalw "less_up1b" thy [Iup_def,less_up_def] - "~(Iup x) << (Abs_Up(Inl ()))" - (fn prems => - [ - (rtac notI 1), - (rtac iffD1 1), - (atac 2), - (stac Abs_Up_inverse2 1), - (stac Abs_Up_inverse2 1), - (stac sum_case_Inr 1), - (stac sum_case_Inl 1), - (rtac refl 1) - ]); +val prems = goalw thy [Iup_def,less_up_def] + "~(Iup x) << (Abs_Up(Inl ()))"; +by (rtac notI 1); +by (rtac iffD1 1); +by (atac 2); +by (stac Abs_Up_inverse2 1); +by (stac Abs_Up_inverse2 1); +by (stac sum_case_Inr 1); +by (stac sum_case_Inl 1); +by (rtac refl 1); +qed "less_up1b"; -qed_goalw "less_up1c" thy [Iup_def,less_up_def] - "(Iup x) << (Iup y)=(x< - [ - (stac Abs_Up_inverse2 1), - (stac Abs_Up_inverse2 1), - (stac sum_case_Inr 1), - (stac sum_case_Inr 1), - (rtac refl 1) - ]); +val prems = goalw thy [Iup_def,less_up_def] + "(Iup x) << (Iup y)=(x< - [ - (strip_tac 1), - (etac (less_up2c RS iffD2) 1) - ]); +val prems = goalw thy [monofun] "monofun(Iup)"; +by (strip_tac 1); +by (etac (less_up2c RS iffD2) 1); +qed "monofun_Iup"; + +val prems = goalw thy [monofun] "monofun(Ifup)"; +by (strip_tac 1); +by (rtac (less_fun RS iffD2) 1); +by (strip_tac 1); +by (res_inst_tac [("p","xa")] upE 1); +by (asm_simp_tac Up0_ss 1); +by (asm_simp_tac Up0_ss 1); +by (etac monofun_cfun_fun 1); +qed "monofun_Ifup1"; -qed_goalw "monofun_Ifup1" thy [monofun] "monofun(Ifup)" - (fn prems => - [ - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (res_inst_tac [("p","xa")] upE 1), - (asm_simp_tac Up0_ss 1), - (asm_simp_tac Up0_ss 1), - (etac monofun_cfun_fun 1) - ]); - -qed_goalw "monofun_Ifup2" thy [monofun] "monofun(Ifup(f))" - (fn prems => - [ - (strip_tac 1), - (res_inst_tac [("p","x")] upE 1), - (asm_simp_tac Up0_ss 1), - (asm_simp_tac Up0_ss 1), - (res_inst_tac [("p","y")] upE 1), - (hyp_subst_tac 1), - (rtac notE 1), - (rtac less_up2b 1), - (atac 1), - (asm_simp_tac Up0_ss 1), - (rtac monofun_cfun_arg 1), - (hyp_subst_tac 1), - (etac (less_up2c RS iffD1) 1) - ]); +val prems = goalw thy [monofun] "monofun(Ifup(f))"; +by (strip_tac 1); +by (res_inst_tac [("p","x")] upE 1); +by (asm_simp_tac Up0_ss 1); +by (asm_simp_tac Up0_ss 1); +by (res_inst_tac [("p","y")] upE 1); +by (hyp_subst_tac 1); +by (rtac notE 1); +by (rtac less_up2b 1); +by (atac 1); +by (asm_simp_tac Up0_ss 1); +by (rtac monofun_cfun_arg 1); +by (hyp_subst_tac 1); +by (etac (less_up2c RS iffD1) 1); +qed "monofun_Ifup2"; (* ------------------------------------------------------------------------ *) (* Some kind of surjectivity lemma *) diff -r 7edd3e5f26d4 -r 428385c4bc50 src/HOLCF/Up3.ML --- a/src/HOLCF/Up3.ML Tue Jul 04 14:58:40 2000 +0200 +++ b/src/HOLCF/Up3.ML Tue Jul 04 15:58:11 2000 +0200 @@ -132,126 +132,106 @@ (* continuous versions of lemmas for ('a)u *) (* ------------------------------------------------------------------------ *) -qed_goalw "Exh_Up1" thy [up_def] "z = UU | (? x. z = up`x)" - (fn prems => - [ - (simp_tac (Up0_ss addsimps [cont_Iup]) 1), - (stac inst_up_pcpo 1), - (rtac Exh_Up 1) - ]); +val prems = goalw thy [up_def] "z = UU | (? x. z = up`x)"; +by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); +by (stac inst_up_pcpo 1); +by (rtac Exh_Up 1); +qed "Exh_Up1"; -qed_goalw "inject_up" thy [up_def] "up`x=up`y ==> x=y" - (fn prems => - [ - (cut_facts_tac prems 1), - (rtac inject_Iup 1), - (etac box_equals 1), - (simp_tac (Up0_ss addsimps [cont_Iup]) 1), - (simp_tac (Up0_ss addsimps [cont_Iup]) 1) - ]); +val prems = goalw thy [up_def] "up`x=up`y ==> x=y"; +by (cut_facts_tac prems 1); +by (rtac inject_Iup 1); +by (etac box_equals 1); +by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); +by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); +qed "inject_up"; -qed_goalw "defined_up" thy [up_def] " up`x ~= UU" - (fn prems => - [ - (simp_tac (Up0_ss addsimps [cont_Iup]) 1), - (rtac defined_Iup2 1) - ]); +val prems = goalw thy [up_def] " up`x ~= UU"; +by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); +by (rtac defined_Iup2 1); +qed "defined_up"; -qed_goalw "upE1" thy [up_def] - "[| p=UU ==> Q; !!x. p=up`x==>Q|] ==>Q" - (fn prems => - [ - (rtac upE 1), - (resolve_tac prems 1), - (etac (inst_up_pcpo RS ssubst) 1), - (resolve_tac (tl prems) 1), - (asm_simp_tac (Up0_ss addsimps [cont_Iup]) 1) - ]); +val prems = goalw thy [up_def] + "[| p=UU ==> Q; !!x. p=up`x==>Q|] ==>Q"; +by (rtac upE 1); +by (resolve_tac prems 1); +by (etac (inst_up_pcpo RS ssubst) 1); +by (resolve_tac (tl prems) 1); +by (asm_simp_tac (Up0_ss addsimps [cont_Iup]) 1); +qed "upE1"; val tac = (simp_tac (simpset() addsimps [cont_Iup,cont_Ifup1, cont_Ifup2,cont2cont_CF1L]) 1); -qed_goalw "fup1" thy [up_def,fup_def] "fup`f`UU=UU" - (fn prems => - [ - (stac inst_up_pcpo 1), - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (simp_tac (Up0_ss addsimps [cont_Iup,cont_Ifup1,cont_Ifup2]) 1) - ]); +val prems = goalw thy [up_def,fup_def] "fup`f`UU=UU"; +by (stac inst_up_pcpo 1); +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (simp_tac (Up0_ss addsimps [cont_Iup,cont_Ifup1,cont_Ifup2]) 1); +qed "fup1"; -qed_goalw "fup2" thy [up_def,fup_def] "fup`f`(up`x)=f`x" - (fn prems => - [ - (stac beta_cfun 1), - (rtac cont_Iup 1), - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - (rtac cont_Ifup2 1), - (simp_tac (Up0_ss addsimps [cont_Iup,cont_Ifup1,cont_Ifup2]) 1) - ]); +val prems = goalw thy [up_def,fup_def] "fup`f`(up`x)=f`x"; +by (stac beta_cfun 1); +by (rtac cont_Iup 1); +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by (rtac cont_Ifup2 1); +by (simp_tac (Up0_ss addsimps [cont_Iup,cont_Ifup1,cont_Ifup2]) 1); +qed "fup2"; -qed_goalw "less_up4b" thy [up_def,fup_def] "~ up`x << UU" - (fn prems => - [ - (simp_tac (Up0_ss addsimps [cont_Iup]) 1), - (rtac less_up3b 1) - ]); +val prems = goalw thy [up_def,fup_def] "~ up`x << UU"; +by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); +by (rtac less_up3b 1); +qed "less_up4b"; -qed_goalw "less_up4c" thy [up_def,fup_def] - "(up`x << up`y) = (x< - [ - (simp_tac (Up0_ss addsimps [cont_Iup]) 1), - (rtac less_up2c 1) - ]); +val prems = goalw thy [up_def,fup_def] + "(up`x << up`y) = (x<\ -\ lub(range(Y)) = up`(lub(range(%i. fup`(LAM x. x)`(Y i))))" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac beta_cfun 1), - tac, - (stac beta_cfun 1), - tac, - (stac (beta_cfun RS ext) 1), - tac, - (rtac thelub_up1a 1), - (atac 1), - (etac exE 1), - (etac exE 1), - (rtac exI 1), - (rtac exI 1), - (etac box_equals 1), - (rtac refl 1), - (simp_tac (Up0_ss addsimps [cont_Iup]) 1) - ]); +\ lub(range(Y)) = up`(lub(range(%i. fup`(LAM x. x)`(Y i))))"; +by (cut_facts_tac prems 1); +by (stac beta_cfun 1); +by tac; +by (stac beta_cfun 1); +by tac; +by (stac (beta_cfun RS ext) 1); +by tac; +by (rtac thelub_up1a 1); +by (atac 1); +by (etac exE 1); +by (etac exE 1); +by (rtac exI 1); +by (rtac exI 1); +by (etac box_equals 1); +by (rtac refl 1); +by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); +qed "thelub_up2a"; -qed_goalw "thelub_up2b" thy [up_def,fup_def] -"[| chain(Y); ! i x. Y(i) ~= up`x |] ==> lub(range(Y)) = UU" - (fn prems => - [ - (cut_facts_tac prems 1), - (stac inst_up_pcpo 1), - (rtac thelub_up1b 1), - (atac 1), - (strip_tac 1), - (dtac spec 1), - (dtac spec 1), - (rtac swap 1), - (atac 1), - (dtac notnotD 1), - (etac box_equals 1), - (rtac refl 1), - (simp_tac (Up0_ss addsimps [cont_Iup]) 1) - ]); +val prems = goalw thy [up_def,fup_def] +"[| chain(Y); ! i x. Y(i) ~= up`x |] ==> lub(range(Y)) = UU"; +by (cut_facts_tac prems 1); +by (stac inst_up_pcpo 1); +by (rtac thelub_up1b 1); +by (atac 1); +by (strip_tac 1); +by (dtac spec 1); +by (dtac spec 1); +by (rtac swap 1); +by (atac 1); +by (dtac notnotD 1); +by (etac box_equals 1); +by (rtac refl 1); +by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); +qed "thelub_up2b"; Goal "(? x. z = up`x) = (z~=UU)"; @@ -277,7 +257,6 @@ qed "thelub_up2a_rev"; Goal "[| chain(Y); lub(range(Y)) = UU |] ==> ! i x. Y(i) ~= up`x"; -by (cut_facts_tac prems 1); by (rtac allI 1); by (rtac (not_ex RS iffD1) 1); by (rtac contrapos 1); @@ -288,7 +267,6 @@ Goal "chain(Y) ==> lub(range(Y)) = UU | \ \ lub(range(Y)) = up`(lub(range(%i. fup`(LAM x. x)`(Y i))))"; -by (cut_facts_tac prems 1); by (rtac disjE 1); by (rtac disjI1 2); by (rtac thelub_up2b 2);