src/HOLCF/Cont.ML
changeset 9248 e1dee89de037
parent 9245 428385c4bc50
child 12030 46d57d0290a2
--- a/src/HOLCF/Cont.ML	Wed Jul 05 14:26:58 2000 +0200
+++ b/src/HOLCF/Cont.ML	Wed Jul 05 16:37:52 2000 +0200
@@ -10,43 +10,37 @@
 (* access to definition                                                     *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goalw thy [contlub]
+Goalw [contlub]
         "! Y. chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))==>\
 \        contlub(f)";
-by (cut_facts_tac prems 1);
 by (atac 1);
 qed "contlubI";
 
-val prems = goalw thy [contlub]
+Goalw [contlub]
         " contlub(f)==>\
 \         ! Y. chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))";
-by (cut_facts_tac prems 1);
 by (atac 1);
 qed "contlubE";
 
 
-val prems = goalw thy [cont]
+Goalw [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";
 
-val prems = goalw thy [cont]
+Goalw [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";
 
 
-val prems = goalw thy [monofun]
+Goalw [monofun]
         "! x y. x << y --> f(x) << f(y) ==> monofun(f)";
-by (cut_facts_tac prems 1);
 by (atac 1);
 qed "monofunI";
 
-val prems = goalw thy [monofun]
+Goalw [monofun]
         "monofun(f) ==> ! x y. x << y --> f(x) << f(y)";
-by (cut_facts_tac prems 1);
 by (atac 1);
 qed "monofunE";
 
@@ -59,35 +53,30 @@
 (* monotone functions map chains to chains                                  *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goal thy 
+Goal 
         "[| 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);
+by (etac (chainE) 1);
 qed "ch2ch_monofun";
 
 (* ------------------------------------------------------------------------ *)
 (* monotone functions map upper bound to upper bounds                       *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goal thy 
+Goal 
  "[| 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);
+by (etac (ub_rangeD) 1);
 qed "ub2ub_monofun";
 
 (* ------------------------------------------------------------------------ *)
 (* left to right: monofun(f) & contlub(f)  ==> cont(f)                     *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goalw thy [cont]
+Goalw [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);
@@ -155,77 +144,70 @@
 (* in both arguments                                                        *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goal thy 
+Goal 
 "[|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";
 
 
-val prems = goal thy 
+Goal 
 "[|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";
 
-val prems = goal thy 
+Goal 
 "[|monofun(MF2); !f. monofun(MF2(f)); chain(F); chain(Y)|] ==> \
 \  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 (etac (ch2ch_MF2L RS chainE) 1);
 by (atac 1);
 by ((rtac (monofunE RS spec RS spec RS mp) 1) THEN (etac spec 1));
-by (etac (chainE RS spec) 1);
+by (etac (chainE) 1);
 qed "ch2ch_MF2LR";
 
 
-val prems = goal thy 
+Goal 
 "[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\
 \  !f. monofun(MF2(f)::('b::po=>'c::cpo));\
 \       chain(F);chain(Y)|] ==> \
 \       chain(%j. lub(range(%i. MF2 (F j) (Y i))))";
-by (cut_facts_tac prems 1);
-by (rtac (lub_mono RS allI RS chainI) 1);
+by (rtac (lub_mono RS chainI) 1);
 by ((rtac ch2ch_MF2R 1) THEN (etac spec 1));
 by (atac 1);
 by ((rtac ch2ch_MF2R 1) THEN (etac spec 1));
 by (atac 1);
 by (strip_tac 1);
-by (rtac (chainE RS spec) 1);
+by (rtac (chainE) 1);
 by (etac ch2ch_MF2L 1);
 by (atac 1);
 qed "ch2ch_lubMF2R";
 
 
-val prems = goal thy 
+Goal 
 "[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\
 \  !f. monofun(MF2(f)::('b::po=>'c::cpo));\
 \       chain(F);chain(Y)|] ==> \
 \       chain(%i. lub(range(%j. MF2 (F j) (Y i))))";
-by (cut_facts_tac prems 1);
-by (rtac (lub_mono RS allI RS chainI) 1);
+by (rtac (lub_mono RS chainI) 1);
 by (etac ch2ch_MF2L 1);
 by (atac 1);
 by (etac ch2ch_MF2L 1);
 by (atac 1);
 by (strip_tac 1);
-by (rtac (chainE RS spec) 1);
+by (rtac (chainE) 1);
 by ((rtac ch2ch_MF2R 1) THEN (etac spec 1));
 by (atac 1);
 qed "ch2ch_lubMF2L";
 
 
-val prems = goal thy 
+Goal 
 "[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\
 \  !f. monofun(MF2(f)::('b::po=>'c::cpo));\
 \       chain(F)|] ==> \
 \       monofun(% x. lub(range(% j. MF2 (F j) (x))))";
-by (cut_facts_tac prems 1);
 by (rtac monofunI 1);
 by (strip_tac 1);
 by (rtac lub_mono 1);
@@ -238,13 +220,12 @@
 by (atac 1);
 qed "lub_MF2_mono";
 
-val prems = goal thy 
+Goal 
 "[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\
 \  !f. monofun(MF2(f)::('b::po=>'c::cpo));\
 \       chain(F); chain(Y)|] ==> \
 \               lub(range(%j. lub(range(%i. MF2(F j) (Y i))))) =\
 \               lub(range(%i. lub(range(%j. MF2(F j) (Y i)))))";
-by (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);
@@ -275,13 +256,12 @@
 qed "ex_lubMF2";
 
 
-val prems = goal thy 
+Goal 
 "[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\
 \  !f. monofun(MF2(f)::('b::po=>'c::cpo));\
 \  chain(FY);chain(TY)|] ==>\
 \ lub(range(%i. lub(range(%j. MF2(FY(j))(TY(i)))))) =\
 \ lub(range(%i. MF2(FY(i))(TY(i))))";
-by (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);
@@ -295,7 +275,7 @@
 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 (rtac (chain_mono) 1);
 by (etac allE 1);
 by (etac ch2ch_MF2R 1);
 by (REPEAT (atac 1));
@@ -303,7 +283,7 @@
 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 (rtac (chain_mono) 1);
 by (etac ch2ch_MF2L 1);
 by (REPEAT (atac 1));
 by (rtac lub_mono 1);
@@ -317,13 +297,12 @@
 by (atac 1);
 qed "diag_lubMF2_1";
 
-val prems = goal thy 
+Goal 
 "[|monofun(MF2::('a::po=>'b::po=>'c::cpo));\
 \  !f. monofun(MF2(f)::('b::po=>'c::cpo));\
 \  chain(FY);chain(TY)|] ==>\
 \ lub(range(%j. lub(range(%i. MF2(FY(j))(TY(i)))))) =\
 \ lub(range(%i. MF2(FY(i))(TY(i))))";
-by (cut_facts_tac prems 1);
 by (rtac trans 1);
 by (rtac ex_lubMF2 1);
 by (REPEAT (atac 1));
@@ -337,46 +316,35 @@
 (* in both arguments                                                        *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goal thy 
-"[|cont(CF2);!f. cont(CF2(f));chain(FY);chain(TY)|] ==>\
+val [prem1,prem2,prem3,prem4] = goal thy 
+"[| cont(CF2); !f. cont(CF2(f)); chain(FY); chain(TY)|] ==>\
 \ CF2(lub(range(FY)))(lub(range(TY))) = lub(range(%i. CF2(FY(i))(TY(i))))";
-by (cut_facts_tac prems 1);
-by (stac ((hd prems) RS cont2contlub RS contlubE RS spec RS mp) 1);
-by (atac 1);
+by (cut_facts_tac [prem1,prem2,prem3, prem4] 1);
+by (stac (prem1 RS cont2contlub RS contlubE RS spec RS mp) 1);
+by (assume_tac 1);
 by (stac thelub_fun 1);
-by (rtac ((hd prems) RS cont2mono RS ch2ch_monofun) 1);
-by (atac 1);
+by (rtac (prem1 RS (cont2mono RS ch2ch_monofun)) 1);
+by (assume_tac 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));
+by (rtac ((prem2 RS spec RS cont2contlub) RS contlubE RS spec RS mp RS ext RS arg_cong RS arg_cong) 1);
+by (rtac prem4 1);
+by (blast_tac (claset() addIs [diag_lubMF2_2, cont2mono]) 1);
 qed "contlub_CF2";
 
 (* ------------------------------------------------------------------------ *)
 (* The following results are about application for functions in 'a=>'b      *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goal thy 
-        "f1 << f2 ==> f1(x) << f2(x)";
-by (cut_facts_tac prems 1);
+Goal "f1 << f2 ==> f1(x) << f2(x)";
 by (etac (less_fun RS iffD1 RS spec) 1);
 qed "monofun_fun_fun";
 
-val prems = goal thy 
-        "[|monofun(f); x1 << x2|] ==> f(x1) << f(x2)";
-by (cut_facts_tac prems 1);
+Goal "[|monofun(f); x1 << x2|] ==> f(x1) << f(x2)";
 by (etac (monofunE RS spec RS spec RS mp) 1);
 by (atac 1);
 qed "monofun_fun_arg";
 
-val prems = goal thy 
-"[|monofun(f1); monofun(f2); f1 << f2; x1 << x2|] ==> f1(x1) << f2(x2)";
-by (cut_facts_tac prems 1);
+Goal "[|monofun(f1); monofun(f2); f1 << f2; x1 << x2|] ==> f1(x1) << f2(x2)";
 by (rtac trans_less 1);
 by (etac monofun_fun_arg 1);
 by (atac 1);
@@ -389,23 +357,20 @@
 (* continuity                                                               *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goal thy 
-        "[|monofun(c1)|] ==> monofun(%x. c1 x y)";
-by (cut_facts_tac prems 1);
+Goal "[|monofun(c1)|] ==> monofun(%x. c1 x y)";
 by (rtac monofunI 1);
 by (strip_tac 1);
 by (etac (monofun_fun_arg RS monofun_fun_fun) 1);
 by (atac 1);
 qed "mono2mono_MF1L";
 
-val prems = goal thy 
-        "[|cont(c1)|] ==> cont(%x. c1 x y)";
-by (cut_facts_tac prems 1);
+Goal "[|cont(c1)|] ==> cont(%x. c1 x y)";
 by (rtac monocontlub2cont 1);
 by (etac (cont2mono RS mono2mono_MF1L) 1);
 by (rtac contlubI 1);
 by (strip_tac 1);
-by (rtac ((hd prems) RS cont2contlub RS contlubE RS spec RS mp RS ssubst) 1);
+by (ftac asm_rl 1);
+by (etac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1);
 by (atac 1);
 by (stac thelub_fun 1);
 by (rtac ch2ch_monofun 1);
@@ -416,20 +381,14 @@
 
 (*********  Note "(%x.%y.c1 x y) = c1" ***********)
 
-val prems = goal thy
-        "!y. monofun(%x. c1 x y) ==> monofun(c1)";
-by (cut_facts_tac prems 1);
+Goal "!y. monofun(%x. c1 x y) ==> monofun(c1)";
 by (rtac monofunI 1);
 by (strip_tac 1);
 by (rtac (less_fun RS iffD2) 1);
-by (strip_tac 1);
-by (rtac ((hd prems) RS spec RS monofunE RS spec RS spec RS mp) 1);
-by (atac 1);
+by (blast_tac (claset() addDs [monofunE]) 1);
 qed "mono2mono_MF1L_rev";
 
-val prems = goal thy
-        "!y. cont(%x. c1 x y) ==> cont(c1)";
-by (cut_facts_tac prems 1);
+Goal "!y. cont(%x. c1 x y) ==> cont(c1)";
 by (rtac monocontlub2cont 1);
 by (rtac (cont2mono RS allI RS mono2mono_MF1L_rev ) 1);
 by (etac spec 1);
@@ -440,8 +399,7 @@
 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);
+by (blast_tac (claset() addDs [ cont2contlub RS contlubE]) 1);
 qed "cont2cont_CF1L_rev";
 
 (* ------------------------------------------------------------------------ *)
@@ -449,10 +407,9 @@
 (* never used here                                                          *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goal thy
+Goal
 "[|chain(Y::nat=>'a);!y. cont(%x.(c::'a::cpo=>'b::cpo=>'c::cpo) x y)|] ==>\
 \ (%y. lub(range(%i. c (Y i) y))) = (lub(range(%i.%y. c (Y i) y)))";
-by (cut_facts_tac prems 1);
 by (rtac trans 1);
 by (rtac (cont2contlub RS contlubE RS spec RS mp) 2);
 by (atac 3);
@@ -463,10 +420,8 @@
 by (atac 1);
 qed "contlub_abstraction";
 
-val prems = goal thy 
-"[|monofun(ft);!x. monofun(ft(x));monofun(tt)|] ==>\
+Goal "[|monofun(ft);!x. monofun(ft(x));monofun(tt)|] ==>\
 \        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);
@@ -479,9 +434,7 @@
 qed "mono2mono_app";
 
 
-val prems = goal thy 
-"[|cont(ft);!x. cont(ft(x));cont(tt)|] ==> contlub(%x.(ft(x))(tt(x)))";
-by (cut_facts_tac prems 1);
+Goal "[|cont(ft);!x. cont(ft(x));cont(tt)|] ==> contlub(%x.(ft(x))(tt(x)))";
 by (rtac contlubI 1);
 by (strip_tac 1);
 by (res_inst_tac [("f3","tt")] (contlubE RS spec RS mp RS ssubst) 1);
@@ -509,7 +462,7 @@
 (* The identity function is continuous                                      *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goal thy "cont(% x. x)";
+Goal "cont(% x. x)";
 by (rtac contI 1);
 by (strip_tac 1);
 by (etac thelubE 1);
@@ -520,16 +473,9 @@
 (* constant functions are continuous                                        *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goalw thy [cont] "cont(%x. c)";
+Goalw [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);
+by (blast_tac (claset() addIs [is_lubI, ub_rangeI] addDs [ub_rangeD]) 1);
 qed "cont_const";
 
 
@@ -541,7 +487,7 @@
 (* A non-emptyness result for Cfun                                          *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goal thy "?x:Collect cont";
+Goal "?x:Collect cont";
 by (rtac CollectI 1);
 by (rtac cont_const 1);
 qed "CfunI";
@@ -550,9 +496,7 @@
 (* some properties of flat                                                  *)
 (* ------------------------------------------------------------------------ *)
 
-val prems = goalw thy [monofun]
-  "f UU = UU ==> monofun (f::'a::flat=>'b::pcpo)";
-by (cut_facts_tac prems 1);
+Goalw [monofun] "f UU = UU ==> monofun (f::'a::flat=>'b::pcpo)";
 by (strip_tac 1);
 by (dtac (ax_flat RS spec RS spec RS mp) 1);
 by (fast_tac ((HOL_cs addss (simpset() addsimps [minimal]))) 1);