# HG changeset patch # User clasohm # Date 823005777 -3600 # Node ID 6bcb44e4d6e5679d560de313a7e8084bf182c4c9 # Parent 5a6f2aabd53846b935a29b9ae1e0cffbbc9719c1 expanded tabs diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Cfun1.ML --- a/src/HOLCF/Cfun1.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Cfun1.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/cfun1.ML +(* Title: HOLCF/cfun1.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for cfun1.thy @@ -14,10 +14,10 @@ qed_goalw "CfunI" Cfun1.thy [Cfun_def] "(% x.x):Cfun" (fn prems => - [ - (rtac (mem_Collect_eq RS ssubst) 1), - (rtac cont_id 1) - ]); + [ + (rtac (mem_Collect_eq RS ssubst) 1), + (rtac cont_id 1) + ]); (* ------------------------------------------------------------------------ *) @@ -26,60 +26,60 @@ qed_goalw "refl_less_cfun" Cfun1.thy [less_cfun_def] "less_cfun f f" (fn prems => - [ - (rtac refl_less 1) - ]); + [ + (rtac refl_less 1) + ]); qed_goalw "antisym_less_cfun" Cfun1.thy [less_cfun_def] - "[|less_cfun f1 f2; less_cfun f2 f1|] ==> f1 = f2" + "[|less_cfun f1 f2; less_cfun 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) - ]); + [ + (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) + ]); qed_goalw "trans_less_cfun" Cfun1.thy [less_cfun_def] - "[|less_cfun f1 f2; less_cfun f2 f3|] ==> less_cfun f1 f3" + "[|less_cfun f1 f2; less_cfun f2 f3|] ==> less_cfun f1 f3" (fn prems => - [ - (cut_facts_tac prems 1), - (etac trans_less 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac trans_less 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* lemmas about application of continuous functions *) (* ------------------------------------------------------------------------ *) qed_goal "cfun_cong" Cfun1.thy - "[| f=g; x=y |] ==> f`x = g`y" + "[| f=g; x=y |] ==> f`x = g`y" (fn prems => - [ - (cut_facts_tac prems 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (fast_tac HOL_cs 1) + ]); qed_goal "cfun_fun_cong" Cfun1.thy "f=g ==> f`x = g`x" (fn prems => - [ - (cut_facts_tac prems 1), - (etac cfun_cong 1), - (rtac refl 1) - ]); + [ + (cut_facts_tac prems 1), + (etac cfun_cong 1), + (rtac refl 1) + ]); qed_goal "cfun_arg_cong" Cfun1.thy "x=y ==> f`x = f`y" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac cfun_cong 1), - (rtac refl 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac cfun_cong 1), + (rtac refl 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) @@ -88,37 +88,37 @@ qed_goal "Abs_Cfun_inverse2" Cfun1.thy "cont(f) ==> fapp(fabs(f)) = f" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac Abs_Cfun_inverse 1), - (rewrite_goals_tac [Cfun_def]), - (etac (mem_Collect_eq RS ssubst) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac Abs_Cfun_inverse 1), + (rewtac Cfun_def), + (etac (mem_Collect_eq RS ssubst) 1) + ]); (* ------------------------------------------------------------------------ *) (* simplification of application *) (* ------------------------------------------------------------------------ *) qed_goal "Cfunapp2" Cfun1.thy - "cont(f) ==> (fabs f)`x = f x" + "cont(f) ==> (fabs f)`x = f x" (fn prems => - [ - (cut_facts_tac prems 1), - (etac (Abs_Cfun_inverse2 RS fun_cong) 1) - ]); + [ + (cut_facts_tac prems 1), + (etac (Abs_Cfun_inverse2 RS fun_cong) 1) + ]); (* ------------------------------------------------------------------------ *) (* beta - equality for continuous functions *) (* ------------------------------------------------------------------------ *) qed_goal "beta_cfun" Cfun1.thy - "cont(c1) ==> (LAM x .c1 x)`u = c1 u" + "cont(c1) ==> (LAM x .c1 x)`u = c1 u" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac Cfunapp2 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac Cfunapp2 1), + (atac 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Cfun2.ML --- a/src/HOLCF/Cfun2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Cfun2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/cfun2.thy +(* Title: HOLCF/cfun2.thy ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for cfun2.thy @@ -14,11 +14,11 @@ qed_goal "less_cfun" Cfun2.thy "( f1 << f2 ) = (fapp(f1) << fapp(f2))" (fn prems => - [ - (rtac (inst_cfun_po RS ssubst) 1), - (fold_goals_tac [less_cfun_def]), - (rtac refl 1) - ]); + [ + (rtac (inst_cfun_po RS ssubst) 1), + (fold_goals_tac [less_cfun_def]), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) (* Type 'a ->'b is pointed *) @@ -26,13 +26,13 @@ qed_goalw "minimal_cfun" Cfun2.thy [UU_cfun_def] "UU_cfun << f" (fn prems => - [ - (rtac (less_cfun RS ssubst) 1), - (rtac (Abs_Cfun_inverse2 RS ssubst) 1), - (rtac cont_const 1), - (fold_goals_tac [UU_fun_def]), - (rtac minimal_fun 1) - ]); + [ + (rtac (less_cfun RS ssubst) 1), + (rtac (Abs_Cfun_inverse2 RS ssubst) 1), + (rtac cont_const 1), + (fold_goals_tac [UU_fun_def]), + (rtac minimal_fun 1) + ]); (* ------------------------------------------------------------------------ *) (* fapp yields continuous functions in 'a => 'b *) @@ -42,11 +42,11 @@ qed_goal "cont_fapp2" Cfun2.thy "cont(fapp(fo))" (fn prems => - [ - (res_inst_tac [("P","cont")] CollectD 1), - (fold_goals_tac [Cfun_def]), - (rtac Rep_Cfun 1) - ]); + [ + (res_inst_tac [("P","cont")] CollectD 1), + (fold_goals_tac [Cfun_def]), + (rtac Rep_Cfun 1) + ]); val monofun_fapp2 = cont_fapp2 RS cont2mono; (* monofun(fapp(?fo1)) *) @@ -73,10 +73,10 @@ qed_goalw "monofun_fapp1" Cfun2.thy [monofun] "monofun(fapp)" (fn prems => - [ - (strip_tac 1), - (etac (less_cfun RS subst) 1) - ]); + [ + (strip_tac 1), + (etac (less_cfun RS subst) 1) + ]); (* ------------------------------------------------------------------------ *) @@ -85,12 +85,12 @@ qed_goal "monofun_cfun_fun" Cfun2.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_fapp1 RS monofunE RS spec RS spec RS mp) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("x","x")] spec 1), + (rtac (less_fun RS subst) 1), + (etac (monofun_fapp1 RS monofunE RS spec RS spec RS mp) 1) + ]); val monofun_cfun_arg = (monofun_fapp2 RS monofunE RS spec RS spec RS mp); @@ -101,14 +101,14 @@ (* ------------------------------------------------------------------------ *) qed_goal "monofun_cfun" Cfun2.thy - "[|f1< f1`x1 << f2`x2" + "[|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) - ]); + [ + (cut_facts_tac prems 1), + (rtac trans_less 1), + (etac monofun_cfun_arg 1), + (etac monofun_cfun_fun 1) + ]); (* ------------------------------------------------------------------------ *) @@ -119,10 +119,10 @@ qed_goal "ch2ch_fappR" Cfun2.thy "is_chain(Y) ==> is_chain(%i. f`(Y i))" (fn prems => - [ - (cut_facts_tac prems 1), - (etac (monofun_fapp2 RS ch2ch_MF2R) 1) - ]); + [ + (cut_facts_tac prems 1), + (etac (monofun_fapp2 RS ch2ch_MF2R) 1) + ]); val ch2ch_fappL = (monofun_fapp1 RS ch2ch_MF2L); @@ -135,15 +135,15 @@ (* ------------------------------------------------------------------------ *) qed_goal "lub_cfun_mono" Cfun2.thy - "is_chain(F) ==> monofun(% x.lub(range(% j.(F j)`x)))" + "is_chain(F) ==> monofun(% x.lub(range(% j.(F j)`x)))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac lub_MF2_mono 1), - (rtac monofun_fapp1 1), - (rtac (monofun_fapp2 RS allI) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac lub_MF2_mono 1), + (rtac monofun_fapp1 1), + (rtac (monofun_fapp2 RS allI) 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* a lemma about the exchange of lubs for type 'a -> 'b *) @@ -151,37 +151,37 @@ (* ------------------------------------------------------------------------ *) qed_goal "ex_lubcfun" Cfun2.thy - "[| is_chain(F); is_chain(Y) |] ==>\ -\ lub(range(%j. lub(range(%i. F(j)`(Y i))))) =\ -\ lub(range(%i. lub(range(%j. F(j)`(Y i)))))" + "[| is_chain(F); is_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_fapp1 1), - (rtac (monofun_fapp2 RS allI) 1), - (atac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac ex_lubMF2 1), + (rtac monofun_fapp1 1), + (rtac (monofun_fapp2 RS allI) 1), + (atac 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* the lub of a chain of cont. functions is continuous *) (* ------------------------------------------------------------------------ *) qed_goal "cont_lubcfun" Cfun2.thy - "is_chain(F) ==> cont(% x.lub(range(% j.F(j)`x)))" + "is_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), - (rtac (contlub_cfun_arg RS ext RS ssubst) 1), - (atac 1), - (etac ex_lubcfun 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac monocontlub2cont 1), + (etac lub_cfun_mono 1), + (rtac contlubI 1), + (strip_tac 1), + (rtac (contlub_cfun_arg RS ext RS ssubst) 1), + (atac 1), + (etac ex_lubcfun 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* type 'a -> 'b is chain complete *) @@ -190,25 +190,25 @@ qed_goal "lub_cfun" Cfun2.thy "is_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), - (rtac (less_cfun RS ssubst) 1), - (rtac (Abs_Cfun_inverse2 RS ssubst) 1), - (etac cont_lubcfun 1), - (rtac (lub_fun RS is_lubE RS conjunct1 RS ub_rangeE RS spec) 1), - (etac (monofun_fapp1 RS ch2ch_monofun) 1), - (strip_tac 1), - (rtac (less_cfun RS ssubst) 1), - (rtac (Abs_Cfun_inverse2 RS ssubst) 1), - (etac cont_lubcfun 1), - (rtac (lub_fun RS is_lubE RS conjunct2 RS spec RS mp) 1), - (etac (monofun_fapp1 RS ch2ch_monofun) 1), - (etac (monofun_fapp1 RS ub2ub_monofun) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_lubI 1), + (rtac conjI 1), + (rtac ub_rangeI 1), + (rtac allI 1), + (rtac (less_cfun RS ssubst) 1), + (rtac (Abs_Cfun_inverse2 RS ssubst) 1), + (etac cont_lubcfun 1), + (rtac (lub_fun RS is_lubE RS conjunct1 RS ub_rangeE RS spec) 1), + (etac (monofun_fapp1 RS ch2ch_monofun) 1), + (strip_tac 1), + (rtac (less_cfun RS ssubst) 1), + (rtac (Abs_Cfun_inverse2 RS ssubst) 1), + (etac cont_lubcfun 1), + (rtac (lub_fun RS is_lubE RS conjunct2 RS spec RS mp) 1), + (etac (monofun_fapp1 RS ch2ch_monofun) 1), + (etac (monofun_fapp1 RS ub2ub_monofun) 1) + ]); val thelub_cfun = (lub_cfun RS thelubI); (* @@ -218,11 +218,11 @@ qed_goal "cpo_cfun" Cfun2.thy "is_chain(CCF::nat=>('a::pcpo->'b::pcpo)) ==> ? x. range(CCF) <<| x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac exI 1), - (etac lub_cfun 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac exI 1), + (etac lub_cfun 1) + ]); (* ------------------------------------------------------------------------ *) @@ -231,29 +231,29 @@ 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","fabs")] arg_cong 1), - (rtac ext 1), - (resolve_tac prems 1) - ]); + [ + (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","fabs")] arg_cong 1), + (rtac ext 1), + (resolve_tac prems 1) + ]); (* ------------------------------------------------------------------------ *) (* Monotonicity of fabs *) (* ------------------------------------------------------------------------ *) qed_goal "semi_monofun_fabs" Cfun2.thy - "[|cont(f);cont(g);f<fabs(f)<fabs(f)< - [ - (rtac (less_cfun RS iffD2) 1), - (rtac (Abs_Cfun_inverse2 RS ssubst) 1), - (resolve_tac prems 1), - (rtac (Abs_Cfun_inverse2 RS ssubst) 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); + [ + (rtac (less_cfun RS iffD2) 1), + (rtac (Abs_Cfun_inverse2 RS ssubst) 1), + (resolve_tac prems 1), + (rtac (Abs_Cfun_inverse2 RS ssubst) 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); (* ------------------------------------------------------------------------ *) (* Extenionality wrt. << in 'a -> 'b *) @@ -261,15 +261,15 @@ qed_goal "less_cfun2" 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_fabs 1), - (rtac cont_fapp2 1), - (rtac cont_fapp2 1), - (rtac (less_fun RS iffD2) 1), - (rtac allI 1), - (resolve_tac prems 1) - ]); + [ + (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_fabs 1), + (rtac cont_fapp2 1), + (rtac cont_fapp2 1), + (rtac (less_fun RS iffD2) 1), + (rtac allI 1), + (resolve_tac prems 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Cfun3.ML --- a/src/HOLCF/Cfun3.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Cfun3.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/cfun3.ML +(* Title: HOLCF/cfun3.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen *) @@ -12,19 +12,19 @@ qed_goal "contlub_fapp1" Cfun3.thy "contlub(fapp)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (expand_fun_eq RS iffD2) 1), - (strip_tac 1), - (rtac (thelub_cfun RS ssubst) 1), - (atac 1), - (rtac (Cfunapp2 RS ssubst) 1), - (etac cont_lubcfun 1), - (rtac (thelub_fun RS ssubst) 1), - (etac (monofun_fapp1 RS ch2ch_monofun) 1), - (rtac refl 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac (expand_fun_eq RS iffD2) 1), + (strip_tac 1), + (rtac (thelub_cfun RS ssubst) 1), + (atac 1), + (rtac (Cfunapp2 RS ssubst) 1), + (etac cont_lubcfun 1), + (rtac (thelub_fun RS ssubst) 1), + (etac (monofun_fapp1 RS ch2ch_monofun) 1), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) @@ -33,11 +33,11 @@ qed_goal "cont_fapp1" Cfun3.thy "cont(fapp)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_fapp1 1), - (rtac contlub_fapp1 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_fapp1 1), + (rtac contlub_fapp1 1) + ]); (* ------------------------------------------------------------------------ *) @@ -48,26 +48,26 @@ "is_chain(FY) ==>\ \ lub(range FY)`x = lub(range (%i.FY(i)`x))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans 1), - (etac (contlub_fapp1 RS contlubE RS spec RS mp RS fun_cong) 1), - (rtac (thelub_fun RS ssubst) 1), - (etac (monofun_fapp1 RS ch2ch_monofun) 1), - (rtac refl 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac trans 1), + (etac (contlub_fapp1 RS contlubE RS spec RS mp RS fun_cong) 1), + (rtac (thelub_fun RS ssubst) 1), + (etac (monofun_fapp1 RS ch2ch_monofun) 1), + (rtac refl 1) + ]); qed_goal "cont_cfun_fun" Cfun3.thy "is_chain(FY) ==>\ \ range(%i.FY(i)`x) <<| lub(range FY)`x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac thelubE 1), - (etac ch2ch_fappL 1), - (etac (contlub_cfun_fun RS sym) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac thelubE 1), + (etac ch2ch_fappL 1), + (etac (contlub_cfun_fun RS sym) 1) + ]); (* ------------------------------------------------------------------------ *) @@ -78,31 +78,31 @@ "[|is_chain(FY);is_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_fapp1 1), - (rtac allI 1), - (rtac cont_fapp2 1), - (atac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac contlub_CF2 1), + (rtac cont_fapp1 1), + (rtac allI 1), + (rtac cont_fapp2 1), + (atac 1), + (atac 1) + ]); qed_goal "cont_cfun" Cfun3.thy "[|is_chain(FY);is_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_fapp1 RS ch2ch_MF2LR) 1), - (rtac allI 1), - (rtac monofun_fapp2 1), - (atac 1), - (atac 1), - (etac (contlub_cfun RS sym) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac thelubE 1), + (rtac (monofun_fapp1 RS ch2ch_MF2LR) 1), + (rtac allI 1), + (rtac monofun_fapp2 1), + (atac 1), + (atac 1), + (etac (contlub_cfun RS sym) 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) @@ -110,18 +110,18 @@ (* ------------------------------------------------------------------------ *) qed_goal "cont2cont_fapp" Cfun3.thy - "[|cont(%x.ft x);cont(%x.tt x)|] ==> cont(%x. (ft x)`(tt x))" + "[|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_fapp1 1), - (atac 1), - (rtac cont_fapp2 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac cont2cont_app2 1), + (rtac cont2cont_app2 1), + (rtac cont_const 1), + (rtac cont_fapp1 1), + (atac 1), + (rtac cont_fapp2 1), + (atac 1) + ]); @@ -131,21 +131,21 @@ qed_goal "cont2mono_LAM" Cfun3.thy "[|!x.cont(c1 x); !y.monofun(%x.c1 x y)|] ==>\ -\ monofun(%x. LAM y. c1 x y)" +\ monofun(%x. LAM y. c1 x y)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac monofunI 1), - (strip_tac 1), - (rtac (less_cfun RS ssubst) 1), - (rtac (less_fun RS ssubst) 1), - (rtac allI 1), - (rtac (beta_cfun RS ssubst) 1), - (etac spec 1), - (rtac (beta_cfun RS ssubst) 1), - (etac spec 1), - (etac ((hd (tl prems)) RS spec RS monofunE RS spec RS spec RS mp) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac monofunI 1), + (strip_tac 1), + (rtac (less_cfun RS ssubst) 1), + (rtac (less_fun RS ssubst) 1), + (rtac allI 1), + (rtac (beta_cfun RS ssubst) 1), + (etac spec 1), + (rtac (beta_cfun RS ssubst) 1), + (etac spec 1), + (etac ((hd (tl prems)) RS spec RS monofunE RS spec RS spec RS mp) 1) + ]); (* ------------------------------------------------------------------------ *) (* cont2cont Lemma for %x. LAM y. c1 x y) *) @@ -154,29 +154,29 @@ qed_goal "cont2cont_LAM" Cfun3.thy "[| !x.cont(c1 x); !y.cont(%x.c1 x y) |] ==> cont(%x. LAM y. c1 x y)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac monocontlub2cont 1), - (etac cont2mono_LAM 1), - (rtac (cont2mono RS allI) 1), - (etac spec 1), - (rtac contlubI 1), - (strip_tac 1), - (rtac (thelub_cfun RS ssubst) 1), - (rtac (cont2mono_LAM RS ch2ch_monofun) 1), - (atac 1), - (rtac (cont2mono RS allI) 1), - (etac spec 1), - (atac 1), - (res_inst_tac [("f","fabs")] arg_cong 1), - (rtac ext 1), - (rtac (beta_cfun RS ext RS ssubst) 1), - (etac spec 1), - (rtac (cont2contlub RS contlubE - RS spec RS mp ) 1), - (etac spec 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac monocontlub2cont 1), + (etac cont2mono_LAM 1), + (rtac (cont2mono RS allI) 1), + (etac spec 1), + (rtac contlubI 1), + (strip_tac 1), + (rtac (thelub_cfun RS ssubst) 1), + (rtac (cont2mono_LAM RS ch2ch_monofun) 1), + (atac 1), + (rtac (cont2mono RS allI) 1), + (etac spec 1), + (atac 1), + (res_inst_tac [("f","fabs")] arg_cong 1), + (rtac ext 1), + (rtac (beta_cfun RS ext RS ssubst) 1), + (etac spec 1), + (rtac (cont2contlub RS contlubE + RS spec RS mp ) 1), + (etac spec 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* elimination of quantifier in premisses of cont2cont_LAM yields good *) @@ -194,7 +194,7 @@ (* ------------------------------------------------------------------------ *) val cont_lemmas = [cont_const, cont_id, cont_fapp2, - cont2cont_fapp,cont2cont_LAM2]; + cont2cont_fapp,cont2cont_LAM2]; val cont_tac = (fn i => (resolve_tac cont_lemmas i)); @@ -207,13 +207,13 @@ qed_goal "strict_fapp1" Cfun3.thy "(UU::'a->'b)`x = (UU::'b)" (fn prems => - [ - (rtac (inst_cfun_pcpo RS ssubst) 1), - (rewrite_goals_tac [UU_cfun_def]), - (rtac (beta_cfun RS ssubst) 1), - (cont_tac 1), - (rtac refl 1) - ]); + [ + (rtac (inst_cfun_pcpo RS ssubst) 1), + (rewtac UU_cfun_def), + (rtac (beta_cfun RS ssubst) 1), + (cont_tac 1), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) @@ -221,163 +221,163 @@ (* ------------------------------------------------------------------------ *) qed_goalw "Istrictify1" Cfun3.thy [Istrictify_def] - "Istrictify(f)(UU)= (UU)" + "Istrictify(f)(UU)= (UU)" (fn prems => - [ - (Simp_tac 1) - ]); + [ + (Simp_tac 1) + ]); qed_goalw "Istrictify2" Cfun3.thy [Istrictify_def] - "~x=UU ==> Istrictify(f)(x)=f`x" + "~x=UU ==> Istrictify(f)(x)=f`x" (fn prems => - [ - (cut_facts_tac prems 1), - (Asm_simp_tac 1) - ]); + [ + (cut_facts_tac prems 1), + (Asm_simp_tac 1) + ]); qed_goal "monofun_Istrictify1" Cfun3.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), - (rtac (Istrictify2 RS ssubst) 1), - (atac 1), - (rtac (Istrictify2 RS ssubst) 1), - (atac 1), - (rtac monofun_cfun_fun 1), - (atac 1), - (hyp_subst_tac 1), - (rtac (Istrictify1 RS ssubst) 1), - (rtac (Istrictify1 RS ssubst) 1), - (rtac refl_less 1) - ]); + [ + (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), + (rtac (Istrictify2 RS ssubst) 1), + (atac 1), + (rtac (Istrictify2 RS ssubst) 1), + (atac 1), + (rtac monofun_cfun_fun 1), + (atac 1), + (hyp_subst_tac 1), + (rtac (Istrictify1 RS ssubst) 1), + (rtac (Istrictify1 RS ssubst) 1), + (rtac refl_less 1) + ]); qed_goal "monofun_Istrictify2" Cfun3.thy "monofun(Istrictify(f))" (fn prems => - [ - (rtac monofunI 1), - (strip_tac 1), - (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), - (rtac (Istrictify2 RS ssubst) 1), - (etac notUU_I 1), - (atac 1), - (rtac (Istrictify2 RS ssubst) 1), - (atac 1), - (rtac monofun_cfun_arg 1), - (atac 1), - (hyp_subst_tac 1), - (rtac (Istrictify1 RS ssubst) 1), - (rtac minimal 1) - ]); + [ + (rtac monofunI 1), + (strip_tac 1), + (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), + (rtac (Istrictify2 RS ssubst) 1), + (etac notUU_I 1), + (atac 1), + (rtac (Istrictify2 RS ssubst) 1), + (atac 1), + (rtac monofun_cfun_arg 1), + (atac 1), + (hyp_subst_tac 1), + (rtac (Istrictify1 RS ssubst) 1), + (rtac minimal 1) + ]); qed_goal "contlub_Istrictify1" Cfun3.thy "contlub(Istrictify)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (expand_fun_eq RS iffD2) 1), - (strip_tac 1), - (rtac (thelub_fun RS ssubst) 1), - (etac (monofun_Istrictify1 RS ch2ch_monofun) 1), - (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), - (rtac (Istrictify2 RS ssubst) 1), - (atac 1), - (rtac (Istrictify2 RS ext RS ssubst) 1), - (atac 1), - (rtac (thelub_cfun RS ssubst) 1), - (atac 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_lubcfun 1), - (atac 1), - (rtac refl 1), - (hyp_subst_tac 1), - (rtac (Istrictify1 RS ssubst) 1), - (rtac (Istrictify1 RS ext RS ssubst) 1), - (rtac (chain_UU_I_inverse RS sym) 1), - (rtac (refl RS allI) 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac (expand_fun_eq RS iffD2) 1), + (strip_tac 1), + (rtac (thelub_fun RS ssubst) 1), + (etac (monofun_Istrictify1 RS ch2ch_monofun) 1), + (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1), + (rtac (Istrictify2 RS ssubst) 1), + (atac 1), + (rtac (Istrictify2 RS ext RS ssubst) 1), + (atac 1), + (rtac (thelub_cfun RS ssubst) 1), + (atac 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_lubcfun 1), + (atac 1), + (rtac refl 1), + (hyp_subst_tac 1), + (rtac (Istrictify1 RS ssubst) 1), + (rtac (Istrictify1 RS ext RS ssubst) 1), + (rtac (chain_UU_I_inverse RS sym) 1), + (rtac (refl RS allI) 1) + ]); qed_goal "contlub_Istrictify2" Cfun3.thy "contlub(Istrictify(f::'a -> 'b))" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (res_inst_tac [("Q","lub(range(Y))=(UU::'a)")] classical2 1), - (res_inst_tac [("t","lub(range(Y))")] subst 1), - (rtac sym 1), - (atac 1), - (rtac (Istrictify1 RS ssubst) 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), - (rtac (Istrictify2 RS ssubst) 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_fapp2 1), - (atac 1), - (rtac ch2ch_monofun 1), - (rtac monofun_Istrictify2 1), - (atac 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (res_inst_tac [("Q","lub(range(Y))=(UU::'a)")] classical2 1), + (res_inst_tac [("t","lub(range(Y))")] subst 1), + (rtac sym 1), + (atac 1), + (rtac (Istrictify1 RS ssubst) 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), + (rtac (Istrictify2 RS ssubst) 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_fapp2 1), + (atac 1), + (rtac ch2ch_monofun 1), + (rtac monofun_Istrictify2 1), + (atac 1) + ]); val cont_Istrictify1 = (contlub_Istrictify1 RS - (monofun_Istrictify1 RS monocontlub2cont)); + (monofun_Istrictify1 RS monocontlub2cont)); val cont_Istrictify2 = (contlub_Istrictify2 RS - (monofun_Istrictify2 RS monocontlub2cont)); + (monofun_Istrictify2 RS monocontlub2cont)); qed_goalw "strictify1" Cfun3.thy [strictify_def] - "strictify`f`UU=UU" + "strictify`f`UU=UU" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (cont_tac 1), - (rtac cont_Istrictify2 1), - (rtac cont2cont_CF1L 1), - (rtac cont_Istrictify1 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Istrictify2 1), - (rtac Istrictify1 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (cont_tac 1), + (rtac cont_Istrictify2 1), + (rtac cont2cont_CF1L 1), + (rtac cont_Istrictify1 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Istrictify2 1), + (rtac Istrictify1 1) + ]); qed_goalw "strictify2" Cfun3.thy [strictify_def] - "~x=UU ==> strictify`f`x=f`x" + "~x=UU ==> strictify`f`x=f`x" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (cont_tac 1), - (rtac cont_Istrictify2 1), - (rtac cont2cont_CF1L 1), - (rtac cont_Istrictify1 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Istrictify2 1), - (rtac Istrictify2 1), - (resolve_tac prems 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (cont_tac 1), + (rtac cont_Istrictify2 1), + (rtac cont2cont_CF1L 1), + (rtac cont_Istrictify1 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Istrictify2 1), + (rtac Istrictify2 1), + (resolve_tac prems 1) + ]); (* ------------------------------------------------------------------------ *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Cont.ML --- a/src/HOLCF/Cont.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Cont.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/cont.ML +(* Title: HOLCF/cont.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for cont.thy @@ -13,56 +13,56 @@ (* ------------------------------------------------------------------------ *) qed_goalw "contlubI" Cont.thy [contlub] - "! Y. is_chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))==>\ -\ contlub(f)" + "! Y. is_chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))==>\ +\ contlub(f)" (fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (atac 1) + ]); qed_goalw "contlubE" Cont.thy [contlub] - " contlub(f)==>\ + " contlub(f)==>\ \ ! Y. is_chain(Y) --> f(lub(range(Y))) = lub(range(%i. f(Y(i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (atac 1) + ]); qed_goalw "contI" Cont.thy [cont] "! Y. is_chain(Y) --> range(% i.f(Y(i))) <<| f(lub(range(Y))) ==> cont(f)" (fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (atac 1) + ]); qed_goalw "contE" Cont.thy [cont] "cont(f) ==> ! Y. is_chain(Y) --> range(% i.f(Y(i))) <<| f(lub(range(Y)))" (fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (atac 1) + ]); qed_goalw "monofunI" Cont.thy [monofun] - "! x y. x << y --> f(x) << f(y) ==> monofun(f)" + "! x y. x << y --> f(x) << f(y) ==> monofun(f)" (fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (atac 1) + ]); qed_goalw "monofunE" Cont.thy [monofun] - "monofun(f) ==> ! x y. x << y --> f(x) << f(y)" + "monofun(f) ==> ! x y. x << y --> f(x) << f(y)" (fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* the main purpose of cont.thy is to show: *) @@ -74,15 +74,15 @@ (* ------------------------------------------------------------------------ *) qed_goal "ch2ch_monofun" Cont.thy - "[| monofun(f); is_chain(Y) |] ==> is_chain(%i. f(Y(i)))" + "[| monofun(f); is_chain(Y) |] ==> is_chain(%i. f(Y(i)))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_chainI 1), - (rtac allI 1), - (etac (monofunE RS spec RS spec RS mp) 1), - (etac (is_chainE RS spec) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_chainI 1), + (rtac allI 1), + (etac (monofunE RS spec RS spec RS mp) 1), + (etac (is_chainE RS spec) 1) + ]); (* ------------------------------------------------------------------------ *) (* monotone functions map upper bound to upper bounds *) @@ -91,30 +91,30 @@ qed_goal "ub2ub_monofun" Cont.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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) (* left to right: monofun(f) & contlub(f) ==> cont(f) *) (* ------------------------------------------------------------------------ *) qed_goalw "monocontlub2cont" Cont.thy [cont] - "[|monofun(f);contlub(f)|] ==> cont(f)" + "[|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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) (* first a lemma about binary chains *) @@ -123,14 +123,14 @@ qed_goal "binchain_cont" Cont.thy "[| cont(f); x << y |] ==> range(%i. 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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) (* right to left: cont(f) ==> monofun(f) & contlub(f) *) @@ -138,17 +138,17 @@ (* ------------------------------------------------------------------------ *) qed_goalw "cont2mono" Cont.thy [monofun] - "cont(f) ==> monofun(f)" + "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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) (* right to left: cont(f) ==> monofun(f) & contlub(f) *) @@ -156,15 +156,15 @@ (* ------------------------------------------------------------------------ *) qed_goalw "cont2contlub" Cont.thy [contlub] - "cont(f) ==> contlub(f)" + "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) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (rtac (thelubI RS sym) 1), + (etac (contE RS spec RS mp) 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* The following results are about a curried function that is monotone *) @@ -174,135 +174,135 @@ qed_goal "ch2ch_MF2L" Cont.thy "[|monofun(MF2); is_chain(F)|] ==> is_chain(%i. MF2 (F i) x)" (fn prems => - [ - (cut_facts_tac prems 1), - (etac (ch2ch_monofun RS ch2ch_fun) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac (ch2ch_monofun RS ch2ch_fun) 1), + (atac 1) + ]); qed_goal "ch2ch_MF2R" Cont.thy "[|monofun(MF2(f)); is_chain(Y)|] ==> is_chain(%i. MF2 f (Y i))" (fn prems => - [ - (cut_facts_tac prems 1), - (etac ch2ch_monofun 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac ch2ch_monofun 1), + (atac 1) + ]); qed_goal "ch2ch_MF2LR" Cont.thy "[|monofun(MF2); !f.monofun(MF2(f)); is_chain(F); is_chain(Y)|] ==> \ \ is_chain(%i. MF2(F(i))(Y(i)))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_chainI 1), - (strip_tac 1 ), - (rtac trans_less 1), - (etac (ch2ch_MF2L RS is_chainE RS spec) 1), - (atac 1), - ((rtac (monofunE RS spec RS spec RS mp) 1) THEN (etac spec 1)), - (etac (is_chainE RS spec) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_chainI 1), + (strip_tac 1 ), + (rtac trans_less 1), + (etac (ch2ch_MF2L RS is_chainE RS spec) 1), + (atac 1), + ((rtac (monofunE RS spec RS spec RS mp) 1) THEN (etac spec 1)), + (etac (is_chainE RS spec) 1) + ]); qed_goal "ch2ch_lubMF2R" Cont.thy "[|monofun(MF2::('a::po=>'b::po=>'c::pcpo));\ \ !f.monofun(MF2(f)::('b::po=>'c::pcpo));\ -\ is_chain(F);is_chain(Y)|] ==> \ -\ is_chain(%j. lub(range(%i. MF2 (F j) (Y i))))" +\ is_chain(F);is_chain(Y)|] ==> \ +\ is_chain(%j. lub(range(%i. MF2 (F j) (Y i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (lub_mono RS allI RS is_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 (is_chainE RS spec) 1), - (etac ch2ch_MF2L 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (lub_mono RS allI RS is_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 (is_chainE RS spec) 1), + (etac ch2ch_MF2L 1), + (atac 1) + ]); qed_goal "ch2ch_lubMF2L" Cont.thy "[|monofun(MF2::('a::po=>'b::po=>'c::pcpo));\ \ !f.monofun(MF2(f)::('b::po=>'c::pcpo));\ -\ is_chain(F);is_chain(Y)|] ==> \ -\ is_chain(%i. lub(range(%j. MF2 (F j) (Y i))))" +\ is_chain(F);is_chain(Y)|] ==> \ +\ is_chain(%i. lub(range(%j. MF2 (F j) (Y i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (lub_mono RS allI RS is_chainI) 1), - (etac ch2ch_MF2L 1), - (atac 1), - (etac ch2ch_MF2L 1), - (atac 1), - (strip_tac 1), - (rtac (is_chainE RS spec) 1), - ((rtac ch2ch_MF2R 1) THEN (etac spec 1)), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (lub_mono RS allI RS is_chainI) 1), + (etac ch2ch_MF2L 1), + (atac 1), + (etac ch2ch_MF2L 1), + (atac 1), + (strip_tac 1), + (rtac (is_chainE RS spec) 1), + ((rtac ch2ch_MF2R 1) THEN (etac spec 1)), + (atac 1) + ]); qed_goal "lub_MF2_mono" Cont.thy "[|monofun(MF2::('a::po=>'b::po=>'c::pcpo));\ \ !f.monofun(MF2(f)::('b::po=>'c::pcpo));\ -\ is_chain(F)|] ==> \ -\ monofun(% x.lub(range(% j.MF2 (F j) (x))))" +\ is_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) - ]); + [ + (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) + ]); qed_goal "ex_lubMF2" Cont.thy "[|monofun(MF2::('a::po=>'b::po=>'c::pcpo));\ \ !f.monofun(MF2(f)::('b::po=>'c::pcpo));\ -\ is_chain(F); is_chain(Y)|] ==> \ -\ lub(range(%j. lub(range(%i. MF2(F j) (Y i))))) =\ -\ lub(range(%i. lub(range(%j. MF2(F j) (Y i)))))" +\ is_chain(F); is_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) - ]); + [ + (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) + ]); qed_goal "diag_lubMF2_1" Cont.thy @@ -312,42 +312,42 @@ \ 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) - ]); + [ + (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) + ]); qed_goal "diag_lubMF2_2" Cont.thy "[|monofun(MF2::('a::po=>'b::po=>'c::pcpo));\ @@ -356,14 +356,14 @@ \ 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)) - ]); + [ + (cut_facts_tac prems 1), + (rtac trans 1), + (rtac ex_lubMF2 1), + (REPEAT (atac 1)), + (etac diag_lubMF2_1 1), + (REPEAT (atac 1)) + ]); @@ -377,55 +377,55 @@ "[|cont(CF2);!f.cont(CF2(f));is_chain(FY);is_chain(TY)|] ==>\ \ CF2(lub(range(FY)))(lub(range(TY))) = lub(range(%i.CF2(FY(i))(TY(i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac ((hd prems) RS cont2contlub RS contlubE RS spec RS mp RS ssubst) 1), - (atac 1), - (rtac (thelub_fun RS ssubst) 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)) - ]); + [ + (cut_facts_tac prems 1), + (rtac ((hd prems) RS cont2contlub RS contlubE RS spec RS mp RS ssubst) 1), + (atac 1), + (rtac (thelub_fun RS ssubst) 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)) + ]); (* ------------------------------------------------------------------------ *) (* The following results are about application for functions in 'a=>'b *) (* ------------------------------------------------------------------------ *) qed_goal "monofun_fun_fun" Cont.thy - "f1 << f2 ==> f1(x) << f2(x)" + "f1 << f2 ==> f1(x) << f2(x)" (fn prems => - [ - (cut_facts_tac prems 1), - (etac (less_fun RS iffD1 RS spec) 1) - ]); + [ + (cut_facts_tac prems 1), + (etac (less_fun RS iffD1 RS spec) 1) + ]); qed_goal "monofun_fun_arg" Cont.thy - "[|monofun(f); x1 << x2|] ==> f(x1) << f(x2)" + "[|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) - ]); + [ + (cut_facts_tac prems 1), + (etac (monofunE RS spec RS spec RS mp) 1), + (atac 1) + ]); qed_goal "monofun_fun" Cont.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) - ]); + [ + (cut_facts_tac prems 1), + (rtac trans_less 1), + (etac monofun_fun_arg 1), + (atac 1), + (etac monofun_fun_fun 1) + ]); (* ------------------------------------------------------------------------ *) @@ -434,69 +434,69 @@ (* ------------------------------------------------------------------------ *) qed_goal "mono2mono_MF1L" Cont.thy - "[|monofun(c1)|] ==> monofun(%x. c1 x y)" + "[|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) - ]); + [ + (cut_facts_tac prems 1), + (rtac monofunI 1), + (strip_tac 1), + (etac (monofun_fun_arg RS monofun_fun_fun) 1), + (atac 1) + ]); qed_goal "cont2cont_CF1L" Cont.thy - "[|cont(c1)|] ==> cont(%x. c1 x y)" + "[|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), - (rtac (thelub_fun RS ssubst) 1), - (rtac ch2ch_monofun 1), - (etac cont2mono 1), - (atac 1), - (rtac refl 1) - ]); + [ + (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), + (rtac (thelub_fun RS ssubst) 1), + (rtac ch2ch_monofun 1), + (etac cont2mono 1), + (atac 1), + (rtac refl 1) + ]); (********* Note "(%x.%y.c1 x y) = c1" ***********) qed_goal "mono2mono_MF1L_rev" Cont.thy - "!y.monofun(%x.c1 x y) ==> monofun(c1)" + "!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) - ]); + [ + (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) + ]); qed_goal "cont2cont_CF1L_rev" Cont.thy - "!y.cont(%x.c1 x y) ==> cont(c1)" + "!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), - (rtac (thelub_fun RS ssubst) 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) - ]); + [ + (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), + (rtac (thelub_fun RS ssubst) 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) + ]); (* ------------------------------------------------------------------------ *) @@ -508,74 +508,74 @@ "[|is_chain(Y::nat=>'a);!y.cont(%x.(c::'a=>'b=>'c) 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) - ]); + [ + (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) + ]); qed_goal "mono2mono_app" Cont.thy "[|monofun(ft);!x.monofun(ft(x));monofun(tt)|] ==>\ -\ monofun(%x.(ft(x))(tt(x)))" +\ 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) - ]); + [ + (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) + ]); qed_goal "cont2contlub_app" Cont.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) - ]); + [ + (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) + ]); qed_goal "cont2cont_app" Cont.thy "[|cont(ft);!x.cont(ft(x));cont(tt)|] ==>\ -\ cont(%x.(ft(x))(tt(x)))" +\ 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) - ]); + [ + (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) + ]); val cont2cont_app2 = (allI RSN (2,cont2cont_app)); @@ -589,12 +589,12 @@ qed_goal "cont_id" Cont.thy "cont(% x.x)" (fn prems => - [ - (rtac contI 1), - (strip_tac 1), - (etac thelubE 1), - (rtac refl 1) - ]); + [ + (rtac contI 1), + (strip_tac 1), + (etac thelubE 1), + (rtac refl 1) + ]); @@ -604,27 +604,27 @@ qed_goalw "cont_const" Cont.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) - ]); + [ + (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) + ]); qed_goal "cont2cont_app3" Cont.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) - ]); + [ + (cut_facts_tac prems 1), + (rtac cont2cont_app2 1), + (rtac cont_const 1), + (atac 1), + (atac 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Cprod1.ML --- a/src/HOLCF/Cprod1.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Cprod1.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/cprod1.ML +(* Title: HOLCF/cprod1.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for theory cprod1.thy @@ -11,54 +11,54 @@ qed_goalw "less_cprod1b" Cprod1.thy [less_cprod_def] "less_cprod p1 p2 = ( fst(p1) << fst(p2) & snd(p1) << snd(p2))" (fn prems => - [ - (rtac refl 1) - ]); + [ + (rtac refl 1) + ]); qed_goalw "less_cprod2a" Cprod1.thy [less_cprod_def] "less_cprod (x,y) (UU,UU) ==> x = UU & y = UU" (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), - (etac UU_I 1), - (etac UU_I 1) - ]); + [ + (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), + (etac UU_I 1), + (etac UU_I 1) + ]); qed_goal "less_cprod2b" Cprod1.thy "less_cprod p (UU,UU) ==> p = (UU,UU)" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p")] PairE 1), - (hyp_subst_tac 1), - (dtac less_cprod2a 1), - (Asm_simp_tac 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","p")] PairE 1), + (hyp_subst_tac 1), + (dtac less_cprod2a 1), + (Asm_simp_tac 1) + ]); qed_goalw "less_cprod2c" Cprod1.thy [less_cprod_def] "less_cprod (x1,y1) (x2,y2) ==> 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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) (* less_cprod is a partial order on 'a * 'b *) @@ -70,44 +70,44 @@ qed_goal "antisym_less_cprod" Cprod1.thy "[|less_cprod p1 p2;less_cprod p2 p1|] ==> p1=p2" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p1")] PairE 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p2")] PairE 1), - (hyp_subst_tac 1), - (dtac less_cprod2c 1), - (dtac less_cprod2c 1), - (etac conjE 1), - (etac conjE 1), - (rtac (Pair_eq RS ssubst) 1), - (fast_tac (HOL_cs addSIs [antisym_less]) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","p1")] PairE 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p2")] PairE 1), + (hyp_subst_tac 1), + (dtac less_cprod2c 1), + (dtac less_cprod2c 1), + (etac conjE 1), + (etac conjE 1), + (rtac (Pair_eq RS ssubst) 1), + (fast_tac (HOL_cs addSIs [antisym_less]) 1) + ]); qed_goal "trans_less_cprod" Cprod1.thy "[|less_cprod p1 p2;less_cprod p2 p3|] ==> less_cprod p1 p3" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p1")] PairE 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p3")] PairE 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p2")] PairE 1), - (hyp_subst_tac 1), - (dtac less_cprod2c 1), - (dtac less_cprod2c 1), - (rtac (less_cprod1b RS ssubst) 1), - (Simp_tac 1), - (etac conjE 1), - (etac conjE 1), - (rtac conjI 1), - (etac trans_less 1), - (atac 1), - (etac trans_less 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","p1")] PairE 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p3")] PairE 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p2")] PairE 1), + (hyp_subst_tac 1), + (dtac less_cprod2c 1), + (dtac less_cprod2c 1), + (rtac (less_cprod1b RS ssubst) 1), + (Simp_tac 1), + (etac conjE 1), + (etac conjE 1), + (rtac conjI 1), + (etac trans_less 1), + (atac 1), + (etac trans_less 1), + (atac 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Cprod2.ML --- a/src/HOLCF/Cprod2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Cprod2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/cprod2.ML +(* Title: HOLCF/cprod2.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for cprod2.thy @@ -9,51 +9,51 @@ open Cprod2; qed_goal "less_cprod3a" Cprod2.thy - "p1=(UU,UU) ==> p1 << p2" + "p1=(UU,UU) ==> p1 << p2" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (inst_cprod_po RS ssubst) 1), - (rtac (less_cprod1b RS ssubst) 1), - (hyp_subst_tac 1), - (Asm_simp_tac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (inst_cprod_po RS ssubst) 1), + (rtac (less_cprod1b RS ssubst) 1), + (hyp_subst_tac 1), + (Asm_simp_tac 1) + ]); qed_goal "less_cprod3b" Cprod2.thy "(p1 << p2) = (fst(p1)< - [ - (rtac (inst_cprod_po RS ssubst) 1), - (rtac less_cprod1b 1) - ]); + [ + (rtac (inst_cprod_po RS ssubst) 1), + (rtac less_cprod1b 1) + ]); qed_goal "less_cprod4a" Cprod2.thy - "(x1,x2) << (UU,UU) ==> x1=UU & x2=UU" + "(x1,x2) << (UU,UU) ==> x1=UU & x2=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac less_cprod2a 1), - (etac (inst_cprod_po RS subst) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac less_cprod2a 1), + (etac (inst_cprod_po RS subst) 1) + ]); qed_goal "less_cprod4b" Cprod2.thy - "p << (UU,UU) ==> p = (UU,UU)" + "p << (UU,UU) ==> p = (UU,UU)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac less_cprod2b 1), - (etac (inst_cprod_po RS subst) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac less_cprod2b 1), + (etac (inst_cprod_po RS subst) 1) + ]); qed_goal "less_cprod4c" Cprod2.thy " (xa,ya) << (x,y) ==> xa< - [ - (cut_facts_tac prems 1), - (rtac less_cprod2c 1), - (etac (inst_cprod_po RS subst) 1), - (REPEAT (atac 1)) - ]); + [ + (cut_facts_tac prems 1), + (rtac less_cprod2c 1), + (etac (inst_cprod_po RS subst) 1), + (REPEAT (atac 1)) + ]); (* ------------------------------------------------------------------------ *) (* type cprod is pointed *) @@ -61,10 +61,10 @@ qed_goal "minimal_cprod" Cprod2.thy "(UU,UU)< - [ - (rtac less_cprod3a 1), - (rtac refl 1) - ]); + [ + (rtac less_cprod3a 1), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) (* Pair <_,_> is monotone in both arguments *) @@ -72,34 +72,34 @@ qed_goalw "monofun_pair1" Cprod2.thy [monofun] "monofun(Pair)" (fn prems => - [ - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (rtac (less_cprod3b RS iffD2) 1), - (Simp_tac 1) - ]); + [ + (strip_tac 1), + (rtac (less_fun RS iffD2) 1), + (strip_tac 1), + (rtac (less_cprod3b RS iffD2) 1), + (Simp_tac 1) + ]); qed_goalw "monofun_pair2" Cprod2.thy [monofun] "monofun(Pair(x))" (fn prems => - [ - (strip_tac 1), - (rtac (less_cprod3b RS iffD2) 1), - (Simp_tac 1) - ]); + [ + (strip_tac 1), + (rtac (less_cprod3b RS iffD2) 1), + (Simp_tac 1) + ]); qed_goal "monofun_pair" Cprod2.thy "[|x1< (x1,y1) << (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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) (* fst and snd are monotone *) @@ -107,27 +107,27 @@ qed_goalw "monofun_fst" Cprod2.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) - ]); + [ + (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) + ]); qed_goalw "monofun_snd" Cprod2.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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) (* the type 'a * 'b is a cpo *) @@ -137,28 +137,28 @@ " is_chain(S) ==> range(S) <<| \ \ (lub(range(%i.fst(S i))),lub(range(%i.snd(S 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 [("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) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_lubI 1), + (rtac conjI 1), + (rtac ub_rangeI 1), + (rtac allI 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 thelub_cprod = (lub_cprod RS thelubI); (* @@ -169,12 +169,12 @@ *) qed_goal "cpo_cprod" Cprod2.thy - "is_chain(S::nat=>'a*'b)==>? x.range(S)<<| x" + "is_chain(S::nat=>'a*'b)==>? x.range(S)<<| x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac exI 1), - (etac lub_cprod 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac exI 1), + (etac lub_cprod 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Cprod3.ML --- a/src/HOLCF/Cprod3.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Cprod3.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/cprod3.ML +(* Title: HOLCF/cprod3.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for Cprod3.thy @@ -17,120 +17,120 @@ \ (lub(range(Y)),(x::'b)) =\ \ (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) - ]); + [ + (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) + ]); 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), - (rtac (lub_fun RS thelubI RS ssubst) 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) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac (expand_fun_eq RS iffD2) 1), + (strip_tac 1), + (rtac (lub_fun RS thelubI RS ssubst) 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) + ]); qed_goal "Cprod3_lemma2" Cprod3.thy "is_chain(Y::(nat=>'a)) ==>\ \ ((x::'b),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) - ]); + [ + (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) + ]); 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) - ]); + [ + (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) + ]); qed_goal "cont_pair1" Cprod3.thy "cont(Pair)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_pair1 1), - (rtac contlub_pair1 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_pair1 1), + (rtac contlub_pair1 1) + ]); qed_goal "cont_pair2" Cprod3.thy "cont(Pair(x))" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_pair2 1), - (rtac contlub_pair2 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_pair2 1), + (rtac contlub_pair2 1) + ]); qed_goal "contlub_fst" Cprod3.thy "contlub(fst)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (lub_cprod RS thelubI RS ssubst) 1), - (atac 1), - (Simp_tac 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac (lub_cprod RS thelubI RS ssubst) 1), + (atac 1), + (Simp_tac 1) + ]); qed_goal "contlub_snd" Cprod3.thy "contlub(snd)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (lub_cprod RS thelubI RS ssubst) 1), - (atac 1), - (Simp_tac 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac (lub_cprod RS thelubI RS ssubst) 1), + (atac 1), + (Simp_tac 1) + ]); qed_goal "cont_fst" Cprod3.thy "cont(fst)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_fst 1), - (rtac contlub_fst 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_fst 1), + (rtac contlub_fst 1) + ]); qed_goal "cont_snd" Cprod3.thy "cont(snd)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_snd 1), - (rtac contlub_snd 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_snd 1), + (rtac contlub_snd 1) + ]); (* -------------------------------------------------------------------------- @@ -144,144 +144,144 @@ (* ------------------------------------------------------------------------ *) qed_goalw "beta_cfun_cprod" Cprod3.thy [cpair_def] - "(LAM x y.(x,y))`a`b = (a,b)" + "(LAM x y.(x,y))`a`b = (a,b)" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (cont_tac 1), - (rtac cont_pair2 1), - (rtac cont2cont_CF1L 1), - (rtac cont_pair1 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_pair2 1), - (rtac refl 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (cont_tac 1), + (rtac cont_pair2 1), + (rtac cont2cont_CF1L 1), + (rtac cont_pair1 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_pair2 1), + (rtac refl 1) + ]); qed_goalw "inject_cpair" Cprod3.thy [cpair_def] - " = ==> a=aa & b=ba" + " = ==> 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) - ]); + [ + (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) + ]); 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) - ]); + [ + (rtac sym 1), + (rtac trans 1), + (rtac beta_cfun_cprod 1), + (rtac sym 1), + (rtac inst_cprod_pcpo 1) + ]); 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) - ]); + [ + (cut_facts_tac prems 1), + (dtac (inst_cprod_pcpo2 RS subst) 1), + (etac inject_cpair 1) + ]); qed_goalw "Exh_Cprod2" Cprod3.thy [cpair_def] - "? a b. z=" + "? a b. z=" (fn prems => - [ - (rtac PairE 1), - (rtac exI 1), - (rtac exI 1), - (etac (beta_cfun_cprod RS ssubst) 1) - ]); + [ + (rtac PairE 1), + (rtac exI 1), + (rtac exI 1), + (etac (beta_cfun_cprod RS ssubst) 1) + ]); 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) - ]); + [ + (rtac PairE 1), + (resolve_tac prems 1), + (etac (beta_cfun_cprod RS ssubst) 1) + ]); qed_goalw "cfst2" Cprod3.thy [cfst_def,cpair_def] - "cfst`=x" + "cfst`=x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun_cprod RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_fst 1), - (Simp_tac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun_cprod RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_fst 1), + (Simp_tac 1) + ]); qed_goalw "csnd2" Cprod3.thy [csnd_def,cpair_def] - "csnd`=y" + "csnd`=y" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun_cprod RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_snd 1), - (Simp_tac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun_cprod RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_snd 1), + (Simp_tac 1) + ]); qed_goalw "surjective_pairing_Cprod2" Cprod3.thy - [cfst_def,csnd_def,cpair_def] " = p" + [cfst_def,csnd_def,cpair_def] " = p" (fn prems => - [ - (rtac (beta_cfun_cprod RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_snd 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_fst 1), - (rtac (surjective_pairing RS sym) 1) - ]); + [ + (rtac (beta_cfun_cprod RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_snd 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_fst 1), + (rtac (surjective_pairing RS sym) 1) + ]); qed_goalw "less_cprod5b" Cprod3.thy [cfst_def,csnd_def,cpair_def] " (p1 << p2) = (cfst`p1 << cfst`p2 & csnd`p1 << csnd`p2)" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_snd 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_snd 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_fst 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_fst 1), - (rtac less_cprod3b 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_snd 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_snd 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_fst 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_fst 1), + (rtac less_cprod3b 1) + ]); 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) - ]); + [ + (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) + ]); qed_goalw "lub_cprod2" Cprod3.thy [cfst_def,csnd_def,cpair_def] "[|is_chain(S)|] ==> range(S) <<| \ \ <(lub(range(%i.cfst`(S i)))) , lub(range(%i.csnd`(S i)))>" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun_cprod RS ssubst) 1), - (rtac (beta_cfun RS ext RS ssubst) 1), - (rtac cont_snd 1), - (rtac (beta_cfun RS ext RS ssubst) 1), - (rtac cont_fst 1), - (rtac lub_cprod 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun_cprod RS ssubst) 1), + (rtac (beta_cfun RS ext RS ssubst) 1), + (rtac cont_snd 1), + (rtac (beta_cfun RS ext RS ssubst) 1), + (rtac cont_fst 1), + (rtac lub_cprod 1), + (atac 1) + ]); val thelub_cprod2 = (lub_cprod2 RS thelubI); (* @@ -290,23 +290,23 @@ " *) qed_goalw "csplit2" Cprod3.thy [csplit_def] - "csplit`f` = f`x`y" + "csplit`f` = f`x`y" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (cont_tacR 1), - (Simp_tac 1), - (simp_tac (!simpset addsimps [cfst2,csnd2]) 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (cont_tacR 1), + (Simp_tac 1), + (simp_tac (!simpset addsimps [cfst2,csnd2]) 1) + ]); qed_goalw "csplit3" Cprod3.thy [csplit_def] "csplit`cpair`z=z" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (cont_tacR 1), - (simp_tac (!simpset addsimps [surjective_pairing_Cprod2]) 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (cont_tacR 1), + (simp_tac (!simpset addsimps [surjective_pairing_Cprod2]) 1) + ]); (* ------------------------------------------------------------------------ *) (* install simplifier for Cprod *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Fix.ML --- a/src/HOLCF/Fix.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Fix.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/fix.ML +(* Title: HOLCF/fix.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for fix.thy @@ -14,25 +14,25 @@ qed_goal "iterate_0" Fix.thy "iterate 0 F x = x" (fn prems => - [ - (resolve_tac (nat_recs iterate_def) 1) - ]); + [ + (resolve_tac (nat_recs iterate_def) 1) + ]); qed_goal "iterate_Suc" Fix.thy "iterate (Suc n) F x = F`(iterate n F x)" (fn prems => - [ - (resolve_tac (nat_recs iterate_def) 1) - ]); + [ + (resolve_tac (nat_recs iterate_def) 1) + ]); Addsimps [iterate_0, iterate_Suc]; qed_goal "iterate_Suc2" Fix.thy "iterate (Suc n) F x = iterate n F (F`x)" (fn prems => - [ - (nat_ind_tac "n" 1), - (Simp_tac 1), - (Asm_simp_tac 1) - ]); + [ + (nat_ind_tac "n" 1), + (Simp_tac 1), + (Asm_simp_tac 1) + ]); (* ------------------------------------------------------------------------ *) (* the sequence of function itertaions is a chain *) @@ -40,26 +40,26 @@ (* ------------------------------------------------------------------------ *) qed_goalw "is_chain_iterate2" Fix.thy [is_chain] - " x << F`x ==> is_chain (%i.iterate i F x)" + " x << F`x ==> is_chain (%i.iterate i F x)" (fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (Simp_tac 1), - (nat_ind_tac "i" 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (etac monofun_cfun_arg 1) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (Simp_tac 1), + (nat_ind_tac "i" 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1), + (etac monofun_cfun_arg 1) + ]); qed_goal "is_chain_iterate" Fix.thy - "is_chain (%i.iterate i F UU)" + "is_chain (%i.iterate i F UU)" (fn prems => - [ - (rtac is_chain_iterate2 1), - (rtac minimal 1) - ]); + [ + (rtac is_chain_iterate2 1), + (rtac minimal 1) + ]); (* ------------------------------------------------------------------------ *) @@ -70,43 +70,43 @@ qed_goalw "Ifix_eq" Fix.thy [Ifix_def] "Ifix F =F`(Ifix F)" (fn prems => - [ - (rtac (contlub_cfun_arg RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac antisym_less 1), - (rtac lub_mono 1), - (rtac is_chain_iterate 1), - (rtac ch2ch_fappR 1), - (rtac is_chain_iterate 1), - (rtac allI 1), - (rtac (iterate_Suc RS subst) 1), - (rtac (is_chain_iterate RS is_chainE RS spec) 1), - (rtac is_lub_thelub 1), - (rtac ch2ch_fappR 1), - (rtac is_chain_iterate 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (rtac (iterate_Suc RS subst) 1), - (rtac is_ub_thelub 1), - (rtac is_chain_iterate 1) - ]); + [ + (rtac (contlub_cfun_arg RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac antisym_less 1), + (rtac lub_mono 1), + (rtac is_chain_iterate 1), + (rtac ch2ch_fappR 1), + (rtac is_chain_iterate 1), + (rtac allI 1), + (rtac (iterate_Suc RS subst) 1), + (rtac (is_chain_iterate RS is_chainE RS spec) 1), + (rtac is_lub_thelub 1), + (rtac ch2ch_fappR 1), + (rtac is_chain_iterate 1), + (rtac ub_rangeI 1), + (rtac allI 1), + (rtac (iterate_Suc RS subst) 1), + (rtac is_ub_thelub 1), + (rtac is_chain_iterate 1) + ]); qed_goalw "Ifix_least" Fix.thy [Ifix_def] "F`x=x ==> Ifix(F) << x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_lub_thelub 1), - (rtac is_chain_iterate 1), - (rtac ub_rangeI 1), - (strip_tac 1), - (nat_ind_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) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_lub_thelub 1), + (rtac is_chain_iterate 1), + (rtac ub_rangeI 1), + (strip_tac 1), + (nat_ind_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) + ]); (* ------------------------------------------------------------------------ *) @@ -115,18 +115,18 @@ qed_goalw "monofun_iterate" Fix.thy [monofun] "monofun(iterate(i))" (fn prems => - [ - (strip_tac 1), - (nat_ind_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) - ]); + [ + (strip_tac 1), + (nat_ind_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) + ]); (* ------------------------------------------------------------------------ *) (* the following lemma uses contlub_cfun which itself is based on a *) @@ -136,42 +136,42 @@ qed_goalw "contlub_iterate" Fix.thy [contlub] "contlub(iterate(i))" (fn prems => - [ - (strip_tac 1), - (nat_ind_tac "i" 1), - (Asm_simp_tac 1), - (rtac (lub_const RS thelubI RS sym) 1), - (Asm_simp_tac 1), - (rtac ext 1), - (rtac (thelub_fun RS ssubst) 1), - (rtac is_chainI 1), - (rtac allI 1), - (rtac (less_fun RS iffD2) 1), - (rtac allI 1), - (rtac (is_chainE RS spec) 1), - (rtac (monofun_fapp1 RS ch2ch_MF2LR) 1), - (rtac allI 1), - (rtac monofun_fapp2 1), - (atac 1), - (rtac ch2ch_fun 1), - (rtac (monofun_iterate RS ch2ch_monofun) 1), - (atac 1), - (rtac (thelub_fun RS ssubst) 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) - ]); + [ + (strip_tac 1), + (nat_ind_tac "i" 1), + (Asm_simp_tac 1), + (rtac (lub_const RS thelubI RS sym) 1), + (Asm_simp_tac 1), + (rtac ext 1), + (rtac (thelub_fun RS ssubst) 1), + (rtac is_chainI 1), + (rtac allI 1), + (rtac (less_fun RS iffD2) 1), + (rtac allI 1), + (rtac (is_chainE RS spec) 1), + (rtac (monofun_fapp1 RS ch2ch_MF2LR) 1), + (rtac allI 1), + (rtac monofun_fapp2 1), + (atac 1), + (rtac ch2ch_fun 1), + (rtac (monofun_iterate RS ch2ch_monofun) 1), + (atac 1), + (rtac (thelub_fun RS ssubst) 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) + ]); qed_goal "cont_iterate" Fix.thy "cont(iterate(i))" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_iterate 1), - (rtac contlub_iterate 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_iterate 1), + (rtac contlub_iterate 1) + ]); (* ------------------------------------------------------------------------ *) (* a lemma about continuity of iterate in its third argument *) @@ -179,37 +179,37 @@ qed_goal "monofun_iterate2" Fix.thy "monofun(iterate n F)" (fn prems => - [ - (rtac monofunI 1), - (strip_tac 1), - (nat_ind_tac "n" 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (etac monofun_cfun_arg 1) - ]); + [ + (rtac monofunI 1), + (strip_tac 1), + (nat_ind_tac "n" 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1), + (etac monofun_cfun_arg 1) + ]); qed_goal "contlub_iterate2" Fix.thy "contlub(iterate n F)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (nat_ind_tac "n" 1), - (Simp_tac 1), - (Simp_tac 1), - (res_inst_tac [("t","iterate n1 F (lub(range(%u. Y u)))"), - ("s","lub(range(%i. iterate n1 F (Y i)))")] ssubst 1), - (atac 1), - (rtac contlub_cfun_arg 1), - (etac (monofun_iterate2 RS ch2ch_monofun) 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (nat_ind_tac "n" 1), + (Simp_tac 1), + (Simp_tac 1), + (res_inst_tac [("t","iterate n1 F (lub(range(%u. Y u)))"), + ("s","lub(range(%i. iterate n1 F (Y i)))")] ssubst 1), + (atac 1), + (rtac contlub_cfun_arg 1), + (etac (monofun_iterate2 RS ch2ch_monofun) 1) + ]); qed_goal "cont_iterate2" Fix.thy "cont (iterate n F)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_iterate2 1), - (rtac contlub_iterate2 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_iterate2 1), + (rtac contlub_iterate2 1) + ]); (* ------------------------------------------------------------------------ *) (* monotonicity and continuity of Ifix *) @@ -217,15 +217,15 @@ qed_goalw "monofun_Ifix" Fix.thy [monofun,Ifix_def] "monofun(Ifix)" (fn prems => - [ - (strip_tac 1), - (rtac lub_mono 1), - (rtac is_chain_iterate 1), - (rtac is_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) - ]); + [ + (strip_tac 1), + (rtac lub_mono 1), + (rtac is_chain_iterate 1), + (rtac is_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) + ]); (* ------------------------------------------------------------------------ *) @@ -236,17 +236,17 @@ qed_goal "is_chain_iterate_lub" Fix.thy "is_chain(Y) ==> is_chain(%i. lub(range(%ia. iterate ia (Y i) UU)))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_chainI 1), - (strip_tac 1), - (rtac lub_mono 1), - (rtac is_chain_iterate 1), - (rtac is_chain_iterate 1), - (strip_tac 1), - (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun RS is_chainE + [ + (cut_facts_tac prems 1), + (rtac is_chainI 1), + (strip_tac 1), + (rtac lub_mono 1), + (rtac is_chain_iterate 1), + (rtac is_chain_iterate 1), + (strip_tac 1), + (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun RS is_chainE RS spec) 1) - ]); + ]); (* ------------------------------------------------------------------------ *) (* this exchange lemma is analog to the one for monotone functions *) @@ -257,69 +257,69 @@ qed_goal "contlub_Ifix_lemma1" Fix.thy "is_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), - (rtac (contlub_iterate RS contlubE RS spec RS mp RS ssubst) 1), - (atac 1), - (rtac refl 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (thelub_fun RS subst) 1), + (rtac (monofun_iterate RS ch2ch_monofun) 1), + (atac 1), + (rtac fun_cong 1), + (rtac (contlub_iterate RS contlubE RS spec RS mp RS ssubst) 1), + (atac 1), + (rtac refl 1) + ]); qed_goal "ex_lub_iterate" Fix.thy "is_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 is_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 is_chain_iterate_lub 1), - (strip_tac 1), - (rtac is_ub_thelub 1), - (rtac is_chain_iterate 1), - (rtac is_lub_thelub 1), - (etac is_chain_iterate_lub 1), - (rtac ub_rangeI 1), - (strip_tac 1), - (rtac lub_mono 1), - (rtac is_chain_iterate 1), - (rtac (contlub_Ifix_lemma1 RS ext RS subst) 1), - (atac 1), - (rtac is_chain_iterate 1), - (strip_tac 1), - (rtac is_ub_thelub 1), - (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1) - ]); + [ + (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 is_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 is_chain_iterate_lub 1), + (strip_tac 1), + (rtac is_ub_thelub 1), + (rtac is_chain_iterate 1), + (rtac is_lub_thelub 1), + (etac is_chain_iterate_lub 1), + (rtac ub_rangeI 1), + (strip_tac 1), + (rtac lub_mono 1), + (rtac is_chain_iterate 1), + (rtac (contlub_Ifix_lemma1 RS ext RS subst) 1), + (atac 1), + (rtac is_chain_iterate 1), + (strip_tac 1), + (rtac is_ub_thelub 1), + (etac (monofun_iterate RS ch2ch_monofun RS ch2ch_fun) 1) + ]); qed_goalw "contlub_Ifix" Fix.thy [contlub,Ifix_def] "contlub(Ifix)" (fn prems => - [ - (strip_tac 1), - (rtac (contlub_Ifix_lemma1 RS ext RS ssubst) 1), - (atac 1), - (etac ex_lub_iterate 1) - ]); + [ + (strip_tac 1), + (rtac (contlub_Ifix_lemma1 RS ext RS ssubst) 1), + (atac 1), + (etac ex_lub_iterate 1) + ]); qed_goal "cont_Ifix" Fix.thy "cont(Ifix)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Ifix 1), - (rtac contlub_Ifix 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Ifix 1), + (rtac contlub_Ifix 1) + ]); (* ------------------------------------------------------------------------ *) (* propagate properties of Ifix to its continuous counterpart *) @@ -327,65 +327,65 @@ qed_goalw "fix_eq" Fix.thy [fix_def] "fix`F = F`(fix`F)" (fn prems => - [ - (asm_simp_tac (!simpset addsimps [cont_Ifix]) 1), - (rtac Ifix_eq 1) - ]); + [ + (asm_simp_tac (!simpset addsimps [cont_Ifix]) 1), + (rtac Ifix_eq 1) + ]); qed_goalw "fix_least" Fix.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) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps [cont_Ifix]) 1), + (etac Ifix_least 1) + ]); qed_goal "fix_eqI" Fix.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) - ]); + [ + (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) + ]); qed_goal "fix_eq2" Fix.thy "f == fix`F ==> f = F`f" (fn prems => - [ - (rewrite_goals_tac prems), - (rtac fix_eq 1) - ]); + [ + (rewrite_goals_tac prems), + (rtac fix_eq 1) + ]); qed_goal "fix_eq3" Fix.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) - ]); + [ + (rtac trans 1), + (rtac ((hd prems) RS fix_eq2 RS cfun_fun_cong) 1), + (rtac refl 1) + ]); fun fix_tac3 thm i = ((rtac trans i) THEN (rtac (thm RS fix_eq3) i)); qed_goal "fix_eq4" Fix.thy "f = fix`F ==> f = F`f" (fn prems => - [ - (cut_facts_tac prems 1), - (hyp_subst_tac 1), - (rtac fix_eq 1) - ]); + [ + (cut_facts_tac prems 1), + (hyp_subst_tac 1), + (rtac fix_eq 1) + ]); qed_goal "fix_eq5" Fix.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) - ]); + [ + (rtac trans 1), + (rtac ((hd prems) RS fix_eq4 RS cfun_fun_cong) 1), + (rtac refl 1) + ]); fun fix_tac5 thm i = ((rtac trans i) THEN (rtac (thm RS fix_eq5) i)); @@ -404,12 +404,12 @@ fun fix_prover2 thy fixdef thm = prove_goal thy thm (fn prems => - [ - (rtac trans 1), - (rtac (fix_eq2) 1), - (rtac fixdef 1), - (rtac beta_cfun 1), - (cont_tacR 1) + [ + (rtac trans 1), + (rtac (fix_eq2) 1), + (rtac fixdef 1), + (rtac beta_cfun 1), + (cont_tacR 1) ]); (* ------------------------------------------------------------------------ *) @@ -419,11 +419,11 @@ qed_goal "Ifix_def2" Fix.thy "Ifix=(%x. lub(range(%i. iterate i x UU)))" (fn prems => - [ - (rtac ext 1), - (rewrite_goals_tac [Ifix_def]), - (rtac refl 1) - ]); + [ + (rtac ext 1), + (rewtac Ifix_def), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) (* direct connection between fix and iteration without Ifix *) @@ -432,10 +432,10 @@ qed_goalw "fix_def2" Fix.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) - ]); + [ + (fold_goals_tac [Ifix_def]), + (asm_simp_tac (!simpset addsimps [cont_Ifix]) 1) + ]); (* ------------------------------------------------------------------------ *) @@ -447,19 +447,19 @@ (* ------------------------------------------------------------------------ *) qed_goalw "adm_def2" Fix.thy [adm_def] - "adm(P) = (!Y. is_chain(Y) --> (!i.P(Y(i))) --> P(lub(range(Y))))" + "adm(P) = (!Y. is_chain(Y) --> (!i.P(Y(i))) --> P(lub(range(Y))))" (fn prems => - [ - (rtac refl 1) - ]); + [ + (rtac refl 1) + ]); qed_goalw "admw_def2" Fix.thy [admw_def] - "admw(P) = (!F.(!n.P(iterate n F UU)) -->\ -\ P (lub(range(%i.iterate i F UU))))" + "admw(P) = (!F.(!n.P(iterate n F UU)) -->\ +\ P (lub(range(%i.iterate i F UU))))" (fn prems => - [ - (rtac refl 1) - ]); + [ + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) (* an admissible formula is also weak admissible *) @@ -467,14 +467,14 @@ qed_goalw "adm_impl_admw" Fix.thy [admw_def] "adm(P)==>admw(P)" (fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), - (atac 1), - (rtac is_chain_iterate 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), + (atac 1), + (rtac is_chain_iterate 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* fixed point induction *) @@ -483,20 +483,20 @@ qed_goal "fix_ind" Fix.thy "[| adm(P);P(UU);!!x. P(x) ==> P(F`x)|] ==> P(fix`F)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (fix_def2 RS ssubst) 1), - (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), - (atac 1), - (rtac is_chain_iterate 1), - (rtac allI 1), - (nat_ind_tac "i" 1), - (rtac (iterate_0 RS ssubst) 1), - (atac 1), - (rtac (iterate_Suc RS ssubst) 1), - (resolve_tac prems 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (fix_def2 RS ssubst) 1), + (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), + (atac 1), + (rtac is_chain_iterate 1), + (rtac allI 1), + (nat_ind_tac "i" 1), + (rtac (iterate_0 RS ssubst) 1), + (atac 1), + (rtac (iterate_Suc RS ssubst) 1), + (resolve_tac prems 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* computational induction for weak admissible formulae *) @@ -505,14 +505,14 @@ qed_goal "wfix_ind" Fix.thy "[| admw(P); !n. P(iterate n F UU)|] ==> P(fix`F)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (fix_def2 RS ssubst) 1), - (rtac (admw_def2 RS iffD1 RS spec RS mp) 1), - (atac 1), - (rtac allI 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (fix_def2 RS ssubst) 1), + (rtac (admw_def2 RS iffD1 RS spec RS mp) 1), + (atac 1), + (rtac allI 1), + (etac spec 1) + ]); (* ------------------------------------------------------------------------ *) (* for chain-finite (easy) types every formula is admissible *) @@ -521,67 +521,67 @@ qed_goalw "adm_max_in_chain" Fix.thy [adm_def] "!Y. is_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), - (rtac (lub_finch1 RS thelubI RS ssubst) 1), - (atac 1), - (atac 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (rtac exE 1), + (rtac mp 1), + (etac spec 1), + (atac 1), + (rtac (lub_finch1 RS thelubI RS ssubst) 1), + (atac 1), + (atac 1), + (etac spec 1) + ]); qed_goalw "adm_chain_finite" Fix.thy [chain_finite_def] - "chain_finite(x::'a) ==> adm(P::'a=>bool)" + "chain_finite(x::'a) ==> adm(P::'a=>bool)" (fn prems => - [ - (cut_facts_tac prems 1), - (etac adm_max_in_chain 1) - ]); + [ + (cut_facts_tac prems 1), + (etac adm_max_in_chain 1) + ]); (* ------------------------------------------------------------------------ *) (* flat types are chain_finite *) (* ------------------------------------------------------------------------ *) qed_goalw "flat_imp_chain_finite" Fix.thy [is_flat_def,chain_finite_def] - "is_flat(x::'a)==>chain_finite(x::'a)" + "is_flat(x::'a)==>chain_finite(x::'a)" (fn prems => - [ - (rewrite_goals_tac [max_in_chain_def]), - (cut_facts_tac prems 1), - (strip_tac 1), - (res_inst_tac [("Q","!i.Y(i)=UU")] classical2 1), - (res_inst_tac [("x","0")] exI 1), - (strip_tac 1), - (rtac trans 1), - (etac spec 1), - (rtac sym 1), - (etac spec 1), - (rtac (chain_mono2 RS exE) 1), - (fast_tac HOL_cs 1), - (atac 1), - (res_inst_tac [("x","Suc(x)")] exI 1), - (strip_tac 1), - (rtac disjE 1), - (atac 3), - (rtac mp 1), - (dtac spec 1), - (etac spec 1), - (etac (le_imp_less_or_eq RS disjE) 1), - (etac (chain_mono RS mp) 1), - (atac 1), - (hyp_subst_tac 1), - (rtac refl_less 1), - (res_inst_tac [("P","Y(Suc(x)) = UU")] notE 1), - (atac 2), - (rtac mp 1), - (etac spec 1), - (Asm_simp_tac 1) - ]); + [ + (rewtac max_in_chain_def), + (cut_facts_tac prems 1), + (strip_tac 1), + (res_inst_tac [("Q","!i.Y(i)=UU")] classical2 1), + (res_inst_tac [("x","0")] exI 1), + (strip_tac 1), + (rtac trans 1), + (etac spec 1), + (rtac sym 1), + (etac spec 1), + (rtac (chain_mono2 RS exE) 1), + (fast_tac HOL_cs 1), + (atac 1), + (res_inst_tac [("x","Suc(x)")] exI 1), + (strip_tac 1), + (rtac disjE 1), + (atac 3), + (rtac mp 1), + (dtac spec 1), + (etac spec 1), + (etac (le_imp_less_or_eq RS disjE) 1), + (etac (chain_mono RS mp) 1), + (atac 1), + (hyp_subst_tac 1), + (rtac refl_less 1), + (res_inst_tac [("P","Y(Suc(x)) = UU")] notE 1), + (atac 2), + (rtac mp 1), + (etac spec 1), + (Asm_simp_tac 1) + ]); val adm_flat = flat_imp_chain_finite RS adm_chain_finite; @@ -589,11 +589,11 @@ qed_goalw "flat_void" Fix.thy [is_flat_def] "is_flat(UU::void)" (fn prems => - [ - (strip_tac 1), - (rtac disjI1 1), - (rtac unique_void2 1) - ]); + [ + (strip_tac 1), + (rtac disjI1 1), + (rtac unique_void2 1) + ]); (* ------------------------------------------------------------------------ *) (* continuous isomorphisms are strict *) @@ -604,46 +604,46 @@ "!!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) - ]); + [ + (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) + ]); qed_goal "isorep_defined" Fix.thy - "[|!x.rep`(abs`x)=x;!y.abs`(rep`y)=y; z~=UU|] ==> rep`z ~= UU" + "[|!x.rep`(abs`x)=x;!y.abs`(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","abs")] cfun_arg_cong 1), - (etac box_equals 1), - (fast_tac HOL_cs 1), - (etac (iso_strict RS conjunct1) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac swap 1), + (dtac notnotD 1), + (dres_inst_tac [("f","abs")] cfun_arg_cong 1), + (etac box_equals 1), + (fast_tac HOL_cs 1), + (etac (iso_strict RS conjunct1) 1), + (atac 1) + ]); qed_goal "isoabs_defined" Fix.thy - "[|!x.rep`(abs`x) = x;!y.abs`(rep`y)=y ; z~=UU|] ==> abs`z ~= UU" + "[|!x.rep`(abs`x) = x;!y.abs`(rep`y)=y ; z~=UU|] ==> abs`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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) (* propagation of flatness and chainfiniteness by continuous isomorphisms *) @@ -653,51 +653,51 @@ "!!f g.[|chain_finite(x::'a); !y.f`(g`y)=(y::'b) ; !x.g`(f`x)=(x::'a) |] \ \ ==> chain_finite(y::'b)" (fn prems => - [ - (rewrite_goals_tac [max_in_chain_def]), - (strip_tac 1), - (rtac exE 1), - (res_inst_tac [("P","is_chain(%i.g`(Y i))")] mp 1), - (etac spec 1), - (etac ch2ch_fappR 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) - ]); + [ + (rewtac max_in_chain_def), + (strip_tac 1), + (rtac exE 1), + (res_inst_tac [("P","is_chain(%i.g`(Y i))")] mp 1), + (etac spec 1), + (etac ch2ch_fappR 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) + ]); qed_goalw "flat2flat" Fix.thy [is_flat_def] "!!f g.[|is_flat(x::'a); !y.f`(g`y)=(y::'b) ; !x.g`(f`x)=(x::'a) |] \ \ ==> is_flat(y::'b)" (fn prems => - [ - (strip_tac 1), - (rtac disjE 1), - (res_inst_tac [("P","g`x< f`(UU::'a)=(UU::'b) | (!z.f`(z::'a)=c)" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("Q","f`(x::'a)=(UU::'b)")] classical2 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), - (res_inst_tac [("Q","f`(UU::'a)=(UU::'b)")] classical2 1), - (etac disjI1 1), - (rtac disjI2 1), - (rtac allI 1), - (res_inst_tac [("s","f`x"),("t","c")] subst 1), - (atac 1), - (res_inst_tac [("a","f`(UU::'a)")] (refl RS box_equals) 1), - (etac allE 1),(etac allE 1), - (dtac mp 1), - (res_inst_tac [("fo5","f")] (minimal RS monofun_cfun_arg) 1), - (etac disjE 1), - (contr_tac 1), - (atac 1), - (etac allE 1), - (etac allE 1), - (dtac mp 1), - (res_inst_tac [("fo5","f")] (minimal RS monofun_cfun_arg) 1), - (etac disjE 1), - (contr_tac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("Q","f`(x::'a)=(UU::'b)")] classical2 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), + (res_inst_tac [("Q","f`(UU::'a)=(UU::'b)")] classical2 1), + (etac disjI1 1), + (rtac disjI2 1), + (rtac allI 1), + (res_inst_tac [("s","f`x"),("t","c")] subst 1), + (atac 1), + (res_inst_tac [("a","f`(UU::'a)")] (refl RS box_equals) 1), + (etac allE 1),(etac allE 1), + (dtac mp 1), + (res_inst_tac [("fo5","f")] (minimal RS monofun_cfun_arg) 1), + (etac disjE 1), + (contr_tac 1), + (atac 1), + (etac allE 1), + (etac allE 1), + (dtac mp 1), + (res_inst_tac [("fo5","f")] (minimal RS monofun_cfun_arg) 1), + (etac disjE 1), + (contr_tac 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* admissibility of special formulae and propagation *) (* ------------------------------------------------------------------------ *) qed_goalw "adm_less" Fix.thy [adm_def] - "[|cont u;cont v|]==> adm(%x.u x << v x)" + "[|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) - ]); + [ + (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) + ]); qed_goal "adm_conj" Fix.thy - "[| adm P; adm Q |] ==> adm(%x. P x & Q x)" + "[| adm P; adm Q |] ==> adm(%x. P x & Q x)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (adm_def2 RS iffD2) 1), - (strip_tac 1), - (rtac conjI 1), - (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), - (atac 1), - (atac 1), - (fast_tac HOL_cs 1), - (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), - (atac 1), - (atac 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (adm_def2 RS iffD2) 1), + (strip_tac 1), + (rtac conjI 1), + (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), + (atac 1), + (atac 1), + (fast_tac HOL_cs 1), + (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), + (atac 1), + (atac 1), + (fast_tac HOL_cs 1) + ]); qed_goal "adm_cong" Fix.thy - "(!x. P x = Q x) ==> adm P = adm Q " + "(!x. P x = Q x) ==> adm P = adm Q " (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("s","P"),("t","Q")] subst 1), - (rtac refl 2), - (rtac ext 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("s","P"),("t","Q")] subst 1), + (rtac refl 2), + (rtac ext 1), + (etac spec 1) + ]); qed_goalw "adm_not_free" Fix.thy [adm_def] "adm(%x.t)" (fn prems => - [ - (fast_tac HOL_cs 1) - ]); + [ + (fast_tac HOL_cs 1) + ]); qed_goalw "adm_not_less" Fix.thy [adm_def] - "cont t ==> adm(%x.~ (t x) << u)" + "cont t ==> adm(%x.~ (t x) << u)" (fn prems => - [ - (cut_facts_tac prems 1), - (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) - ]); + [ + (cut_facts_tac prems 1), + (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) + ]); qed_goal "adm_all" Fix.thy - " !y.adm(P y) ==> adm(%x.!y.P y x)" + " !y.adm(P y) ==> adm(%x.!y.P y x)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (adm_def2 RS iffD2) 1), - (strip_tac 1), - (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), - (etac spec 1), - (atac 1), - (rtac allI 1), - (dtac spec 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (adm_def2 RS iffD2) 1), + (strip_tac 1), + (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), + (etac spec 1), + (atac 1), + (rtac allI 1), + (dtac spec 1), + (etac spec 1) + ]); val adm_all2 = (allI RS adm_all); qed_goal "adm_subst" Fix.thy - "[|cont t; adm P|] ==> adm(%x. P (t x))" + "[|cont t; adm P|] ==> adm(%x. P (t x))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (adm_def2 RS iffD2) 1), - (strip_tac 1), - (rtac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1), - (atac 1), - (atac 1), - (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), - (atac 1), - (rtac (cont2mono RS ch2ch_monofun) 1), - (atac 1), - (atac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (adm_def2 RS iffD2) 1), + (strip_tac 1), + (rtac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1), + (atac 1), + (atac 1), + (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), + (atac 1), + (rtac (cont2mono RS ch2ch_monofun) 1), + (atac 1), + (atac 1), + (atac 1) + ]); qed_goal "adm_UU_not_less" Fix.thy "adm(%x.~ UU << t(x))" (fn prems => - [ - (res_inst_tac [("P2","%x.False")] (adm_cong RS iffD1) 1), - (Asm_simp_tac 1), - (rtac adm_not_free 1) - ]); + [ + (res_inst_tac [("P2","%x.False")] (adm_cong RS iffD1) 1), + (Asm_simp_tac 1), + (rtac adm_not_free 1) + ]); qed_goalw "adm_not_UU" Fix.thy [adm_def] - "cont(t)==> adm(%x.~ (t x) = UU)" + "cont(t)==> adm(%x.~ (t x) = UU)" (fn prems => - [ - (cut_facts_tac prems 1), - (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) - ]); + [ + (cut_facts_tac prems 1), + (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" Fix.thy - "[|cont u ; cont v|]==> adm(%x. u x = v x)" + "[|cont u ; cont v|]==> adm(%x. u x = v x)" (fn prems => - [ - (rtac (adm_cong RS iffD1) 1), - (rtac allI 1), - (rtac iffI 1), - (rtac antisym_less 1), - (rtac antisym_less_inverse 3), - (atac 3), - (etac conjunct1 1), - (etac conjunct2 1), - (rtac adm_conj 1), - (rtac adm_less 1), - (resolve_tac prems 1), - (resolve_tac prems 1), - (rtac adm_less 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); + [ + (rtac (adm_cong RS iffD1) 1), + (rtac allI 1), + (rtac iffI 1), + (rtac antisym_less 1), + (rtac antisym_less_inverse 3), + (atac 3), + (etac conjunct1 1), + (etac conjunct2 1), + (rtac adm_conj 1), + (rtac adm_less 1), + (resolve_tac prems 1), + (resolve_tac prems 1), + (rtac adm_less 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); (* ------------------------------------------------------------------------ *) @@ -901,279 +901,279 @@ "[| is_chain Y; !n.P (Y n) | Q(Y n)|]\ \ ==> (? i.!j. i Q(Y(j))) | (!i.? j.i - [ - (cut_facts_tac prems 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (fast_tac HOL_cs 1) + ]); qed_goal "adm_disj_lemma2" Fix.thy "[| adm(Q); ? X.is_chain(X) & (!n.Q(X(n))) &\ \ lub(range(Y))=lub(range(X))|] ==> Q(lub(range(Y)))" (fn prems => - [ - (cut_facts_tac prems 1), - (etac exE 1), - (etac conjE 1), - (etac conjE 1), - (res_inst_tac [("s","lub(range(X))"),("t","lub(range(Y))")] ssubst 1), - (atac 1), - (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), - (atac 1), - (atac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac exE 1), + (etac conjE 1), + (etac conjE 1), + (res_inst_tac [("s","lub(range(X))"),("t","lub(range(Y))")] ssubst 1), + (atac 1), + (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), + (atac 1), + (atac 1), + (atac 1) + ]); qed_goal "adm_disj_lemma3" Fix.thy "[| is_chain(Y); ! j. i < j --> Q(Y(j)) |] ==>\ \ is_chain(%m. if m < Suc i then Y(Suc i) else Y m)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_chainI 1), - (rtac allI 1), - (res_inst_tac [("m","i"),("n","ia")] nat_less_cases 1), - (res_inst_tac [("s","False"),("t","ia < Suc(i)")] ssubst 1), - (rtac iffI 1), - (etac FalseE 2), - (rtac notE 1), - (rtac (not_less_eq RS iffD2) 1), - (atac 1), - (atac 1), - (res_inst_tac [("s","False"),("t","Suc(ia) < Suc(i)")] ssubst 1), - (Asm_simp_tac 1), - (rtac iffI 1), - (etac FalseE 2), - (rtac notE 1), - (etac less_not_sym 1), - (atac 1), - (Asm_simp_tac 1), - (etac (is_chainE RS spec) 1), - (hyp_subst_tac 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_chainI 1), + (rtac allI 1), + (res_inst_tac [("m","i"),("n","ia")] nat_less_cases 1), + (res_inst_tac [("s","False"),("t","ia < Suc(i)")] ssubst 1), + (rtac iffI 1), + (etac FalseE 2), + (rtac notE 1), + (rtac (not_less_eq RS iffD2) 1), + (atac 1), + (atac 1), + (res_inst_tac [("s","False"),("t","Suc(ia) < Suc(i)")] ssubst 1), + (Asm_simp_tac 1), + (rtac iffI 1), + (etac FalseE 2), + (rtac notE 1), + (etac less_not_sym 1), + (atac 1), + (Asm_simp_tac 1), + (etac (is_chainE RS spec) 1), + (hyp_subst_tac 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1) + ]); qed_goal "adm_disj_lemma4" Fix.thy "[| ! j. i < j --> Q(Y(j)) |] ==>\ -\ ! n. Q( if n < Suc i then Y(Suc i) else Y n)" +\ ! n. Q( if n < Suc i then Y(Suc i) else Y n)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac allI 1), - (res_inst_tac [("m","n"),("n","Suc(i)")] nat_less_cases 1), - (res_inst_tac[("s","Y(Suc(i))"),("t","if n'a); ! 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 => - [ - (cut_facts_tac prems 1), - (rtac lub_equal2 1), - (atac 2), - (rtac adm_disj_lemma3 2), - (atac 2), - (atac 2), - (res_inst_tac [("x","i")] exI 1), - (strip_tac 1), - (res_inst_tac [("s","False"),("t","ia < Suc(i)")] ssubst 1), - (rtac iffI 1), - (etac FalseE 2), - (rtac notE 1), - (rtac (not_less_eq RS iffD2) 1), - (atac 1), - (atac 1), - (rtac (if_False RS ssubst) 1), - (rtac refl 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac lub_equal2 1), + (atac 2), + (rtac adm_disj_lemma3 2), + (atac 2), + (atac 2), + (res_inst_tac [("x","i")] exI 1), + (strip_tac 1), + (res_inst_tac [("s","False"),("t","ia < Suc(i)")] ssubst 1), + (rtac iffI 1), + (etac FalseE 2), + (rtac notE 1), + (rtac (not_less_eq RS iffD2) 1), + (atac 1), + (atac 1), + (rtac (if_False RS ssubst) 1), + (rtac refl 1) + ]); qed_goal "adm_disj_lemma6" Fix.thy "[| is_chain(Y::nat=>'a); ? i. ! j. i < j --> Q(Y(j)) |] ==>\ \ ? X. is_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); ! i. ? j. i < j & P(Y(j)) |] ==>\ \ is_chain(%m. Y(theleast(%j. m - [ - (cut_facts_tac prems 1), - (rtac is_chainI 1), - (rtac allI 1), - (rtac chain_mono3 1), - (atac 1), - (rtac theleast2 1), - (rtac conjI 1), - (rtac Suc_lessD 1), - (etac allE 1), - (etac exE 1), - (rtac (theleast1 RS conjunct1) 1), - (atac 1), - (etac allE 1), - (etac exE 1), - (rtac (theleast1 RS conjunct2) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_chainI 1), + (rtac allI 1), + (rtac chain_mono3 1), + (atac 1), + (rtac theleast2 1), + (rtac conjI 1), + (rtac Suc_lessD 1), + (etac allE 1), + (etac exE 1), + (rtac (theleast1 RS conjunct1) 1), + (atac 1), + (etac allE 1), + (etac exE 1), + (rtac (theleast1 RS conjunct2) 1), + (atac 1) + ]); qed_goal "adm_disj_lemma8" Fix.thy "[| ! i. ? j. i < j & P(Y(j)) |] ==> ! m. P(Y(theleast(%j. m - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (etac allE 1), - (etac exE 1), - (etac (theleast1 RS conjunct2) 1) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (etac allE 1), + (etac exE 1), + (etac (theleast1 RS conjunct2) 1) + ]); qed_goal "adm_disj_lemma9" Fix.thy "[| is_chain(Y::nat=>'a); ! i. ? j. i < j & P(Y(j)) |] ==>\ \ lub(range(Y)) = lub(range(%m. Y(theleast(%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 (theleast1 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) - ]); + [ + (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 (theleast1 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) + ]); qed_goal "adm_disj_lemma10" Fix.thy "[| is_chain(Y::nat=>'a); ! i. ? j. i < j & P(Y(j)) |] ==>\ \ ? X. is_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(theleast(%j. mP(lub(range(Y)))" (fn prems => - [ - (cut_facts_tac prems 1), - (etac adm_disj_lemma2 1), - (etac adm_disj_lemma10 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac adm_disj_lemma2 1), + (etac adm_disj_lemma10 1), + (atac 1) + ]); qed_goal "adm_disj_lemma12" Fix.thy "[| adm(P); is_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_lemma6 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac adm_disj_lemma2 1), + (etac adm_disj_lemma6 1), + (atac 1) + ]); qed_goal "adm_disj" Fix.thy - "[| adm P; adm Q |] ==> adm(%x.P x | Q x)" + "[| adm P; adm Q |] ==> adm(%x.P x | Q x)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (adm_def2 RS iffD2) 1), - (strip_tac 1), - (rtac (adm_disj_lemma1 RS disjE) 1), - (atac 1), - (atac 1), - (rtac disjI2 1), - (etac adm_disj_lemma12 1), - (atac 1), - (atac 1), - (rtac disjI1 1), - (etac adm_disj_lemma11 1), - (atac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (adm_def2 RS iffD2) 1), + (strip_tac 1), + (rtac (adm_disj_lemma1 RS disjE) 1), + (atac 1), + (atac 1), + (rtac disjI2 1), + (etac adm_disj_lemma12 1), + (atac 1), + (atac 1), + (rtac disjI1 1), + (etac adm_disj_lemma11 1), + (atac 1), + (atac 1) + ]); qed_goal "adm_impl" Fix.thy - "[| adm(%x.~(P x)); adm Q |] ==> adm(%x.P x --> Q x)" + "[| adm(%x.~(P x)); adm Q |] ==> adm(%x.P x --> Q x)" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("P2","%x.~(P x)|Q x")] (adm_cong RS iffD1) 1), - (fast_tac HOL_cs 1), - (rtac adm_disj 1), - (atac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("P2","%x.~(P x)|Q x")] (adm_cong RS iffD1) 1), + (fast_tac HOL_cs 1), + (rtac adm_disj 1), + (atac 1), + (atac 1) + ]); val adm_thms = [adm_impl,adm_disj,adm_eq,adm_not_UU,adm_UU_not_less, - adm_all2,adm_not_less,adm_not_free,adm_conj,adm_less - ]; + adm_all2,adm_not_less,adm_not_free,adm_conj,adm_less + ]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Fun1.ML --- a/src/HOLCF/Fun1.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Fun1.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/fun1.ML +(* Title: HOLCF/fun1.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for fun1.thy @@ -14,30 +14,30 @@ qed_goalw "refl_less_fun" Fun1.thy [less_fun_def] "less_fun f f" (fn prems => - [ - (fast_tac (HOL_cs addSIs [refl_less]) 1) - ]); + [ + (fast_tac (HOL_cs addSIs [refl_less]) 1) + ]); qed_goalw "antisym_less_fun" Fun1.thy [less_fun_def] - "[|less_fun f1 f2; less_fun f2 f1|] ==> f1 = f2" + "[|less_fun f1 f2; less_fun f2 f1|] ==> f1 = f2" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (expand_fun_eq RS ssubst) 1), - (fast_tac (HOL_cs addSIs [antisym_less]) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (expand_fun_eq RS ssubst) 1), + (fast_tac (HOL_cs addSIs [antisym_less]) 1) + ]); qed_goalw "trans_less_fun" Fun1.thy [less_fun_def] - "[|less_fun f1 f2; less_fun f2 f3 |] ==> less_fun f1 f3" + "[|less_fun f1 f2; less_fun f2 f3 |] ==> less_fun 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)) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (rtac trans_less 1), + (etac allE 1), + (atac 1), + ((etac allE 1) THEN (atac 1)) + ]); (* -------------------------------------------------------------------------- diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Fun2.ML --- a/src/HOLCF/Fun2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Fun2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/fun2.ML +(* Title: HOLCF/fun2.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for fun2.thy @@ -14,11 +14,11 @@ qed_goalw "minimal_fun" Fun2.thy [UU_fun_def] "UU_fun << f" (fn prems => - [ - (rtac (inst_fun_po RS ssubst) 1), - (rewrite_goals_tac [less_fun_def]), - (fast_tac (HOL_cs addSIs [minimal]) 1) - ]); + [ + (rtac (inst_fun_po RS ssubst) 1), + (rewtac less_fun_def), + (fast_tac (HOL_cs addSIs [minimal]) 1) + ]); (* ------------------------------------------------------------------------ *) (* make the symbol << accessible for type fun *) @@ -26,28 +26,28 @@ qed_goal "less_fun" Fun2.thy "(f1 << f2) = (! x. f1(x) << f2(x))" (fn prems => - [ - (rtac (inst_fun_po RS ssubst) 1), - (fold_goals_tac [less_fun_def]), - (rtac refl 1) - ]); + [ + (rtac (inst_fun_po RS ssubst) 1), + (fold_goals_tac [less_fun_def]), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) (* chains of functions yield chains in the po range *) (* ------------------------------------------------------------------------ *) qed_goal "ch2ch_fun" Fun2.thy - "is_chain(S::nat=>('a::term => 'b::po)) ==> is_chain(% i.S(i)(x))" + "is_chain(S::nat=>('a::term => 'b::po)) ==> is_chain(% i.S(i)(x))" (fn prems => - [ - (cut_facts_tac prems 1), - (rewrite_goals_tac [is_chain]), - (rtac allI 1), - (rtac spec 1), - (rtac (less_fun RS subst) 1), - (etac allE 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rewtac is_chain), + (rtac allI 1), + (rtac spec 1), + (rtac (less_fun RS subst) 1), + (etac allE 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* upper bounds of function chains yield upper bound in the po range *) @@ -56,51 +56,51 @@ 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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) (* Type 'a::term => 'b::pcpo is chain complete *) (* ------------------------------------------------------------------------ *) qed_goal "lub_fun" Fun2.thy - "is_chain(S::nat=>('a::term => 'b::pcpo)) ==> \ + "is_chain(S::nat=>('a::term => 'b::pcpo)) ==> \ \ 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), - (rtac (less_fun RS ssubst) 1), - (rtac allI 1), - (rtac is_ub_thelub 1), - (etac ch2ch_fun 1), - (strip_tac 1), - (rtac (less_fun RS ssubst) 1), - (rtac allI 1), - (rtac is_lub_thelub 1), - (etac ch2ch_fun 1), - (etac ub2ub_fun 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_lubI 1), + (rtac conjI 1), + (rtac ub_rangeI 1), + (rtac allI 1), + (rtac (less_fun RS ssubst) 1), + (rtac allI 1), + (rtac is_ub_thelub 1), + (etac ch2ch_fun 1), + (strip_tac 1), + (rtac (less_fun RS ssubst) 1), + (rtac allI 1), + (rtac is_lub_thelub 1), + (etac ch2ch_fun 1), + (etac ub2ub_fun 1) + ]); val thelub_fun = (lub_fun RS thelubI); (* is_chain ?S1 ==> lub (range ?S1) = (%x. lub (range (%i. ?S1 i x))) *) qed_goal "cpo_fun" Fun2.thy - "is_chain(S::nat=>('a::term => 'b::pcpo)) ==> ? x. range(S) <<| x" + "is_chain(S::nat=>('a::term => 'b::pcpo)) ==> ? x. range(S) <<| x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac exI 1), - (etac lub_fun 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac exI 1), + (etac lub_fun 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Fun3.ML --- a/src/HOLCF/Fun3.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Fun3.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/fun3.ML +(* Title: HOLCF/fun3.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/HOLCF.ML --- a/src/HOLCF/HOLCF.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/HOLCF.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/HOLCF.ML +(* Title: HOLCF/HOLCF.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Holcfb.ML --- a/src/HOLCF/Holcfb.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Holcfb.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/holcfb.ML +(* Title: HOLCF/holcfb.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for Holcfb.thy @@ -13,27 +13,27 @@ (* ------------------------------------------------------------------------ *) qed_goal "well_ordering_nat" Nat.thy - "!P. P(x::nat) --> (? y. P(y) & (! x. P(x) --> y <= x))" + "!P. P(x::nat) --> (? y. P(y) & (! x. P(x) --> y <= x))" (fn prems => - [ - (res_inst_tac [("n","x")] less_induct 1), - (strip_tac 1), - (res_inst_tac [("Q","? k.k P(theleast(P)) & (!x.P(x)--> theleast(P) <= x)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (well_ordering_nat RS spec RS mp RS exE) 1), - (atac 1), - (rtac selectI 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (well_ordering_nat RS spec RS mp RS exE) 1), + (atac 1), + (rtac selectI 1), + (atac 1) + ]); val theleast1= theleast RS conjunct1; (* ?P1(?x1) ==> ?P1(theleast(?P1)) *) qed_goal "theleast2" Holcfb.thy "P(x) ==> theleast(P) <= x" (fn prems => - [ - (rtac (theleast RS conjunct2 RS spec RS mp) 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); + [ + (rtac (theleast RS conjunct2 RS spec RS mp) 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); (* ------------------------------------------------------------------------ *) @@ -70,80 +70,80 @@ qed_goal "de_morgan1" HOL.thy "(~a & ~b)=(~(a | b))" (fn prems => - [ - (rtac iffI 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); + [ + (rtac iffI 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); qed_goal "de_morgan2" HOL.thy "(~a | ~b)=(~(a & b))" (fn prems => - [ - (rtac iffI 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); + [ + (rtac iffI 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); qed_goal "notall2ex" HOL.thy "(~ (!x.P(x))) = (? x.~P(x))" (fn prems => - [ - (rtac iffI 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); + [ + (rtac iffI 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); qed_goal "notex2all" HOL.thy "(~ (? x.P(x))) = (!x.~P(x))" (fn prems => - [ - (rtac iffI 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); + [ + (rtac iffI 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); qed_goal "selectI3" HOL.thy "(? x. P(x)) ==> P(@ x.P(x))" (fn prems => - [ - (cut_facts_tac prems 1), - (etac exE 1), - (etac selectI 1) - ]); + [ + (cut_facts_tac prems 1), + (etac exE 1), + (etac selectI 1) + ]); qed_goal "selectE" HOL.thy "P(@ x.P(x)) ==> (? x. P(x))" (fn prems => - [ - (cut_facts_tac prems 1), - (etac exI 1) - ]); + [ + (cut_facts_tac prems 1), + (etac exI 1) + ]); qed_goal "select_eq_Ex" HOL.thy "(P(@ x.P(x))) = (? x. P(x))" (fn prems => - [ - (rtac (iff RS mp RS mp) 1), - (strip_tac 1), - (etac selectE 1), - (strip_tac 1), - (etac selectI3 1) - ]); + [ + (rtac (iff RS mp RS mp) 1), + (strip_tac 1), + (etac selectE 1), + (strip_tac 1), + (etac selectI3 1) + ]); qed_goal "notnotI" HOL.thy "P ==> ~~P" (fn prems => - [ - (cut_facts_tac prems 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (fast_tac HOL_cs 1) + ]); qed_goal "classical2" HOL.thy "[|Q ==> R; ~Q ==> R|]==>R" (fn prems => - [ - (rtac disjE 1), - (rtac excluded_middle 1), - (eresolve_tac prems 1), - (eresolve_tac prems 1) - ]); + [ + (rtac disjE 1), + (rtac excluded_middle 1), + (eresolve_tac prems 1), + (eresolve_tac prems 1) + ]); @@ -154,13 +154,13 @@ qed_goal "nat_less_cases" Nat.thy "[| (m::nat) < n ==> P(n,m); m=n ==> P(n,m);n < m ==> P(n,m)|]==>P(n,m)" ( fn prems => - [ - (res_inst_tac [("m1","n"),("n1","m")] (less_linear RS disjE) 1), - (etac disjE 2), - (etac (hd (tl (tl prems))) 1), - (etac (sym RS hd (tl prems)) 1), - (etac (hd prems) 1) - ]); + [ + (res_inst_tac [("m1","n"),("n1","m")] (less_linear RS disjE) 1), + (etac disjE 2), + (etac (hd (tl (tl prems))) 1), + (etac (sym RS hd (tl prems)) 1), + (etac (hd prems) 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Lift1.ML --- a/src/HOLCF/Lift1.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Lift1.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,188 +1,188 @@ -(* Title: HOLCF/lift1.ML +(* Title: HOLCF/lift1.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen *) open Lift1; qed_goalw "Exh_Lift" Lift1.thy [UU_lift_def,Iup_def ] - "z = UU_lift | (? x. z = Iup(x))" + "z = UU_lift | (? x. z = Iup(x))" (fn prems => - [ - (rtac (Rep_Lift_inverse RS subst) 1), - (res_inst_tac [("s","Rep_Lift(z)")] sumE 1), - (rtac disjI1 1), - (res_inst_tac [("f","Abs_Lift")] arg_cong 1), - (rtac (unique_void2 RS subst) 1), - (atac 1), - (rtac disjI2 1), - (rtac exI 1), - (res_inst_tac [("f","Abs_Lift")] arg_cong 1), - (atac 1) - ]); + [ + (rtac (Rep_Lift_inverse RS subst) 1), + (res_inst_tac [("s","Rep_Lift(z)")] sumE 1), + (rtac disjI1 1), + (res_inst_tac [("f","Abs_Lift")] arg_cong 1), + (rtac (unique_void2 RS subst) 1), + (atac 1), + (rtac disjI2 1), + (rtac exI 1), + (res_inst_tac [("f","Abs_Lift")] arg_cong 1), + (atac 1) + ]); qed_goal "inj_Abs_Lift" Lift1.thy "inj(Abs_Lift)" (fn prems => - [ - (rtac inj_inverseI 1), - (rtac Abs_Lift_inverse 1) - ]); + [ + (rtac inj_inverseI 1), + (rtac Abs_Lift_inverse 1) + ]); qed_goal "inj_Rep_Lift" Lift1.thy "inj(Rep_Lift)" (fn prems => - [ - (rtac inj_inverseI 1), - (rtac Rep_Lift_inverse 1) - ]); + [ + (rtac inj_inverseI 1), + (rtac Rep_Lift_inverse 1) + ]); qed_goalw "inject_Iup" Lift1.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_Lift RS injD) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (inj_Inr RS injD) 1), + (rtac (inj_Abs_Lift RS injD) 1), + (atac 1) + ]); qed_goalw "defined_Iup" Lift1.thy [Iup_def,UU_lift_def] "Iup(x)~=UU_lift" (fn prems => - [ - (rtac notI 1), - (rtac notE 1), - (rtac Inl_not_Inr 1), - (rtac sym 1), - (etac (inj_Abs_Lift RS injD) 1) - ]); + [ + (rtac notI 1), + (rtac notE 1), + (rtac Inl_not_Inr 1), + (rtac sym 1), + (etac (inj_Abs_Lift RS injD) 1) + ]); qed_goal "liftE" Lift1.thy - "[| p=UU_lift ==> Q; !!x. p=Iup(x)==>Q|] ==>Q" + "[| p=UU_lift ==> Q; !!x. p=Iup(x)==>Q|] ==>Q" (fn prems => - [ - (rtac (Exh_Lift RS disjE) 1), - (eresolve_tac prems 1), - (etac exE 1), - (eresolve_tac prems 1) - ]); + [ + (rtac (Exh_Lift RS disjE) 1), + (eresolve_tac prems 1), + (etac exE 1), + (eresolve_tac prems 1) + ]); qed_goalw "Ilift1" Lift1.thy [Ilift_def,UU_lift_def] - "Ilift(f)(UU_lift)=UU" + "Ilift(f)(UU_lift)=UU" (fn prems => - [ - (rtac (Abs_Lift_inverse RS ssubst) 1), - (rtac (sum_case_Inl RS ssubst) 1), - (rtac refl 1) - ]); + [ + (rtac (Abs_Lift_inverse RS ssubst) 1), + (rtac (sum_case_Inl RS ssubst) 1), + (rtac refl 1) + ]); qed_goalw "Ilift2" Lift1.thy [Ilift_def,Iup_def] - "Ilift(f)(Iup(x))=f`x" + "Ilift(f)(Iup(x))=f`x" (fn prems => - [ - (rtac (Abs_Lift_inverse RS ssubst) 1), - (rtac (sum_case_Inr RS ssubst) 1), - (rtac refl 1) - ]); + [ + (rtac (Abs_Lift_inverse RS ssubst) 1), + (rtac (sum_case_Inr RS ssubst) 1), + (rtac refl 1) + ]); val Lift0_ss = (simpset_of "Cfun3") addsimps [Ilift1,Ilift2]; qed_goalw "less_lift1a" Lift1.thy [less_lift_def,UU_lift_def] - "less_lift(UU_lift)(z)" + "less_lift(UU_lift)(z)" (fn prems => - [ - (rtac (Abs_Lift_inverse RS ssubst) 1), - (rtac (sum_case_Inl RS ssubst) 1), - (rtac TrueI 1) - ]); + [ + (rtac (Abs_Lift_inverse RS ssubst) 1), + (rtac (sum_case_Inl RS ssubst) 1), + (rtac TrueI 1) + ]); qed_goalw "less_lift1b" Lift1.thy [Iup_def,less_lift_def,UU_lift_def] - "~less_lift (Iup x) UU_lift" + "~less_lift (Iup x) UU_lift" (fn prems => - [ - (rtac notI 1), - (rtac iffD1 1), - (atac 2), - (rtac (Abs_Lift_inverse RS ssubst) 1), - (rtac (Abs_Lift_inverse RS ssubst) 1), - (rtac (sum_case_Inr RS ssubst) 1), - (rtac (sum_case_Inl RS ssubst) 1), - (rtac refl 1) - ]); + [ + (rtac notI 1), + (rtac iffD1 1), + (atac 2), + (rtac (Abs_Lift_inverse RS ssubst) 1), + (rtac (Abs_Lift_inverse RS ssubst) 1), + (rtac (sum_case_Inr RS ssubst) 1), + (rtac (sum_case_Inl RS ssubst) 1), + (rtac refl 1) + ]); qed_goalw "less_lift1c" Lift1.thy [Iup_def,less_lift_def,UU_lift_def] - "less_lift (Iup x) (Iup y)=(x< - [ - (rtac (Abs_Lift_inverse RS ssubst) 1), - (rtac (Abs_Lift_inverse RS ssubst) 1), - (rtac (sum_case_Inr RS ssubst) 1), - (rtac (sum_case_Inr RS ssubst) 1), - (rtac refl 1) - ]); + [ + (rtac (Abs_Lift_inverse RS ssubst) 1), + (rtac (Abs_Lift_inverse RS ssubst) 1), + (rtac (sum_case_Inr RS ssubst) 1), + (rtac (sum_case_Inr RS ssubst) 1), + (rtac refl 1) + ]); qed_goal "refl_less_lift" Lift1.thy "less_lift p p" (fn prems => - [ - (res_inst_tac [("p","p")] liftE 1), - (hyp_subst_tac 1), - (rtac less_lift1a 1), - (hyp_subst_tac 1), - (rtac (less_lift1c RS iffD2) 1), - (rtac refl_less 1) - ]); + [ + (res_inst_tac [("p","p")] liftE 1), + (hyp_subst_tac 1), + (rtac less_lift1a 1), + (hyp_subst_tac 1), + (rtac (less_lift1c RS iffD2) 1), + (rtac refl_less 1) + ]); qed_goal "antisym_less_lift" Lift1.thy - "[|less_lift p1 p2;less_lift p2 p1|] ==> p1=p2" + "[|less_lift p1 p2;less_lift p2 p1|] ==> p1=p2" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p1")] liftE 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p2")] liftE 1), - (hyp_subst_tac 1), - (rtac refl 1), - (hyp_subst_tac 1), - (res_inst_tac [("P","less_lift (Iup x) UU_lift")] notE 1), - (rtac less_lift1b 1), - (atac 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p2")] liftE 1), - (hyp_subst_tac 1), - (res_inst_tac [("P","less_lift (Iup x) UU_lift")] notE 1), - (rtac less_lift1b 1), - (atac 1), - (hyp_subst_tac 1), - (rtac arg_cong 1), - (rtac antisym_less 1), - (etac (less_lift1c RS iffD1) 1), - (etac (less_lift1c RS iffD1) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","p1")] liftE 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p2")] liftE 1), + (hyp_subst_tac 1), + (rtac refl 1), + (hyp_subst_tac 1), + (res_inst_tac [("P","less_lift (Iup x) UU_lift")] notE 1), + (rtac less_lift1b 1), + (atac 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p2")] liftE 1), + (hyp_subst_tac 1), + (res_inst_tac [("P","less_lift (Iup x) UU_lift")] notE 1), + (rtac less_lift1b 1), + (atac 1), + (hyp_subst_tac 1), + (rtac arg_cong 1), + (rtac antisym_less 1), + (etac (less_lift1c RS iffD1) 1), + (etac (less_lift1c RS iffD1) 1) + ]); qed_goal "trans_less_lift" Lift1.thy - "[|less_lift p1 p2;less_lift p2 p3|] ==> less_lift p1 p3" + "[|less_lift p1 p2;less_lift p2 p3|] ==> less_lift p1 p3" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p1")] liftE 1), - (hyp_subst_tac 1), - (rtac less_lift1a 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p2")] liftE 1), - (hyp_subst_tac 1), - (rtac notE 1), - (rtac less_lift1b 1), - (atac 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p3")] liftE 1), - (hyp_subst_tac 1), - (rtac notE 1), - (rtac less_lift1b 1), - (atac 1), - (hyp_subst_tac 1), - (rtac (less_lift1c RS iffD2) 1), - (rtac trans_less 1), - (etac (less_lift1c RS iffD1) 1), - (etac (less_lift1c RS iffD1) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","p1")] liftE 1), + (hyp_subst_tac 1), + (rtac less_lift1a 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p2")] liftE 1), + (hyp_subst_tac 1), + (rtac notE 1), + (rtac less_lift1b 1), + (atac 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p3")] liftE 1), + (hyp_subst_tac 1), + (rtac notE 1), + (rtac less_lift1b 1), + (atac 1), + (hyp_subst_tac 1), + (rtac (less_lift1c RS iffD2) 1), + (rtac trans_less 1), + (etac (less_lift1c RS iffD1) 1), + (etac (less_lift1c RS iffD1) 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Lift2.ML --- a/src/HOLCF/Lift2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Lift2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/lift2.ML +(* Title: HOLCF/lift2.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for lift2.thy @@ -14,10 +14,10 @@ qed_goal "minimal_lift" Lift2.thy "UU_lift << z" (fn prems => - [ - (rtac (inst_lift_po RS ssubst) 1), - (rtac less_lift1a 1) - ]); + [ + (rtac (inst_lift_po RS ssubst) 1), + (rtac less_lift1a 1) + ]); (* -------------------------------------------------------------------------*) (* access to less_lift in class po *) @@ -25,17 +25,17 @@ qed_goal "less_lift2b" Lift2.thy "~ Iup(x) << UU_lift" (fn prems => - [ - (rtac (inst_lift_po RS ssubst) 1), - (rtac less_lift1b 1) - ]); + [ + (rtac (inst_lift_po RS ssubst) 1), + (rtac less_lift1b 1) + ]); qed_goal "less_lift2c" Lift2.thy "(Iup(x)< - [ - (rtac (inst_lift_po RS ssubst) 1), - (rtac less_lift1c 1) - ]); + [ + (rtac (inst_lift_po RS ssubst) 1), + (rtac less_lift1c 1) + ]); (* ------------------------------------------------------------------------ *) (* Iup and Ilift are monotone *) @@ -43,40 +43,40 @@ qed_goalw "monofun_Iup" Lift2.thy [monofun] "monofun(Iup)" (fn prems => - [ - (strip_tac 1), - (etac (less_lift2c RS iffD2) 1) - ]); + [ + (strip_tac 1), + (etac (less_lift2c RS iffD2) 1) + ]); qed_goalw "monofun_Ilift1" Lift2.thy [monofun] "monofun(Ilift)" (fn prems => - [ - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (res_inst_tac [("p","xa")] liftE 1), + [ + (strip_tac 1), + (rtac (less_fun RS iffD2) 1), + (strip_tac 1), + (res_inst_tac [("p","xa")] liftE 1), (asm_simp_tac Lift0_ss 1), (asm_simp_tac Lift0_ss 1), - (etac monofun_cfun_fun 1) - ]); + (etac monofun_cfun_fun 1) + ]); qed_goalw "monofun_Ilift2" Lift2.thy [monofun] "monofun(Ilift(f))" (fn prems => - [ - (strip_tac 1), - (res_inst_tac [("p","x")] liftE 1), + [ + (strip_tac 1), + (res_inst_tac [("p","x")] liftE 1), (asm_simp_tac Lift0_ss 1), (asm_simp_tac Lift0_ss 1), - (res_inst_tac [("p","y")] liftE 1), - (hyp_subst_tac 1), - (rtac notE 1), - (rtac less_lift2b 1), - (atac 1), + (res_inst_tac [("p","y")] liftE 1), + (hyp_subst_tac 1), + (rtac notE 1), + (rtac less_lift2b 1), + (atac 1), (asm_simp_tac Lift0_ss 1), - (rtac monofun_cfun_arg 1), - (hyp_subst_tac 1), - (etac (less_lift2c RS iffD1) 1) - ]); + (rtac monofun_cfun_arg 1), + (hyp_subst_tac 1), + (etac (less_lift2c RS iffD1) 1) + ]); (* ------------------------------------------------------------------------ *) (* Some kind of surjectivity lemma *) @@ -85,10 +85,10 @@ qed_goal "lift_lemma1" Lift2.thy "z=Iup(x) ==> Iup(Ilift(LAM x.x)(z)) = z" (fn prems => - [ - (cut_facts_tac prems 1), + [ + (cut_facts_tac prems 1), (asm_simp_tac Lift0_ss 1) - ]); + ]); (* ------------------------------------------------------------------------ *) (* ('a)u is a cpo *) @@ -98,61 +98,61 @@ "[|is_chain(Y);? i x.Y(i)=Iup(x)|] ==>\ \ range(Y) <<| Iup(lub(range(%i.(Ilift (LAM x.x) (Y(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 [("p","Y(i)")] liftE 1), - (res_inst_tac [("s","UU_lift"),("t","Y(i)")] subst 1), - (etac sym 1), - (rtac minimal_lift 1), - (res_inst_tac [("t","Y(i)")] (lift_lemma1 RS subst) 1), - (atac 1), - (rtac (less_lift2c RS iffD2) 1), - (rtac is_ub_thelub 1), - (etac (monofun_Ilift2 RS ch2ch_monofun) 1), - (strip_tac 1), - (res_inst_tac [("p","u")] liftE 1), - (etac exE 1), - (etac exE 1), - (res_inst_tac [("P","Y(i)<\ \ range(Y) <<| UU_lift" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_lubI 1), - (rtac conjI 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (res_inst_tac [("p","Y(i)")] liftE 1), - (res_inst_tac [("s","UU_lift"),("t","Y(i)")] ssubst 1), - (atac 1), - (rtac refl_less 1), - (rtac notE 1), - (dtac spec 1), - (dtac spec 1), - (atac 1), - (atac 1), - (strip_tac 1), - (rtac minimal_lift 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_lubI 1), + (rtac conjI 1), + (rtac ub_rangeI 1), + (rtac allI 1), + (res_inst_tac [("p","Y(i)")] liftE 1), + (res_inst_tac [("s","UU_lift"),("t","Y(i)")] ssubst 1), + (atac 1), + (rtac refl_less 1), + (rtac notE 1), + (dtac spec 1), + (dtac spec 1), + (atac 1), + (atac 1), + (strip_tac 1), + (rtac minimal_lift 1) + ]); val thelub_lift1a = lub_lift1a RS thelubI; (* @@ -167,17 +167,17 @@ *) qed_goal "cpo_lift" Lift2.thy - "is_chain(Y::nat=>('a)u) ==> ? x.range(Y) <<|x" + "is_chain(Y::nat=>('a)u) ==> ? x.range(Y) <<|x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac disjE 1), - (rtac exI 2), - (etac lub_lift1a 2), - (atac 2), - (rtac exI 2), - (etac lub_lift1b 2), - (atac 2), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac disjE 1), + (rtac exI 2), + (etac lub_lift1a 2), + (atac 2), + (rtac exI 2), + (etac lub_lift1b 2), + (atac 2), + (fast_tac HOL_cs 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Lift3.ML --- a/src/HOLCF/Lift3.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Lift3.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/lift3.ML +(* Title: HOLCF/lift3.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for lift3.thy @@ -14,17 +14,17 @@ qed_goal "less_lift3b" Lift3.thy "~ Iup(x) << UU" (fn prems => - [ - (rtac (inst_lift_pcpo RS ssubst) 1), - (rtac less_lift2b 1) - ]); + [ + (rtac (inst_lift_pcpo RS ssubst) 1), + (rtac less_lift2b 1) + ]); qed_goal "defined_Iup2" Lift3.thy "Iup(x) ~= UU" (fn prems => - [ - (rtac (inst_lift_pcpo RS ssubst) 1), - (rtac defined_Iup 1) - ]); + [ + (rtac (inst_lift_pcpo RS ssubst) 1), + (rtac defined_Iup 1) + ]); (* ------------------------------------------------------------------------ *) (* continuity for Iup *) @@ -32,28 +32,28 @@ qed_goal "contlub_Iup" Lift3.thy "contlub(Iup)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac trans 1), - (rtac (thelub_lift1a RS sym) 2), - (fast_tac HOL_cs 3), - (etac (monofun_Iup RS ch2ch_monofun) 2), - (res_inst_tac [("f","Iup")] arg_cong 1), - (rtac lub_equal 1), - (atac 1), - (rtac (monofun_Ilift2 RS ch2ch_monofun) 1), - (etac (monofun_Iup RS ch2ch_monofun) 1), + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac trans 1), + (rtac (thelub_lift1a RS sym) 2), + (fast_tac HOL_cs 3), + (etac (monofun_Iup RS ch2ch_monofun) 2), + (res_inst_tac [("f","Iup")] arg_cong 1), + (rtac lub_equal 1), + (atac 1), + (rtac (monofun_Ilift2 RS ch2ch_monofun) 1), + (etac (monofun_Iup RS ch2ch_monofun) 1), (asm_simp_tac Lift0_ss 1) - ]); + ]); qed_goal "cont_Iup" Lift3.thy "cont(Iup)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Iup 1), - (rtac contlub_Iup 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Iup 1), + (rtac contlub_Iup 1) + ]); (* ------------------------------------------------------------------------ *) @@ -62,83 +62,83 @@ qed_goal "contlub_Ilift1" Lift3.thy "contlub(Ilift)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac trans 1), - (rtac (thelub_fun RS sym) 2), - (etac (monofun_Ilift1 RS ch2ch_monofun) 2), - (rtac ext 1), - (res_inst_tac [("p","x")] liftE 1), + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac trans 1), + (rtac (thelub_fun RS sym) 2), + (etac (monofun_Ilift1 RS ch2ch_monofun) 2), + (rtac ext 1), + (res_inst_tac [("p","x")] liftE 1), (asm_simp_tac Lift0_ss 1), - (rtac (lub_const RS thelubI RS sym) 1), + (rtac (lub_const RS thelubI RS sym) 1), (asm_simp_tac Lift0_ss 1), - (etac contlub_cfun_fun 1) - ]); + (etac contlub_cfun_fun 1) + ]); qed_goal "contlub_Ilift2" Lift3.thy "contlub(Ilift(f))" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac disjE 1), - (rtac (thelub_lift1a RS ssubst) 2), - (atac 2), - (atac 2), + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac disjE 1), + (rtac (thelub_lift1a RS ssubst) 2), + (atac 2), + (atac 2), (asm_simp_tac Lift0_ss 2), - (rtac (thelub_lift1b RS ssubst) 3), - (atac 3), - (atac 3), - (fast_tac HOL_cs 1), + (rtac (thelub_lift1b RS ssubst) 3), + (atac 3), + (atac 3), + (fast_tac HOL_cs 1), (asm_simp_tac Lift0_ss 2), - (rtac (chain_UU_I_inverse RS sym) 2), - (rtac allI 2), - (res_inst_tac [("p","Y(i)")] liftE 2), + (rtac (chain_UU_I_inverse RS sym) 2), + (rtac allI 2), + (res_inst_tac [("p","Y(i)")] liftE 2), (asm_simp_tac Lift0_ss 2), - (rtac notE 2), - (dtac spec 2), - (etac spec 2), - (atac 2), - (rtac (contlub_cfun_arg RS ssubst) 1), - (etac (monofun_Ilift2 RS ch2ch_monofun) 1), - (rtac lub_equal2 1), - (rtac (monofun_fapp2 RS ch2ch_monofun) 2), - (etac (monofun_Ilift2 RS ch2ch_monofun) 2), - (etac (monofun_Ilift2 RS ch2ch_monofun) 2), - (rtac (chain_mono2 RS exE) 1), - (atac 2), - (etac exE 1), - (etac exE 1), - (rtac exI 1), - (res_inst_tac [("s","Iup(x)"),("t","Y(i)")] ssubst 1), - (atac 1), - (rtac defined_Iup2 1), - (rtac exI 1), - (strip_tac 1), - (res_inst_tac [("p","Y(i)")] liftE 1), + (rtac notE 2), + (dtac spec 2), + (etac spec 2), + (atac 2), + (rtac (contlub_cfun_arg RS ssubst) 1), + (etac (monofun_Ilift2 RS ch2ch_monofun) 1), + (rtac lub_equal2 1), + (rtac (monofun_fapp2 RS ch2ch_monofun) 2), + (etac (monofun_Ilift2 RS ch2ch_monofun) 2), + (etac (monofun_Ilift2 RS ch2ch_monofun) 2), + (rtac (chain_mono2 RS exE) 1), + (atac 2), + (etac exE 1), + (etac exE 1), + (rtac exI 1), + (res_inst_tac [("s","Iup(x)"),("t","Y(i)")] ssubst 1), + (atac 1), + (rtac defined_Iup2 1), + (rtac exI 1), + (strip_tac 1), + (res_inst_tac [("p","Y(i)")] liftE 1), (asm_simp_tac Lift0_ss 2), - (res_inst_tac [("P","Y(i) = UU")] notE 1), - (fast_tac HOL_cs 1), - (rtac (inst_lift_pcpo RS ssubst) 1), - (atac 1) - ]); + (res_inst_tac [("P","Y(i) = UU")] notE 1), + (fast_tac HOL_cs 1), + (rtac (inst_lift_pcpo RS ssubst) 1), + (atac 1) + ]); qed_goal "cont_Ilift1" Lift3.thy "cont(Ilift)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Ilift1 1), - (rtac contlub_Ilift1 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Ilift1 1), + (rtac contlub_Ilift1 1) + ]); qed_goal "cont_Ilift2" Lift3.thy "cont(Ilift(f))" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Ilift2 1), - (rtac contlub_Ilift2 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Ilift2 1), + (rtac contlub_Ilift2 1) + ]); (* ------------------------------------------------------------------------ *) @@ -147,200 +147,200 @@ qed_goalw "Exh_Lift1" Lift3.thy [up_def] "z = UU | (? x. z = up`x)" (fn prems => - [ + [ (simp_tac (Lift0_ss addsimps [cont_Iup]) 1), - (rtac (inst_lift_pcpo RS ssubst) 1), - (rtac Exh_Lift 1) - ]); + (rtac (inst_lift_pcpo RS ssubst) 1), + (rtac Exh_Lift 1) + ]); qed_goalw "inject_up" Lift3.thy [up_def] "up`x=up`y ==> x=y" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac inject_Iup 1), - (etac box_equals 1), + [ + (cut_facts_tac prems 1), + (rtac inject_Iup 1), + (etac box_equals 1), (simp_tac (Lift0_ss addsimps [cont_Iup]) 1), (simp_tac (Lift0_ss addsimps [cont_Iup]) 1) - ]); + ]); qed_goalw "defined_up" Lift3.thy [up_def] " up`x ~= UU" (fn prems => - [ + [ (simp_tac (Lift0_ss addsimps [cont_Iup]) 1), - (rtac defined_Iup2 1) - ]); + (rtac defined_Iup2 1) + ]); qed_goalw "liftE1" Lift3.thy [up_def] - "[| p=UU ==> Q; !!x. p=up`x==>Q|] ==>Q" + "[| p=UU ==> Q; !!x. p=up`x==>Q|] ==>Q" (fn prems => - [ - (rtac liftE 1), - (resolve_tac prems 1), - (etac (inst_lift_pcpo RS ssubst) 1), - (resolve_tac (tl prems) 1), + [ + (rtac liftE 1), + (resolve_tac prems 1), + (etac (inst_lift_pcpo RS ssubst) 1), + (resolve_tac (tl prems) 1), (asm_simp_tac (Lift0_ss addsimps [cont_Iup]) 1) - ]); + ]); qed_goalw "lift1" Lift3.thy [up_def,lift_def] "lift`f`UU=UU" (fn prems => - [ - (rtac (inst_lift_pcpo RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, - cont_Ilift2,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, - cont_Ilift2,cont2cont_CF1L]) 1)), - (simp_tac (Lift0_ss addsimps [cont_Iup,cont_Ilift1,cont_Ilift2]) 1) - ]); + [ + (rtac (inst_lift_pcpo RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), + (simp_tac (Lift0_ss addsimps [cont_Iup,cont_Ilift1,cont_Ilift2]) 1) + ]); qed_goalw "lift2" Lift3.thy [up_def,lift_def] "lift`f`(up`x)=f`x" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Iup 1), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, - cont_Ilift2,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Ilift2 1), - (simp_tac (Lift0_ss addsimps [cont_Iup,cont_Ilift1,cont_Ilift2]) 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Iup 1), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Ilift2 1), + (simp_tac (Lift0_ss addsimps [cont_Iup,cont_Ilift1,cont_Ilift2]) 1) + ]); qed_goalw "less_lift4b" Lift3.thy [up_def,lift_def] "~ up`x << UU" (fn prems => - [ + [ (simp_tac (Lift0_ss addsimps [cont_Iup]) 1), - (rtac less_lift3b 1) - ]); + (rtac less_lift3b 1) + ]); qed_goalw "less_lift4c" Lift3.thy [up_def,lift_def] - "(up`x << up`y) = (x< - [ + [ (simp_tac (Lift0_ss addsimps [cont_Iup]) 1), - (rtac less_lift2c 1) - ]); + (rtac less_lift2c 1) + ]); qed_goalw "thelub_lift2a" Lift3.thy [up_def,lift_def] "[| is_chain(Y); ? i x. Y(i) = up`x |] ==>\ \ lub(range(Y)) = up`(lub(range(%i. lift`(LAM x. x)`(Y i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, - cont_Ilift2,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, - cont_Ilift2,cont2cont_CF1L]) 1)), + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ext RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, - cont_Ilift2,cont2cont_CF1L]) 1)), - (rtac thelub_lift1a 1), - (atac 1), - (etac exE 1), - (etac exE 1), - (rtac exI 1), - (rtac exI 1), - (etac box_equals 1), - (rtac refl 1), - (simp_tac (Lift0_ss addsimps [cont_Iup]) 1) - ]); + (rtac (beta_cfun RS ext RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), + (rtac thelub_lift1a 1), + (atac 1), + (etac exE 1), + (etac exE 1), + (rtac exI 1), + (rtac exI 1), + (etac box_equals 1), + (rtac refl 1), + (simp_tac (Lift0_ss addsimps [cont_Iup]) 1) + ]); qed_goalw "thelub_lift2b" Lift3.thy [up_def,lift_def] "[| is_chain(Y); ! i x. Y(i) ~= up`x |] ==> lub(range(Y)) = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (inst_lift_pcpo RS ssubst) 1), - (rtac thelub_lift1b 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), + [ + (cut_facts_tac prems 1), + (rtac (inst_lift_pcpo RS ssubst) 1), + (rtac thelub_lift1b 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 (Lift0_ss addsimps [cont_Iup]) 1) - ]); + ]); qed_goal "lift_lemma2" Lift3.thy " (? x.z = up`x) = (z~=UU)" (fn prems => - [ - (rtac iffI 1), - (etac exE 1), - (hyp_subst_tac 1), - (rtac defined_up 1), - (res_inst_tac [("p","z")] liftE1 1), - (etac notE 1), - (atac 1), - (etac exI 1) - ]); + [ + (rtac iffI 1), + (etac exE 1), + (hyp_subst_tac 1), + (rtac defined_up 1), + (res_inst_tac [("p","z")] liftE1 1), + (etac notE 1), + (atac 1), + (etac exI 1) + ]); qed_goal "thelub_lift2a_rev" Lift3.thy "[| is_chain(Y); lub(range(Y)) = up`x |] ==> ? i x. Y(i) = up`x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac exE 1), - (rtac chain_UU_I_inverse2 1), - (rtac (lift_lemma2 RS iffD1) 1), - (etac exI 1), - (rtac exI 1), - (rtac (lift_lemma2 RS iffD2) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac exE 1), + (rtac chain_UU_I_inverse2 1), + (rtac (lift_lemma2 RS iffD1) 1), + (etac exI 1), + (rtac exI 1), + (rtac (lift_lemma2 RS iffD2) 1), + (atac 1) + ]); qed_goal "thelub_lift2b_rev" Lift3.thy "[| is_chain(Y); lub(range(Y)) = UU |] ==> ! i x. Y(i) ~= up`x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac allI 1), - (rtac (notex2all RS iffD1) 1), - (rtac contrapos 1), - (etac (lift_lemma2 RS iffD1) 2), - (rtac notnotI 1), - (rtac (chain_UU_I RS spec) 1), - (atac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac allI 1), + (rtac (notex2all RS iffD1) 1), + (rtac contrapos 1), + (etac (lift_lemma2 RS iffD1) 2), + (rtac notnotI 1), + (rtac (chain_UU_I RS spec) 1), + (atac 1), + (atac 1) + ]); qed_goal "thelub_lift3" Lift3.thy "is_chain(Y) ==> lub(range(Y)) = UU |\ \ lub(range(Y)) = up`(lub(range(%i. lift`(LAM x.x)`(Y i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac disjE 1), - (rtac disjI1 2), - (rtac thelub_lift2b 2), - (atac 2), - (atac 2), - (rtac disjI2 2), - (rtac thelub_lift2a 2), - (atac 2), - (atac 2), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac disjE 1), + (rtac disjI1 2), + (rtac thelub_lift2b 2), + (atac 2), + (atac 2), + (rtac disjI2 2), + (rtac thelub_lift2a 2), + (atac 2), + (atac 2), + (fast_tac HOL_cs 1) + ]); qed_goal "lift3" Lift3.thy "lift`up`x=x" (fn prems => - [ - (res_inst_tac [("p","x")] liftE1 1), - (asm_simp_tac ((simpset_of "Cfun3") addsimps [lift1,lift2]) 1), - (asm_simp_tac ((simpset_of "Cfun3") addsimps [lift1,lift2]) 1) - ]); + [ + (res_inst_tac [("p","x")] liftE1 1), + (asm_simp_tac ((simpset_of "Cfun3") addsimps [lift1,lift2]) 1), + (asm_simp_tac ((simpset_of "Cfun3") addsimps [lift1,lift2]) 1) + ]); (* ------------------------------------------------------------------------ *) (* install simplifier for ('a)u *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/One.ML --- a/src/HOLCF/One.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/One.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/one.thy +(* Title: HOLCF/one.thy ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for one.thy @@ -14,28 +14,28 @@ qed_goalw "Exh_one" One.thy [one_def] "z=UU | z = one" (fn prems => - [ - (res_inst_tac [("p","rep_one`z")] liftE1 1), - (rtac disjI1 1), - (rtac ((abs_one_iso RS allI) RS ((rep_one_iso RS allI) RS iso_strict ) - RS conjunct2 RS subst) 1), - (rtac (abs_one_iso RS subst) 1), - (etac cfun_arg_cong 1), - (rtac disjI2 1), - (rtac (abs_one_iso RS subst) 1), - (rtac cfun_arg_cong 1), - (rtac (unique_void2 RS subst) 1), - (atac 1) - ]); + [ + (res_inst_tac [("p","rep_one`z")] liftE1 1), + (rtac disjI1 1), + (rtac ((abs_one_iso RS allI) RS ((rep_one_iso RS allI) RS iso_strict ) + RS conjunct2 RS subst) 1), + (rtac (abs_one_iso RS subst) 1), + (etac cfun_arg_cong 1), + (rtac disjI2 1), + (rtac (abs_one_iso RS subst) 1), + (rtac cfun_arg_cong 1), + (rtac (unique_void2 RS subst) 1), + (atac 1) + ]); qed_goal "oneE" One.thy - "[| p=UU ==> Q; p = one ==>Q|] ==>Q" + "[| p=UU ==> Q; p = one ==>Q|] ==>Q" (fn prems => - [ - (rtac (Exh_one RS disjE) 1), - (eresolve_tac prems 1), - (eresolve_tac prems 1) - ]); + [ + (rtac (Exh_one RS disjE) 1), + (eresolve_tac prems 1), + (eresolve_tac prems 1) + ]); (* ------------------------------------------------------------------------ *) (* distinctness for type one : stored in a list *) @@ -44,23 +44,23 @@ val dist_less_one = [ prove_goalw One.thy [one_def] "~one << UU" (fn prems => - [ - (rtac classical3 1), - (rtac less_lift4b 1), - (rtac (rep_one_iso RS subst) 1), - (rtac (rep_one_iso RS subst) 1), - (rtac monofun_cfun_arg 1), - (etac ((abs_one_iso RS allI) RS ((rep_one_iso RS allI) RS iso_strict ) - RS conjunct2 RS ssubst) 1) - ]) + [ + (rtac classical3 1), + (rtac less_lift4b 1), + (rtac (rep_one_iso RS subst) 1), + (rtac (rep_one_iso RS subst) 1), + (rtac monofun_cfun_arg 1), + (etac ((abs_one_iso RS allI) RS ((rep_one_iso RS allI) RS iso_strict ) + RS conjunct2 RS ssubst) 1) + ]) ]; val dist_eq_one = [prove_goal One.thy "one~=UU" (fn prems => - [ - (rtac not_less2not_eq 1), - (resolve_tac dist_less_one 1) - ])]; + [ + (rtac not_less2not_eq 1), + (resolve_tac dist_less_one 1) + ])]; val dist_eq_one = dist_eq_one @ (map (fn thm => (thm RS not_sym)) dist_eq_one); @@ -70,15 +70,15 @@ qed_goalw "flat_one" One.thy [is_flat_def] "is_flat(one)" (fn prems => - [ - (rtac allI 1), - (rtac allI 1), - (res_inst_tac [("p","x")] oneE 1), - (Asm_simp_tac 1), - (res_inst_tac [("p","y")] oneE 1), - (asm_simp_tac (!simpset addsimps dist_less_one) 1), - (Asm_simp_tac 1) - ]); + [ + (rtac allI 1), + (rtac allI 1), + (res_inst_tac [("p","x")] oneE 1), + (Asm_simp_tac 1), + (res_inst_tac [("p","y")] oneE 1), + (asm_simp_tac (!simpset addsimps dist_less_one) 1), + (Asm_simp_tac 1) + ]); (* ------------------------------------------------------------------------ *) @@ -88,11 +88,11 @@ fun prover s = prove_goalw One.thy [one_when_def,one_def] s (fn prems => - [ - (simp_tac (!simpset addsimps [(rep_one_iso ), - (abs_one_iso RS allI) RS ((rep_one_iso RS allI) - RS iso_strict) RS conjunct1] )1) - ]); + [ + (simp_tac (!simpset addsimps [(rep_one_iso ), + (abs_one_iso RS allI) RS ((rep_one_iso RS allI) + RS iso_strict) RS conjunct1] )1) + ]); val one_when = map prover ["one_when`x`UU = UU","one_when`x`one = x"]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Pcpo.ML --- a/src/HOLCF/Pcpo.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Pcpo.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/pcpo.ML +(* Title: HOLCF/pcpo.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for pcpo.thy @@ -13,14 +13,14 @@ (* ------------------------------------------------------------------------ *) qed_goal "thelubE" Pcpo.thy - "[| is_chain(S);lub(range(S)) = (l::'a::pcpo)|] ==> range(S) <<| l " + "[| is_chain(S);lub(range(S)) = (l::'a::pcpo)|] ==> range(S) <<| l " (fn prems => - [ - (cut_facts_tac prems 1), - (hyp_subst_tac 1), - (rtac lubI 1), - (etac cpo 1) - ]); + [ + (cut_facts_tac prems 1), + (hyp_subst_tac 1), + (rtac lubI 1), + (etac cpo 1) + ]); (* ------------------------------------------------------------------------ *) (* Properties of the lub *) @@ -39,18 +39,18 @@ (* ------------------------------------------------------------------------ *) qed_goal "lub_mono" Pcpo.thy - "[|is_chain(C1::(nat=>'a::pcpo));is_chain(C2); ! k. C1(k) << C2(k)|]\ + "[|is_chain(C1::(nat=>'a::pcpo));is_chain(C2); ! k. C1(k) << C2(k)|]\ \ ==> lub(range(C1)) << lub(range(C2))" (fn prems => - [ - (cut_facts_tac prems 1), - (etac is_lub_thelub 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (rtac trans_less 1), - (etac spec 1), - (etac is_ub_thelub 1) - ]); + [ + (cut_facts_tac prems 1), + (etac is_lub_thelub 1), + (rtac ub_rangeI 1), + (rtac allI 1), + (rtac trans_less 1), + (etac spec 1), + (etac is_ub_thelub 1) + ]); (* ------------------------------------------------------------------------ *) (* the = relation between two chains is preserved by their lubs *) @@ -58,24 +58,24 @@ qed_goal "lub_equal" Pcpo.thy "[| is_chain(C1::(nat=>'a::pcpo));is_chain(C2);!k.C1(k)=C2(k)|]\ -\ ==> lub(range(C1))=lub(range(C2))" +\ ==> lub(range(C1))=lub(range(C2))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac antisym_less 1), - (rtac lub_mono 1), - (atac 1), - (atac 1), - (strip_tac 1), - (rtac (antisym_less_inverse RS conjunct1) 1), - (etac spec 1), - (rtac lub_mono 1), - (atac 1), - (atac 1), - (strip_tac 1), - (rtac (antisym_less_inverse RS conjunct2) 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac antisym_less 1), + (rtac lub_mono 1), + (atac 1), + (atac 1), + (strip_tac 1), + (rtac (antisym_less_inverse RS conjunct1) 1), + (etac spec 1), + (rtac lub_mono 1), + (atac 1), + (atac 1), + (strip_tac 1), + (rtac (antisym_less_inverse RS conjunct2) 1), + (etac spec 1) + ]); (* ------------------------------------------------------------------------ *) (* more results about mono and = of lubs of chains *) @@ -85,64 +85,64 @@ "[|? j.!i. j X(i::nat)=Y(i);is_chain(X::nat=>'a::pcpo);is_chain(Y)|]\ \ ==> lub(range(X))< - [ - (rtac exE 1), - (resolve_tac prems 1), - (rtac is_lub_thelub 1), - (resolve_tac prems 1), - (rtac ub_rangeI 1), - (strip_tac 1), - (res_inst_tac [("Q","x X(i)=Y(i);is_chain(X::nat=>'a::pcpo);is_chain(Y)|]\ \ ==> lub(range(X))=lub(range(Y))" (fn prems => - [ - (rtac antisym_less 1), - (rtac lub_mono2 1), - (REPEAT (resolve_tac prems 1)), - (cut_facts_tac prems 1), - (rtac lub_mono2 1), - (safe_tac HOL_cs), - (step_tac HOL_cs 1), - (safe_tac HOL_cs), - (rtac sym 1), - (fast_tac HOL_cs 1) - ]); + [ + (rtac antisym_less 1), + (rtac lub_mono2 1), + (REPEAT (resolve_tac prems 1)), + (cut_facts_tac prems 1), + (rtac lub_mono2 1), + (safe_tac HOL_cs), + (step_tac HOL_cs 1), + (safe_tac HOL_cs), + (rtac sym 1), + (fast_tac HOL_cs 1) + ]); qed_goal "lub_mono3" Pcpo.thy "[|is_chain(Y::nat=>'a::pcpo);is_chain(X);\ \! i. ? j. Y(i)<< X(j)|]==> lub(range(Y))< - [ - (cut_facts_tac prems 1), - (rtac is_lub_thelub 1), - (atac 1), - (rtac ub_rangeI 1), - (strip_tac 1), - (etac allE 1), - (etac exE 1), - (rtac trans_less 1), - (rtac is_ub_thelub 2), - (atac 2), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_lub_thelub 1), + (atac 1), + (rtac ub_rangeI 1), + (strip_tac 1), + (etac allE 1), + (etac exE 1), + (rtac trans_less 1), + (rtac is_ub_thelub 2), + (atac 2), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* usefull lemmas about UU *) @@ -150,96 +150,96 @@ qed_goal "eq_UU_iff" Pcpo.thy "(x=UU)=(x< - [ - (rtac iffI 1), - (hyp_subst_tac 1), - (rtac refl_less 1), - (rtac antisym_less 1), - (atac 1), - (rtac minimal 1) - ]); + [ + (rtac iffI 1), + (hyp_subst_tac 1), + (rtac refl_less 1), + (rtac antisym_less 1), + (atac 1), + (rtac minimal 1) + ]); qed_goal "UU_I" Pcpo.thy "x << UU ==> x = UU" (fn prems => - [ - (rtac (eq_UU_iff RS ssubst) 1), - (resolve_tac prems 1) - ]); + [ + (rtac (eq_UU_iff RS ssubst) 1), + (resolve_tac prems 1) + ]); qed_goal "not_less2not_eq" Pcpo.thy "~x< ~x=y" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac classical3 1), - (atac 1), - (hyp_subst_tac 1), - (rtac refl_less 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac classical3 1), + (atac 1), + (hyp_subst_tac 1), + (rtac refl_less 1) + ]); qed_goal "chain_UU_I" Pcpo.thy - "[|is_chain(Y);lub(range(Y))=UU|] ==> ! i.Y(i)=UU" + "[|is_chain(Y);lub(range(Y))=UU|] ==> ! i.Y(i)=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac allI 1), - (rtac antisym_less 1), - (rtac minimal 2), - (etac subst 1), - (etac is_ub_thelub 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac allI 1), + (rtac antisym_less 1), + (rtac minimal 2), + (etac subst 1), + (etac is_ub_thelub 1) + ]); qed_goal "chain_UU_I_inverse" Pcpo.thy - "!i.Y(i::nat)=UU ==> lub(range(Y::(nat=>'a::pcpo)))=UU" + "!i.Y(i::nat)=UU ==> lub(range(Y::(nat=>'a::pcpo)))=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac lub_chain_maxelem 1), - (rtac exI 1), - (etac spec 1), - (rtac allI 1), - (rtac (antisym_less_inverse RS conjunct1) 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac lub_chain_maxelem 1), + (rtac exI 1), + (etac spec 1), + (rtac allI 1), + (rtac (antisym_less_inverse RS conjunct1) 1), + (etac spec 1) + ]); qed_goal "chain_UU_I_inverse2" Pcpo.thy - "~lub(range(Y::(nat=>'a::pcpo)))=UU ==> ? i.~ Y(i)=UU" + "~lub(range(Y::(nat=>'a::pcpo)))=UU ==> ? i.~ Y(i)=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (notall2ex RS iffD1) 1), - (rtac swap 1), - (rtac chain_UU_I_inverse 2), - (etac notnotD 2), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (notall2ex RS iffD1) 1), + (rtac swap 1), + (rtac chain_UU_I_inverse 2), + (etac notnotD 2), + (atac 1) + ]); qed_goal "notUU_I" Pcpo.thy "[| x< ~y=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (etac contrapos 1), - (rtac UU_I 1), - (hyp_subst_tac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac contrapos 1), + (rtac UU_I 1), + (hyp_subst_tac 1), + (atac 1) + ]); qed_goal "chain_mono2" Pcpo.thy "[|? j.~Y(j)=UU;is_chain(Y::nat=>'a::pcpo)|]\ \ ==> ? j.!i.j~Y(i)=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (safe_tac HOL_cs), - (step_tac HOL_cs 1), - (strip_tac 1), - (rtac notUU_I 1), - (atac 2), - (etac (chain_mono RS mp) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (safe_tac HOL_cs), + (step_tac HOL_cs 1), + (strip_tac 1), + (rtac notUU_I 1), + (atac 2), + (etac (chain_mono RS mp) 1), + (atac 1) + ]); @@ -250,15 +250,15 @@ qed_goal "unique_void2" Pcpo.thy "(x::void)=UU" (fn prems => - [ - (rtac (inst_void_pcpo RS ssubst) 1), - (rtac (Rep_Void_inverse RS subst) 1), - (rtac (Rep_Void_inverse RS subst) 1), - (rtac arg_cong 1), - (rtac box_equals 1), - (rtac refl 1), - (rtac (unique_void RS sym) 1), - (rtac (unique_void RS sym) 1) - ]); + [ + (rtac (inst_void_pcpo RS ssubst) 1), + (rtac (Rep_Void_inverse RS subst) 1), + (rtac (Rep_Void_inverse RS subst) 1), + (rtac arg_cong 1), + (rtac box_equals 1), + (rtac refl 1), + (rtac (unique_void RS sym) 1), + (rtac (unique_void RS sym) 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Porder.ML --- a/src/HOLCF/Porder.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Porder.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/porder.thy +(* Title: HOLCF/porder.thy ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for theory porder.thy @@ -16,76 +16,76 @@ qed_goal "antisym_less_inverse" Porder.thy "x=y ==> x << y & y << x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - ((rtac subst 1) THEN (rtac refl_less 2) THEN (atac 1)), - ((rtac subst 1) THEN (rtac refl_less 2) THEN (etac sym 1)) - ]); + [ + (cut_facts_tac prems 1), + (rtac conjI 1), + ((rtac subst 1) THEN (rtac refl_less 2) THEN (atac 1)), + ((rtac subst 1) THEN (rtac refl_less 2) THEN (etac sym 1)) + ]); qed_goal "box_less" Porder.thy "[| a << b; c << a; b << d|] ==> c << d" (fn prems => - [ - (cut_facts_tac prems 1), - (etac trans_less 1), - (etac trans_less 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac trans_less 1), + (etac trans_less 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* lubs are unique *) (* ------------------------------------------------------------------------ *) qed_goalw "unique_lub " Porder.thy [is_lub, is_ub] - "[| S <<| x ; S <<| y |] ==> x=y" + "[| 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)) - ]); + [ + (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)) + ]); (* ------------------------------------------------------------------------ *) (* chains are monotone functions *) (* ------------------------------------------------------------------------ *) qed_goalw "chain_mono" Porder.thy [is_chain] - " is_chain(F) ==> x F(x)< x F(x)< - [ - (cut_facts_tac prems 1), - (nat_ind_tac "y" 1), - (rtac impI 1), - (etac less_zeroE 1), - (rtac (less_Suc_eq RS ssubst) 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) - ]); + [ + (cut_facts_tac prems 1), + (nat_ind_tac "y" 1), + (rtac impI 1), + (etac less_zeroE 1), + (rtac (less_Suc_eq RS ssubst) 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) + ]); qed_goal "chain_mono3" Porder.thy - "[| is_chain(F); x <= y |] ==> F(x) << F(y)" + "[| is_chain(F); x <= y |] ==> F(x) << F(y)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (le_imp_less_or_eq RS disjE) 1), - (atac 1), - (etac (chain_mono RS mp) 1), - (atac 1), - (hyp_subst_tac 1), - (rtac refl_less 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (le_imp_less_or_eq RS disjE) 1), + (atac 1), + (etac (chain_mono RS mp) 1), + (atac 1), + (hyp_subst_tac 1), + (rtac refl_less 1) + ]); (* ------------------------------------------------------------------------ *) @@ -95,26 +95,26 @@ qed_goalw "chain_is_tord" Porder.thy [is_tord,range_def] "is_chain(F) ==> is_tord(range(F))" (fn prems => - [ - (cut_facts_tac prems 1), - (REPEAT (rtac allI 1 )), - (rtac (mem_Collect_eq RS ssubst) 1), - (rtac (mem_Collect_eq RS ssubst) 1), - (strip_tac 1), - (etac conjE 1), - (REPEAT (etac exE 1)), - (REPEAT (hyp_subst_tac 1)), - (rtac nat_less_cases 1), - (rtac disjI1 1), - (etac (chain_mono RS mp) 1), - (atac 1), - (rtac disjI1 1), - (hyp_subst_tac 1), - (rtac refl_less 1), - (rtac disjI2 1), - (etac (chain_mono RS mp) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (REPEAT (rtac allI 1 )), + (rtac (mem_Collect_eq RS ssubst) 1), + (rtac (mem_Collect_eq RS ssubst) 1), + (strip_tac 1), + (etac conjE 1), + (REPEAT (etac exE 1)), + (REPEAT (hyp_subst_tac 1)), + (rtac nat_less_cases 1), + (rtac disjI1 1), + (etac (chain_mono RS mp) 1), + (atac 1), + (rtac disjI1 1), + (hyp_subst_tac 1), + (rtac refl_less 1), + (rtac disjI2 1), + (etac (chain_mono RS mp) 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* technical lemmas about lub and is_lub *) @@ -122,38 +122,38 @@ qed_goal "lubI" Porder.thy "(? x. M <<| x) ==> M <<| lub(M)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (lub RS ssubst) 1), - (etac selectI3 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (lub RS ssubst) 1), + (etac selectI3 1) + ]); qed_goal "lubE" Porder.thy " M <<| lub(M) ==> ? x. M <<| x" (fn prems => - [ - (cut_facts_tac prems 1), - (etac exI 1) - ]); + [ + (cut_facts_tac prems 1), + (etac exI 1) + ]); qed_goal "lub_eq" Porder.thy - "(? x. M <<| x) = M <<| lub(M)" + "(? x. M <<| x) = M <<| lub(M)" (fn prems => - [ - (rtac (lub RS ssubst) 1), - (rtac (select_eq_Ex RS subst) 1), - (rtac refl 1) - ]); + [ + (rtac (lub RS ssubst) 1), + (rtac (select_eq_Ex RS subst) 1), + (rtac refl 1) + ]); qed_goal "thelubI" Porder.thy " M <<| l ==> lub(M) = l" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac unique_lub 1), - (rtac (lub RS ssubst) 1), - (etac selectI 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac unique_lub 1), + (rtac (lub RS ssubst) 1), + (etac selectI 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) @@ -161,60 +161,60 @@ (* ------------------------------------------------------------------------ *) qed_goalw "is_lubE" Porder.thy [is_lub] - "S <<| x ==> S <| x & (! u. S <| u --> x << u)" + "S <<| x ==> S <| x & (! u. S <| u --> x << u)" (fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (atac 1) + ]); qed_goalw "is_lubI" Porder.thy [is_lub] - "S <| x & (! u. S <| u --> x << u) ==> S <<| x" + "S <| x & (! u. S <| u --> x << u) ==> S <<| x" (fn prems => - [ - (cut_facts_tac prems 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (atac 1) + ]); qed_goalw "is_chainE" Porder.thy [is_chain] "is_chain(F) ==> ! i. F(i) << F(Suc(i))" (fn prems => - [ - (cut_facts_tac prems 1), - (atac 1)]); + [ + (cut_facts_tac prems 1), + (atac 1)]); qed_goalw "is_chainI" Porder.thy [is_chain] "! i. F(i) << F(Suc(i)) ==> is_chain(F) " (fn prems => - [ - (cut_facts_tac prems 1), - (atac 1)]); + [ + (cut_facts_tac prems 1), + (atac 1)]); (* ------------------------------------------------------------------------ *) (* technical lemmas about (least) upper bounds of chains *) (* ------------------------------------------------------------------------ *) qed_goalw "ub_rangeE" Porder.thy [is_ub] - "range(S) <| x ==> ! i. S(i) << x" + "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) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (rtac mp 1), + (etac spec 1), + (rtac rangeI 1) + ]); qed_goalw "ub_rangeI" Porder.thy [is_ub] - "! i. S(i) << x ==> range(S) <| x" + "! 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) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (etac rangeE 1), + (hyp_subst_tac 1), + (etac spec 1) + ]); val is_ub_lub = (is_lubE RS conjunct1 RS ub_rangeE RS spec); (* range(?S1) <<| ?x1 ==> ?S1(?x) << ?x1 *) @@ -232,28 +232,28 @@ qed_goal "less_void" Porder.thy "((u1::void) << u2) = (u1 = u2)" (fn prems => - [ - (rtac (inst_void_po RS ssubst) 1), - (rewrite_goals_tac [less_void_def]), - (rtac iffI 1), - (rtac injD 1), - (atac 2), - (rtac inj_inverseI 1), - (rtac Rep_Void_inverse 1), - (etac arg_cong 1) - ]); + [ + (rtac (inst_void_po RS ssubst) 1), + (rewtac less_void_def), + (rtac iffI 1), + (rtac injD 1), + (atac 2), + (rtac inj_inverseI 1), + (rtac Rep_Void_inverse 1), + (etac arg_cong 1) + ]); (* ------------------------------------------------------------------------ *) (* void is pointed. The least element is UU_void *) (* ------------------------------------------------------------------------ *) -qed_goal "minimal_void" Porder.thy "UU_void << x" +qed_goal "minimal_void" Porder.thy "UU_void << x" (fn prems => - [ - (rtac (inst_void_po RS ssubst) 1), - (rewrite_goals_tac [less_void_def]), - (simp_tac (!simpset addsimps [unique_void]) 1) - ]); + [ + (rtac (inst_void_po RS ssubst) 1), + (rewtac less_void_def), + (simp_tac (!simpset addsimps [unique_void]) 1) + ]); (* ------------------------------------------------------------------------ *) (* UU_void is the trivial lub of all chains in void *) @@ -261,16 +261,16 @@ qed_goalw "lub_void" Porder.thy [is_lub] "M <<| UU_void" (fn prems => - [ - (rtac conjI 1), - (rewrite_goals_tac [is_ub]), - (strip_tac 1), - (rtac (inst_void_po RS ssubst) 1), - (rewrite_goals_tac [less_void_def]), - (simp_tac (!simpset addsimps [unique_void]) 1), - (strip_tac 1), - (rtac minimal_void 1) - ]); + [ + (rtac conjI 1), + (rewtac is_ub), + (strip_tac 1), + (rtac (inst_void_po RS ssubst) 1), + (rewtac less_void_def), + (simp_tac (!simpset addsimps [unique_void]) 1), + (strip_tac 1), + (rtac minimal_void 1) + ]); (* ------------------------------------------------------------------------ *) (* lub(?M) = UU_void *) @@ -283,13 +283,13 @@ (* ------------------------------------------------------------------------ *) qed_goal "cpo_void" Porder.thy - "is_chain((S::nat=>void)) ==> ? x. range(S) <<| x " + "is_chain((S::nat=>void)) ==> ? x. range(S) <<| x " (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("x","UU_void")] exI 1), - (rtac lub_void 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("x","UU_void")] exI 1), + (rtac lub_void 1) + ]); (* ------------------------------------------------------------------------ *) (* end of prototype lemmas for class pcpo *) @@ -301,72 +301,72 @@ (* ------------------------------------------------------------------------ *) qed_goalw "lub_finch1" Porder.thy [max_in_chain_def] - "[| is_chain(C) ; max_in_chain i C|] ==> range(C) <<| C(i)" + "[| is_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) - ]); + [ + (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) + ]); qed_goalw "lub_finch2" Porder.thy [finite_chain_def] - "finite_chain(C) ==> range(C) <<| C(@ i. max_in_chain i C)" + "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 selectI3 1), - (etac conjunct2 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac lub_finch1 1), + (etac conjunct1 1), + (rtac selectI3 1), + (etac conjunct2 1) + ]); qed_goal "bin_chain" Porder.thy "x< is_chain (%i. if i=0 then x else y)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_chainI 1), - (rtac allI 1), - (nat_ind_tac "i" 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (rtac refl_less 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_chainI 1), + (rtac allI 1), + (nat_ind_tac "i" 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1), + (rtac refl_less 1) + ]); qed_goalw "bin_chainmax" Porder.thy [max_in_chain_def,le_def] - "x< max_in_chain (Suc 0) (%i. if (i=0) then x else y)" + "x< max_in_chain (Suc 0) (%i. if (i=0) then x else y)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac allI 1), - (nat_ind_tac "j" 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac allI 1), + (nat_ind_tac "j" 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1) + ]); qed_goal "lub_bin_chain" Porder.thy - "x << y ==> range(%i. if (i=0) then x else y) <<| y" + "x << y ==> range(%i. if (i=0) then x else y) <<| y" (fn prems=> - [ (cut_facts_tac prems 1), - (res_inst_tac [("s","if (Suc 0) = 0 then x else y")] subst 1), - (rtac lub_finch1 2), - (etac bin_chain 2), - (etac bin_chainmax 2), - (Simp_tac 1) - ]); + [ (cut_facts_tac prems 1), + (res_inst_tac [("s","if (Suc 0) = 0 then x else y")] subst 1), + (rtac lub_finch1 2), + (etac bin_chain 2), + (etac bin_chainmax 2), + (Simp_tac 1) + ]); (* ------------------------------------------------------------------------ *) (* the maximal element in a chain is its lub *) @@ -375,17 +375,17 @@ qed_goal "lub_chain_maxelem" Porder.thy "[|? i.Y(i)=c;!i.Y(i)< lub(range(Y)) = c" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac thelubI 1), - (rtac is_lubI 1), - (rtac conjI 1), - (etac ub_rangeI 1), - (strip_tac 1), - (etac exE 1), - (hyp_subst_tac 1), - (etac (ub_rangeE RS spec) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac thelubI 1), + (rtac is_lubI 1), + (rtac conjI 1), + (etac ub_rangeI 1), + (strip_tac 1), + (etac exE 1), + (hyp_subst_tac 1), + (etac (ub_rangeE RS spec) 1) + ]); (* ------------------------------------------------------------------------ *) (* the lub of a constant chain is the constant *) @@ -393,15 +393,15 @@ qed_goal "lub_const" Porder.thy "range(%x.c) <<| c" (fn prems => - [ - (rtac is_lubI 1), - (rtac conjI 1), - (rtac ub_rangeI 1), - (strip_tac 1), - (rtac refl_less 1), - (strip_tac 1), - (etac (ub_rangeE RS spec) 1) - ]); + [ + (rtac is_lubI 1), + (rtac conjI 1), + (rtac ub_rangeI 1), + (strip_tac 1), + (rtac refl_less 1), + (strip_tac 1), + (etac (ub_rangeE RS spec) 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/ROOT.ML --- a/src/HOLCF/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,7 +1,7 @@ -(* Title: HOLCF/ROOT +(* Title: HOLCF/ROOT ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen ROOT file for the conservative extension of HOL by the LCF logic. Should be executed in subdirectory HOLCF. @@ -40,4 +40,4 @@ print_depth 100; -val HOLCF_build_completed = (); (*indicate successful build*) +val HOLCF_build_completed = (); (*indicate successful build*) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Sprod0.ML --- a/src/HOLCF/Sprod0.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Sprod0.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/sprod0.thy +(* Title: HOLCF/sprod0.thy ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for theory sprod0.thy @@ -13,20 +13,20 @@ (* ------------------------------------------------------------------------ *) qed_goalw "SprodI" Sprod0.thy [Sprod_def] - "(Spair_Rep a b):Sprod" + "(Spair_Rep a b):Sprod" (fn prems => - [ - (EVERY1 [rtac CollectI, rtac exI,rtac exI, rtac refl]) - ]); + [ + (EVERY1 [rtac CollectI, rtac exI,rtac exI, rtac refl]) + ]); qed_goal "inj_onto_Abs_Sprod" Sprod0.thy - "inj_onto Abs_Sprod Sprod" + "inj_onto Abs_Sprod Sprod" (fn prems => - [ - (rtac inj_onto_inverseI 1), - (etac Abs_Sprod_inverse 1) - ]); + [ + (rtac inj_onto_inverseI 1), + (etac Abs_Sprod_inverse 1) + ]); (* ------------------------------------------------------------------------ *) @@ -37,27 +37,27 @@ 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) - ]); + [ + (cut_facts_tac prems 1), + (rtac ext 1), + (rtac ext 1), + (rtac iffI 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); 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 => - [ - (res_inst_tac [("Q","a=UU|b=UU")] classical2 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) - ]); + [ + (res_inst_tac [("Q","a=UU|b=UU")] classical2 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) + ]); (* ------------------------------------------------------------------------ *) @@ -67,26 +67,26 @@ 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) - ]); + [ + (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) + ]); qed_goalw "inject_Ispair" Sprod0.thy [Ispair_def] - "[|~aa=UU ; ~ba=UU ; Ispair a b = Ispair aa ba |] ==> a=aa & b=ba" + "[|~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_onto_Abs_Sprod RS inj_ontoD) 1), - (rtac SprodI 1), - (rtac SprodI 1) - ]); + [ + (cut_facts_tac prems 1), + (etac inject_Spair_Rep 1), + (atac 1), + (etac (inj_onto_Abs_Sprod RS inj_ontoD) 1), + (rtac SprodI 1), + (rtac SprodI 1) + ]); (* ------------------------------------------------------------------------ *) @@ -96,62 +96,62 @@ 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) - ]); + [ + (cut_facts_tac prems 1), + (etac (strict_Spair_Rep RS arg_cong) 1) + ]); qed_goalw "strict_Ispair1" Sprod0.thy [Ispair_def] - "Ispair UU b = Ispair UU UU" + "Ispair UU b = Ispair UU UU" (fn prems => - [ - (rtac (strict_Spair_Rep RS arg_cong) 1), - (rtac disjI1 1), - (rtac refl 1) - ]); + [ + (rtac (strict_Spair_Rep RS arg_cong) 1), + (rtac disjI1 1), + (rtac refl 1) + ]); qed_goalw "strict_Ispair2" Sprod0.thy [Ispair_def] - "Ispair a UU = Ispair UU UU" + "Ispair a UU = Ispair UU UU" (fn prems => - [ - (rtac (strict_Spair_Rep RS arg_cong) 1), - (rtac disjI2 1), - (rtac refl 1) - ]); + [ + (rtac (strict_Spair_Rep RS arg_cong) 1), + (rtac disjI2 1), + (rtac refl 1) + ]); qed_goal "strict_Ispair_rev" Sprod0.thy - "~Ispair x y = Ispair UU UU ==> ~x=UU & ~y=UU" + "~Ispair x y = Ispair UU UU ==> ~x=UU & ~y=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (de_morgan1 RS ssubst) 1), - (etac contrapos 1), - (etac strict_Ispair 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (de_morgan1 RS ssubst) 1), + (etac contrapos 1), + (etac strict_Ispair 1) + ]); qed_goalw "defined_Ispair_rev" Sprod0.thy [Ispair_def] - "Ispair a b = Ispair UU UU ==> (a = UU | b = UU)" + "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_onto_Abs_Sprod RS inj_ontoD) 1), - (atac 1), - (rtac SprodI 1), - (rtac SprodI 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac defined_Spair_Rep_rev 1), + (rtac (inj_onto_Abs_Sprod RS inj_ontoD) 1), + (atac 1), + (rtac SprodI 1), + (rtac SprodI 1) + ]); 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_morgan1 RS iffD1) 1), - (etac conjI 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac contrapos 1), + (etac defined_Ispair_rev 2), + (rtac (de_morgan1 RS iffD1) 1), + (etac conjI 1), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) @@ -159,27 +159,27 @@ (* ------------------------------------------------------------------------ *) qed_goalw "Exh_Sprod" Sprod0.thy [Ispair_def] - "z=Ispair UU UU | (? a b. z=Ispair a b & a~=UU & b~=UU)" + "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_morgan1 RS ssubst) 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) - ]); + [ + (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_morgan1 RS ssubst) 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) + ]); (* ------------------------------------------------------------------------ *) (* general elimination rule for strict product *) @@ -188,17 +188,17 @@ 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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) @@ -206,134 +206,134 @@ (* ------------------------------------------------------------------------ *) qed_goalw "strict_Isfst" Sprod0.thy [Isfst_def] - "p=Ispair UU UU ==> Isfst p = UU" + "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)) - ]); + [ + (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)) + ]); qed_goal "strict_Isfst1" Sprod0.thy - "Isfst(Ispair UU y) = UU" + "Isfst(Ispair UU y) = UU" (fn prems => - [ - (rtac (strict_Ispair1 RS ssubst) 1), - (rtac strict_Isfst 1), - (rtac refl 1) - ]); + [ + (rtac (strict_Ispair1 RS ssubst) 1), + (rtac strict_Isfst 1), + (rtac refl 1) + ]); qed_goal "strict_Isfst2" Sprod0.thy - "Isfst(Ispair x UU) = UU" + "Isfst(Ispair x UU) = UU" (fn prems => - [ - (rtac (strict_Ispair2 RS ssubst) 1), - (rtac strict_Isfst 1), - (rtac refl 1) - ]); + [ + (rtac (strict_Ispair2 RS ssubst) 1), + (rtac strict_Isfst 1), + (rtac refl 1) + ]); qed_goalw "strict_Issnd" Sprod0.thy [Issnd_def] - "p=Ispair UU UU ==>Issnd p=UU" + "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)) - ]); + [ + (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)) + ]); qed_goal "strict_Issnd1" Sprod0.thy - "Issnd(Ispair UU y) = UU" + "Issnd(Ispair UU y) = UU" (fn prems => - [ - (rtac (strict_Ispair1 RS ssubst) 1), - (rtac strict_Issnd 1), - (rtac refl 1) - ]); + [ + (rtac (strict_Ispair1 RS ssubst) 1), + (rtac strict_Issnd 1), + (rtac refl 1) + ]); qed_goal "strict_Issnd2" Sprod0.thy - "Issnd(Ispair x UU) = UU" + "Issnd(Ispair x UU) = UU" (fn prems => - [ - (rtac (strict_Ispair2 RS ssubst) 1), - (rtac strict_Issnd 1), - (rtac refl 1) - ]); + [ + (rtac (strict_Ispair2 RS ssubst) 1), + (rtac strict_Issnd 1), + (rtac refl 1) + ]); qed_goalw "Isfst" Sprod0.thy [Isfst_def] - "[|x~=UU ;y~=UU |] ==> Isfst(Ispair x y) = x" + "[|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) - ]); + [ + (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) + ]); qed_goalw "Issnd" Sprod0.thy [Issnd_def] - "[|x~=UU ;y~=UU |] ==> Issnd(Ispair x y) = y" + "[|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) - ]); + [ + (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) + ]); 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) - ]); + [ + (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) + ]); 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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) @@ -346,17 +346,17 @@ Isfst2,Issnd2]; qed_goal "defined_IsfstIssnd" Sprod0.thy - "p~=Ispair UU UU ==> Isfst p ~= UU & Issnd p ~= UU" + "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), + [ + (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) - ]); + ]); (* ------------------------------------------------------------------------ *) @@ -364,13 +364,13 @@ (* ------------------------------------------------------------------------ *) qed_goal "surjective_pairing_Sprod" Sprod0.thy - "z = Ispair(Isfst z)(Issnd z)" + "z = Ispair(Isfst z)(Issnd z)" (fn prems => - [ - (res_inst_tac [("z1","z")] (Exh_Sprod RS disjE) 1), + [ + (res_inst_tac [("z1","z")] (Exh_Sprod RS disjE) 1), (asm_simp_tac Sprod0_ss 1), - (etac exE 1), - (etac exE 1), + (etac exE 1), + (etac exE 1), (asm_simp_tac Sprod0_ss 1) - ]); + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Sprod1.ML --- a/src/HOLCF/Sprod1.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Sprod1.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/sprod1.ML +(* Title: HOLCF/sprod1.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for theory sprod1.thy @@ -14,77 +14,77 @@ qed_goalw "less_sprod1a" Sprod1.thy [less_sprod_def] - "p1=Ispair UU UU ==> less_sprod p1 p2" + "p1=Ispair UU UU ==> less_sprod p1 p2" (fn prems => - [ - (cut_facts_tac prems 1), + [ + (cut_facts_tac prems 1), (asm_simp_tac HOL_ss 1) - ]); + ]); qed_goalw "less_sprod1b" Sprod1.thy [less_sprod_def] "p1~=Ispair UU UU ==> \ \ less_sprod p1 p2 = ( Isfst p1 << Isfst p2 & Issnd p1 << Issnd p2)" (fn prems => - [ - (cut_facts_tac prems 1), + [ + (cut_facts_tac prems 1), (asm_simp_tac HOL_ss 1) - ]); + ]); qed_goal "less_sprod2a" Sprod1.thy - "less_sprod(Ispair x y)(Ispair UU UU) ==> x = UU | y = UU" + "less_sprod(Ispair x y)(Ispair UU UU) ==> x = UU | y = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (excluded_middle RS disjE) 1), - (atac 2), - (rtac disjI1 1), - (rtac antisym_less 1), - (rtac minimal 2), - (res_inst_tac [("s","Isfst(Ispair x y)"),("t","x")] subst 1), - (rtac Isfst 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1), - (res_inst_tac [("s","Isfst(Ispair UU UU)"),("t","UU")] subst 1), + [ + (cut_facts_tac prems 1), + (rtac (excluded_middle RS disjE) 1), + (atac 2), + (rtac disjI1 1), + (rtac antisym_less 1), + (rtac minimal 2), + (res_inst_tac [("s","Isfst(Ispair x y)"),("t","x")] subst 1), + (rtac Isfst 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1), + (res_inst_tac [("s","Isfst(Ispair UU UU)"),("t","UU")] subst 1), (simp_tac Sprod0_ss 1), - (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct1) 1), - (REPEAT (fast_tac HOL_cs 1)) - ]); + (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct1) 1), + (REPEAT (fast_tac HOL_cs 1)) + ]); qed_goal "less_sprod2b" Sprod1.thy "less_sprod p (Ispair UU UU) ==> p = Ispair UU UU" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p")] IsprodE 1), - (atac 1), - (hyp_subst_tac 1), - (rtac strict_Ispair 1), - (etac less_sprod2a 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","p")] IsprodE 1), + (atac 1), + (hyp_subst_tac 1), + (rtac strict_Ispair 1), + (etac less_sprod2a 1) + ]); qed_goal "less_sprod2c" Sprod1.thy "[|less_sprod(Ispair xa ya)(Ispair x y);\ \ xa ~= UU ; ya ~= UU; x ~= UU ; y ~= UU |] ==> xa << x & ya << y" (fn prems => - [ - (rtac conjI 1), - (res_inst_tac [("s","Isfst(Ispair xa ya)"),("t","xa")] subst 1), + [ + (rtac conjI 1), + (res_inst_tac [("s","Isfst(Ispair xa ya)"),("t","xa")] subst 1), (simp_tac (Sprod0_ss addsimps prems)1), - (res_inst_tac [("s","Isfst(Ispair x y)"),("t","x")] subst 1), + (res_inst_tac [("s","Isfst(Ispair x y)"),("t","x")] subst 1), (simp_tac (Sprod0_ss addsimps prems)1), - (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct1) 1), - (resolve_tac prems 1), - (resolve_tac prems 1), + (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct1) 1), + (resolve_tac prems 1), + (resolve_tac prems 1), (simp_tac (Sprod0_ss addsimps prems)1), - (res_inst_tac [("s","Issnd(Ispair xa ya)"),("t","ya")] subst 1), + (res_inst_tac [("s","Issnd(Ispair xa ya)"),("t","ya")] subst 1), (simp_tac (Sprod0_ss addsimps prems)1), - (res_inst_tac [("s","Issnd(Ispair x y)"),("t","y")] subst 1), + (res_inst_tac [("s","Issnd(Ispair x y)"),("t","y")] subst 1), (simp_tac (Sprod0_ss addsimps prems)1), - (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct2) 1), - (resolve_tac prems 1), - (resolve_tac prems 1), + (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct2) 1), + (resolve_tac prems 1), + (resolve_tac prems 1), (simp_tac (Sprod0_ss addsimps prems)1) - ]); + ]); (* ------------------------------------------------------------------------ *) (* less_sprod is a partial order on Sprod *) @@ -92,85 +92,85 @@ qed_goal "refl_less_sprod" Sprod1.thy "less_sprod p p" (fn prems => - [ - (res_inst_tac [("p","p")] IsprodE 1), - (etac less_sprod1a 1), - (hyp_subst_tac 1), - (rtac (less_sprod1b RS ssubst) 1), - (rtac defined_Ispair 1), - (REPEAT (fast_tac (HOL_cs addIs [refl_less]) 1)) - ]); + [ + (res_inst_tac [("p","p")] IsprodE 1), + (etac less_sprod1a 1), + (hyp_subst_tac 1), + (rtac (less_sprod1b RS ssubst) 1), + (rtac defined_Ispair 1), + (REPEAT (fast_tac (HOL_cs addIs [refl_less]) 1)) + ]); qed_goal "antisym_less_sprod" Sprod1.thy "[|less_sprod p1 p2;less_sprod p2 p1|] ==> p1=p2" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p1")] IsprodE 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p2")] IsprodE 1), - (hyp_subst_tac 1), - (rtac refl 1), - (hyp_subst_tac 1), - (rtac (strict_Ispair RS sym) 1), - (etac less_sprod2a 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p2")] IsprodE 1), - (hyp_subst_tac 1), - (rtac (strict_Ispair) 1), - (etac less_sprod2a 1), - (hyp_subst_tac 1), - (res_inst_tac [("x1","x"),("y1","xa"),("x","y"),("y","ya")] (arg_cong RS cong) 1), - (rtac antisym_less 1), + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","p1")] IsprodE 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p2")] IsprodE 1), + (hyp_subst_tac 1), + (rtac refl 1), + (hyp_subst_tac 1), + (rtac (strict_Ispair RS sym) 1), + (etac less_sprod2a 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p2")] IsprodE 1), + (hyp_subst_tac 1), + (rtac (strict_Ispair) 1), + (etac less_sprod2a 1), + (hyp_subst_tac 1), + (res_inst_tac [("x1","x"),("y1","xa"),("x","y"),("y","ya")] (arg_cong RS cong) 1), + (rtac antisym_less 1), (asm_simp_tac (HOL_ss addsimps [less_sprod2c RS conjunct1]) 1), (asm_simp_tac (HOL_ss addsimps [less_sprod2c RS conjunct1]) 1), - (rtac antisym_less 1), + (rtac antisym_less 1), (asm_simp_tac (HOL_ss addsimps [less_sprod2c RS conjunct2]) 1), (asm_simp_tac (HOL_ss addsimps [less_sprod2c RS conjunct2]) 1) - ]); + ]); qed_goal "trans_less_sprod" Sprod1.thy "[|less_sprod (p1::'a**'b) p2;less_sprod p2 p3|] ==> less_sprod p1 p3" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p1")] IsprodE 1), - (etac less_sprod1a 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p3")] IsprodE 1), - (hyp_subst_tac 1), - (res_inst_tac [("s","p2"),("t","Ispair (UU::'a)(UU::'b)")] subst 1), - (etac less_sprod2b 1), - (atac 1), - (hyp_subst_tac 1), - (res_inst_tac [("Q","p2=Ispair(UU::'a)(UU::'b)")] - (excluded_middle RS disjE) 1), - (rtac (defined_Ispair RS less_sprod1b RS ssubst) 1), - (REPEAT (atac 1)), - (rtac conjI 1), - (res_inst_tac [("y","Isfst(p2)")] trans_less 1), - (rtac conjunct1 1), - (rtac (less_sprod1b RS subst) 1), - (rtac defined_Ispair 1), - (REPEAT (atac 1)), - (rtac conjunct1 1), - (rtac (less_sprod1b RS subst) 1), - (REPEAT (atac 1)), - (res_inst_tac [("y","Issnd(p2)")] trans_less 1), - (rtac conjunct2 1), - (rtac (less_sprod1b RS subst) 1), - (rtac defined_Ispair 1), - (REPEAT (atac 1)), - (rtac conjunct2 1), - (rtac (less_sprod1b RS subst) 1), - (REPEAT (atac 1)), - (hyp_subst_tac 1), - (res_inst_tac [("s","Ispair(UU::'a)(UU::'b)"),("t","Ispair x y")] - subst 1), - (etac (less_sprod2b RS sym) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","p1")] IsprodE 1), + (etac less_sprod1a 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p3")] IsprodE 1), + (hyp_subst_tac 1), + (res_inst_tac [("s","p2"),("t","Ispair (UU::'a)(UU::'b)")] subst 1), + (etac less_sprod2b 1), + (atac 1), + (hyp_subst_tac 1), + (res_inst_tac [("Q","p2=Ispair(UU::'a)(UU::'b)")] + (excluded_middle RS disjE) 1), + (rtac (defined_Ispair RS less_sprod1b RS ssubst) 1), + (REPEAT (atac 1)), + (rtac conjI 1), + (res_inst_tac [("y","Isfst(p2)")] trans_less 1), + (rtac conjunct1 1), + (rtac (less_sprod1b RS subst) 1), + (rtac defined_Ispair 1), + (REPEAT (atac 1)), + (rtac conjunct1 1), + (rtac (less_sprod1b RS subst) 1), + (REPEAT (atac 1)), + (res_inst_tac [("y","Issnd(p2)")] trans_less 1), + (rtac conjunct2 1), + (rtac (less_sprod1b RS subst) 1), + (rtac defined_Ispair 1), + (REPEAT (atac 1)), + (rtac conjunct2 1), + (rtac (less_sprod1b RS subst) 1), + (REPEAT (atac 1)), + (hyp_subst_tac 1), + (res_inst_tac [("s","Ispair(UU::'a)(UU::'b)"),("t","Ispair x y")] + subst 1), + (etac (less_sprod2b RS sym) 1), + (atac 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Sprod2.ML --- a/src/HOLCF/Sprod2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Sprod2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/sprod2.ML +(* Title: HOLCF/sprod2.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for sprod2.thy @@ -14,47 +14,47 @@ (* ------------------------------------------------------------------------ *) qed_goal "less_sprod3a" Sprod2.thy - "p1=Ispair UU UU ==> p1 << p2" + "p1=Ispair UU UU ==> p1 << p2" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (inst_sprod_po RS ssubst) 1), - (etac less_sprod1a 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (inst_sprod_po RS ssubst) 1), + (etac less_sprod1a 1) + ]); qed_goal "less_sprod3b" Sprod2.thy "p1~=Ispair UU UU ==>\ -\ (p1< - [ - (cut_facts_tac prems 1), - (rtac (inst_sprod_po RS ssubst) 1), - (etac less_sprod1b 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (inst_sprod_po RS ssubst) 1), + (etac less_sprod1b 1) + ]); qed_goal "less_sprod4b" Sprod2.thy - "p << Ispair UU UU ==> p = Ispair UU UU" + "p << Ispair UU UU ==> p = Ispair UU UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac less_sprod2b 1), - (etac (inst_sprod_po RS subst) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac less_sprod2b 1), + (etac (inst_sprod_po RS subst) 1) + ]); val less_sprod4a = (less_sprod4b RS defined_Ispair_rev); (* Ispair ?a ?b << Ispair UU UU ==> ?a = UU | ?b = UU *) qed_goal "less_sprod4c" Sprod2.thy "[|Ispair xa ya << Ispair x y; xa~=UU; ya~=UU; x~=UU; y~=UU|] ==>\ -\ xa< - [ - (cut_facts_tac prems 1), - (rtac less_sprod2c 1), - (etac (inst_sprod_po RS subst) 1), - (REPEAT (atac 1)) - ]); + [ + (cut_facts_tac prems 1), + (rtac less_sprod2c 1), + (etac (inst_sprod_po RS subst) 1), + (REPEAT (atac 1)) + ]); (* ------------------------------------------------------------------------ *) (* type sprod is pointed *) @@ -62,10 +62,10 @@ qed_goal "minimal_sprod" Sprod2.thy "Ispair UU UU << p" (fn prems => - [ - (rtac less_sprod3a 1), - (rtac refl 1) - ]); + [ + (rtac less_sprod3a 1), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) (* Ispair is monotone in both arguments *) @@ -73,93 +73,93 @@ qed_goalw "monofun_Ispair1" Sprod2.thy [monofun] "monofun(Ispair)" (fn prems => - [ - (strip_tac 1), - (rtac (less_fun RS iffD2) 1), - (strip_tac 1), - (res_inst_tac [("Q", - " Ispair y xa = Ispair UU UU")] (excluded_middle RS disjE) 1), - (res_inst_tac [("Q", - " Ispair x xa = Ispair UU UU")] (excluded_middle RS disjE) 1), - (rtac (less_sprod3b RS iffD2) 1), - (atac 1), - (rtac conjI 1), - (rtac (Isfst RS ssubst) 1), - (etac (strict_Ispair_rev RS conjunct1) 1), - (etac (strict_Ispair_rev RS conjunct2) 1), - (rtac (Isfst RS ssubst) 1), - (etac (strict_Ispair_rev RS conjunct1) 1), - (etac (strict_Ispair_rev RS conjunct2) 1), - (atac 1), - (rtac (Issnd RS ssubst) 1), - (etac (strict_Ispair_rev RS conjunct1) 1), - (etac (strict_Ispair_rev RS conjunct2) 1), - (rtac (Issnd RS ssubst) 1), - (etac (strict_Ispair_rev RS conjunct1) 1), - (etac (strict_Ispair_rev RS conjunct2) 1), - (rtac refl_less 1), - (etac less_sprod3a 1), - (res_inst_tac [("Q", - " Ispair x xa = Ispair UU UU")] (excluded_middle RS disjE) 1), - (etac less_sprod3a 2), - (res_inst_tac [("P","Ispair y xa = Ispair UU UU")] notE 1), - (atac 2), - (rtac defined_Ispair 1), - (etac notUU_I 1), - (etac (strict_Ispair_rev RS conjunct1) 1), - (etac (strict_Ispair_rev RS conjunct2) 1) - ]); + [ + (strip_tac 1), + (rtac (less_fun RS iffD2) 1), + (strip_tac 1), + (res_inst_tac [("Q", + " Ispair y xa = Ispair UU UU")] (excluded_middle RS disjE) 1), + (res_inst_tac [("Q", + " Ispair x xa = Ispair UU UU")] (excluded_middle RS disjE) 1), + (rtac (less_sprod3b RS iffD2) 1), + (atac 1), + (rtac conjI 1), + (rtac (Isfst RS ssubst) 1), + (etac (strict_Ispair_rev RS conjunct1) 1), + (etac (strict_Ispair_rev RS conjunct2) 1), + (rtac (Isfst RS ssubst) 1), + (etac (strict_Ispair_rev RS conjunct1) 1), + (etac (strict_Ispair_rev RS conjunct2) 1), + (atac 1), + (rtac (Issnd RS ssubst) 1), + (etac (strict_Ispair_rev RS conjunct1) 1), + (etac (strict_Ispair_rev RS conjunct2) 1), + (rtac (Issnd RS ssubst) 1), + (etac (strict_Ispair_rev RS conjunct1) 1), + (etac (strict_Ispair_rev RS conjunct2) 1), + (rtac refl_less 1), + (etac less_sprod3a 1), + (res_inst_tac [("Q", + " Ispair x xa = Ispair UU UU")] (excluded_middle RS disjE) 1), + (etac less_sprod3a 2), + (res_inst_tac [("P","Ispair y xa = Ispair UU UU")] notE 1), + (atac 2), + (rtac defined_Ispair 1), + (etac notUU_I 1), + (etac (strict_Ispair_rev RS conjunct1) 1), + (etac (strict_Ispair_rev RS conjunct2) 1) + ]); qed_goalw "monofun_Ispair2" Sprod2.thy [monofun] "monofun(Ispair(x))" (fn prems => - [ - (strip_tac 1), - (res_inst_tac [("Q", - " Ispair x y = Ispair UU UU")] (excluded_middle RS disjE) 1), - (res_inst_tac [("Q", - " Ispair x xa = Ispair UU UU")] (excluded_middle RS disjE) 1), - (rtac (less_sprod3b RS iffD2) 1), - (atac 1), - (rtac conjI 1), - (rtac (Isfst RS ssubst) 1), - (etac (strict_Ispair_rev RS conjunct1) 1), - (etac (strict_Ispair_rev RS conjunct2) 1), - (rtac (Isfst RS ssubst) 1), - (etac (strict_Ispair_rev RS conjunct1) 1), - (etac (strict_Ispair_rev RS conjunct2) 1), - (rtac refl_less 1), - (rtac (Issnd RS ssubst) 1), - (etac (strict_Ispair_rev RS conjunct1) 1), - (etac (strict_Ispair_rev RS conjunct2) 1), - (rtac (Issnd RS ssubst) 1), - (etac (strict_Ispair_rev RS conjunct1) 1), - (etac (strict_Ispair_rev RS conjunct2) 1), - (atac 1), - (etac less_sprod3a 1), - (res_inst_tac [("Q", - " Ispair x xa = Ispair UU UU")] (excluded_middle RS disjE) 1), - (etac less_sprod3a 2), - (res_inst_tac [("P","Ispair x y = Ispair UU UU")] notE 1), - (atac 2), - (rtac defined_Ispair 1), - (etac (strict_Ispair_rev RS conjunct1) 1), - (etac notUU_I 1), - (etac (strict_Ispair_rev RS conjunct2) 1) - ]); + [ + (strip_tac 1), + (res_inst_tac [("Q", + " Ispair x y = Ispair UU UU")] (excluded_middle RS disjE) 1), + (res_inst_tac [("Q", + " Ispair x xa = Ispair UU UU")] (excluded_middle RS disjE) 1), + (rtac (less_sprod3b RS iffD2) 1), + (atac 1), + (rtac conjI 1), + (rtac (Isfst RS ssubst) 1), + (etac (strict_Ispair_rev RS conjunct1) 1), + (etac (strict_Ispair_rev RS conjunct2) 1), + (rtac (Isfst RS ssubst) 1), + (etac (strict_Ispair_rev RS conjunct1) 1), + (etac (strict_Ispair_rev RS conjunct2) 1), + (rtac refl_less 1), + (rtac (Issnd RS ssubst) 1), + (etac (strict_Ispair_rev RS conjunct1) 1), + (etac (strict_Ispair_rev RS conjunct2) 1), + (rtac (Issnd RS ssubst) 1), + (etac (strict_Ispair_rev RS conjunct1) 1), + (etac (strict_Ispair_rev RS conjunct2) 1), + (atac 1), + (etac less_sprod3a 1), + (res_inst_tac [("Q", + " Ispair x xa = Ispair UU UU")] (excluded_middle RS disjE) 1), + (etac less_sprod3a 2), + (res_inst_tac [("P","Ispair x y = Ispair UU UU")] notE 1), + (atac 2), + (rtac defined_Ispair 1), + (etac (strict_Ispair_rev RS conjunct1) 1), + (etac notUU_I 1), + (etac (strict_Ispair_rev RS conjunct2) 1) + ]); 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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) @@ -168,57 +168,57 @@ qed_goalw " monofun_Isfst" Sprod2.thy [monofun] "monofun(Isfst)" (fn prems => - [ - (strip_tac 1), - (res_inst_tac [("p","x")] IsprodE 1), - (hyp_subst_tac 1), - (rtac trans_less 1), - (rtac minimal 2), - (rtac (strict_Isfst1 RS ssubst) 1), - (rtac refl_less 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","y")] IsprodE 1), - (hyp_subst_tac 1), - (res_inst_tac [("t","Isfst(Ispair xa ya)")] subst 1), - (rtac refl_less 2), - (etac (less_sprod4b RS sym RS arg_cong) 1), - (hyp_subst_tac 1), - (rtac (Isfst RS ssubst) 1), - (atac 1), - (atac 1), - (rtac (Isfst RS ssubst) 1), - (atac 1), - (atac 1), - (etac (less_sprod4c RS conjunct1) 1), - (REPEAT (atac 1)) - ]); + [ + (strip_tac 1), + (res_inst_tac [("p","x")] IsprodE 1), + (hyp_subst_tac 1), + (rtac trans_less 1), + (rtac minimal 2), + (rtac (strict_Isfst1 RS ssubst) 1), + (rtac refl_less 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","y")] IsprodE 1), + (hyp_subst_tac 1), + (res_inst_tac [("t","Isfst(Ispair xa ya)")] subst 1), + (rtac refl_less 2), + (etac (less_sprod4b RS sym RS arg_cong) 1), + (hyp_subst_tac 1), + (rtac (Isfst RS ssubst) 1), + (atac 1), + (atac 1), + (rtac (Isfst RS ssubst) 1), + (atac 1), + (atac 1), + (etac (less_sprod4c RS conjunct1) 1), + (REPEAT (atac 1)) + ]); qed_goalw "monofun_Issnd" Sprod2.thy [monofun] "monofun(Issnd)" (fn prems => - [ - (strip_tac 1), - (res_inst_tac [("p","x")] IsprodE 1), - (hyp_subst_tac 1), - (rtac trans_less 1), - (rtac minimal 2), - (rtac (strict_Issnd1 RS ssubst) 1), - (rtac refl_less 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","y")] IsprodE 1), - (hyp_subst_tac 1), - (res_inst_tac [("t","Issnd(Ispair xa ya)")] subst 1), - (rtac refl_less 2), - (etac (less_sprod4b RS sym RS arg_cong) 1), - (hyp_subst_tac 1), - (rtac (Issnd RS ssubst) 1), - (atac 1), - (atac 1), - (rtac (Issnd RS ssubst) 1), - (atac 1), - (atac 1), - (etac (less_sprod4c RS conjunct2) 1), - (REPEAT (atac 1)) - ]); + [ + (strip_tac 1), + (res_inst_tac [("p","x")] IsprodE 1), + (hyp_subst_tac 1), + (rtac trans_less 1), + (rtac minimal 2), + (rtac (strict_Issnd1 RS ssubst) 1), + (rtac refl_less 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","y")] IsprodE 1), + (hyp_subst_tac 1), + (res_inst_tac [("t","Issnd(Ispair xa ya)")] subst 1), + (rtac refl_less 2), + (etac (less_sprod4b RS sym RS arg_cong) 1), + (hyp_subst_tac 1), + (rtac (Issnd RS ssubst) 1), + (atac 1), + (atac 1), + (rtac (Issnd RS ssubst) 1), + (atac 1), + (atac 1), + (etac (less_sprod4c RS conjunct2) 1), + (REPEAT (atac 1)) + ]); (* ------------------------------------------------------------------------ *) @@ -229,40 +229,40 @@ "[|is_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 is_lubI 1), - (rtac conjI 1), - (rtac ub_rangeI 1), - (rtac allI 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) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_lubI 1), + (rtac conjI 1), + (rtac ub_rangeI 1), + (rtac allI 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) + ]); val thelub_sprod = (lub_sprod RS thelubI); qed_goal "cpo_sprod" Sprod2.thy - "is_chain(S::nat=>'a**'b)==>? x.range(S)<<| x" + "is_chain(S::nat=>'a**'b)==>? x.range(S)<<| x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac exI 1), - (etac lub_sprod 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac exI 1), + (etac lub_sprod 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Sprod3.ML --- a/src/HOLCF/Sprod3.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Sprod3.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/sprod3.thy +(* Title: HOLCF/sprod3.thy ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for Sprod3.thy @@ -18,36 +18,36 @@ \ 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), + [ + (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 (notall2ex 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), + (rtac sym 1), + (rtac lub_chain_maxelem 1), + (res_inst_tac [("P","%j.~Y(j)=UU")] exE 1), + (rtac (notall2ex 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), + (res_inst_tac [("s","UU"),("t","Y(i)")] subst 1), + (etac sym 1), (asm_simp_tac Sprod0_ss 1), (rtac minimal 1) - ]); + ]); qed_goal "sprod3_lemma2" Sprod3.thy "[| is_chain(Y); x ~= UU; lub(range(Y)) = UU |] ==>\ @@ -55,20 +55,20 @@ \ 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), + [ + (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) - ]); + (etac (chain_UU_I RS spec) 1), + (atac 1) + ]); qed_goal "sprod3_lemma3" Sprod3.thy @@ -77,45 +77,45 @@ \ 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), + [ + (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) - ]); + ]); qed_goal "contlub_Ispair1" Sprod3.thy "contlub(Ispair)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (expand_fun_eq RS iffD2) 1), - (strip_tac 1), - (rtac (lub_fun RS thelubI RS ssubst) 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) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac (expand_fun_eq RS iffD2) 1), + (strip_tac 1), + (rtac (lub_fun RS thelubI RS ssubst) 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) + ]); qed_goal "sprod3_lemma4" Sprod3.thy "[| is_chain(Y); x ~= UU; lub(range(Y)) ~= UU |] ==>\ @@ -123,35 +123,35 @@ \ 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 (notall2ex 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), + [ + (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 (notall2ex 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), + (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), + (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) - ]); + ]); qed_goal "sprod3_lemma5" Sprod3.thy "[| is_chain(Y); x ~= UU; lub(range(Y)) = UU |] ==>\ @@ -159,20 +159,20 @@ \ 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), + [ + (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) - ]); + (etac (chain_UU_I RS spec) 1), + (atac 1) + ]); qed_goal "sprod3_lemma6" Sprod3.thy "[| is_chain(Y); x = UU |] ==>\ @@ -180,130 +180,130 @@ \ 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), + [ + (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) - ]); + ]); qed_goal "contlub_Ispair2" Sprod3.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) - ]); + [ + (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) + ]); qed_goal "cont_Ispair1" Sprod3.thy "cont(Ispair)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Ispair1 1), - (rtac contlub_Ispair1 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Ispair1 1), + (rtac contlub_Ispair1 1) + ]); qed_goal "cont_Ispair2" Sprod3.thy "cont(Ispair(x))" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Ispair2 1), - (rtac contlub_Ispair2 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Ispair2 1), + (rtac contlub_Ispair2 1) + ]); qed_goal "contlub_Isfst" Sprod3.thy "contlub(Isfst)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (lub_sprod RS thelubI RS ssubst) 1), - (atac 1), - (res_inst_tac [("Q","lub(range(%i. Issnd(Y(i))))=UU")] - (excluded_middle RS disjE) 1), + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac (lub_sprod RS thelubI RS ssubst) 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), + (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), - (rtac notnotI 1), - (rtac (chain_UU_I RS spec) 1), - (rtac (monofun_Issnd RS ch2ch_monofun) 1), - (atac 1), - (atac 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), + (rtac notnotI 1), + (rtac (chain_UU_I RS spec) 1), + (rtac (monofun_Issnd RS ch2ch_monofun) 1), + (atac 1), + (atac 1) + ]); qed_goal "contlub_Issnd" Sprod3.thy "contlub(Issnd)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac (lub_sprod RS thelubI RS ssubst) 1), - (atac 1), - (res_inst_tac [("Q","lub(range(%i. Isfst(Y(i))))=UU")] - (excluded_middle RS disjE) 1), + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac (lub_sprod RS thelubI RS ssubst) 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), + (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), - (rtac notnotI 1), - (rtac (chain_UU_I RS spec) 1), - (rtac (monofun_Isfst RS ch2ch_monofun) 1), - (atac 1), - (atac 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), + (rtac notnotI 1), + (rtac (chain_UU_I RS spec) 1), + (rtac (monofun_Isfst RS ch2ch_monofun) 1), + (atac 1), + (atac 1) + ]); qed_goal "cont_Isfst" Sprod3.thy "cont(Isfst)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Isfst 1), - (rtac contlub_Isfst 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Isfst 1), + (rtac contlub_Isfst 1) + ]); qed_goal "cont_Issnd" Sprod3.thy "cont(Issnd)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Issnd 1), - (rtac contlub_Issnd 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Issnd 1), + (rtac contlub_Issnd 1) + ]); (* -------------------------------------------------------------------------- @@ -314,310 +314,310 @@ qed_goal "spair_eq" Sprod3.thy "[|x1=x2;y1=y2|] ==> (|x1,y1|) = (|x2,y2|)" (fn prems => - [ - (cut_facts_tac prems 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (fast_tac HOL_cs 1) + ]); (* ------------------------------------------------------------------------ *) (* convert all lemmas to the continuous versions *) (* ------------------------------------------------------------------------ *) qed_goalw "beta_cfun_sprod" Sprod3.thy [spair_def] - "(LAM x y.Ispair x y)`a`b = Ispair a b" + "(LAM x y.Ispair x y)`a`b = Ispair a b" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (cont_tac 1), - (rtac cont_Ispair2 1), - (rtac cont2cont_CF1L 1), - (rtac cont_Ispair1 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Ispair2 1), - (rtac refl 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (cont_tac 1), + (rtac cont_Ispair2 1), + (rtac cont2cont_CF1L 1), + (rtac cont_Ispair1 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Ispair2 1), + (rtac refl 1) + ]); qed_goalw "inject_spair" Sprod3.thy [spair_def] - "[| aa~=UU ; ba~=UU ; (|a,b|)=(|aa,ba|) |] ==> a=aa & b=ba" + "[| 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) - ]); + [ + (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) + ]); qed_goalw "inst_sprod_pcpo2" Sprod3.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) - ]); + [ + (rtac sym 1), + (rtac trans 1), + (rtac beta_cfun_sprod 1), + (rtac sym 1), + (rtac inst_sprod_pcpo 1) + ]); qed_goalw "strict_spair" Sprod3.thy [spair_def] - "(a=UU | b=UU) ==> (|a,b|)=UU" + "(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) - ]); + [ + (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) + ]); qed_goalw "strict_spair1" Sprod3.thy [spair_def] "(|UU,b|) = UU" (fn prems => - [ - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac trans 1), - (rtac (inst_sprod_pcpo RS sym) 2), - (rtac strict_Ispair1 1) - ]); + [ + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac trans 1), + (rtac (inst_sprod_pcpo RS sym) 2), + (rtac strict_Ispair1 1) + ]); qed_goalw "strict_spair2" Sprod3.thy [spair_def] "(|a,UU|) = UU" (fn prems => - [ - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac trans 1), - (rtac (inst_sprod_pcpo RS sym) 2), - (rtac strict_Ispair2 1) - ]); + [ + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac trans 1), + (rtac (inst_sprod_pcpo RS sym) 2), + (rtac strict_Ispair2 1) + ]); qed_goalw "strict_spair_rev" Sprod3.thy [spair_def] - "(|x,y|)~=UU ==> ~x=UU & ~y=UU" + "(|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) - ]); + [ + (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) + ]); qed_goalw "defined_spair_rev" Sprod3.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) - ]); + [ + (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) + ]); qed_goalw "defined_spair" Sprod3.thy [spair_def] - "[|a~=UU; b~=UU|] ==> (|a,b|) ~= UU" + "[|a~=UU; b~=UU|] ==> (|a,b|) ~= UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac (inst_sprod_pcpo RS ssubst) 1), - (etac defined_Ispair 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac (inst_sprod_pcpo RS ssubst) 1), + (etac defined_Ispair 1), + (atac 1) + ]); qed_goalw "Exh_Sprod2" Sprod3.thy [spair_def] - "z=UU | (? a b. z=(|a,b|) & a~=UU & b~=UU)" + "z=UU | (? a b. z=(|a,b|) & a~=UU & b~=UU)" (fn prems => - [ - (rtac (Exh_Sprod RS disjE) 1), - (rtac disjI1 1), - (rtac (inst_sprod_pcpo RS ssubst) 1), - (atac 1), - (rtac disjI2 1), - (etac exE 1), - (etac exE 1), - (rtac exI 1), - (rtac exI 1), - (rtac conjI 1), - (rtac (beta_cfun_sprod RS ssubst) 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); + [ + (rtac (Exh_Sprod RS disjE) 1), + (rtac disjI1 1), + (rtac (inst_sprod_pcpo RS ssubst) 1), + (atac 1), + (rtac disjI2 1), + (etac exE 1), + (etac exE 1), + (rtac exI 1), + (rtac exI 1), + (rtac conjI 1), + (rtac (beta_cfun_sprod RS ssubst) 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); qed_goalw "sprodE" Sprod3.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), - (rtac (inst_sprod_pcpo RS ssubst) 1), - (atac 1), - (resolve_tac prems 1), - (atac 2), - (atac 2), - (rtac (beta_cfun_sprod RS ssubst) 1), - (atac 1) - ]); + [ + (rtac IsprodE 1), + (resolve_tac prems 1), + (rtac (inst_sprod_pcpo RS ssubst) 1), + (atac 1), + (resolve_tac prems 1), + (atac 2), + (atac 2), + (rtac (beta_cfun_sprod RS ssubst) 1), + (atac 1) + ]); qed_goalw "strict_sfst" Sprod3.thy [sfst_def] - "p=UU==>sfst`p=UU" + "p=UU==>sfst`p=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Isfst 1), - (rtac strict_Isfst 1), - (rtac (inst_sprod_pcpo RS subst) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Isfst 1), + (rtac strict_Isfst 1), + (rtac (inst_sprod_pcpo RS subst) 1), + (atac 1) + ]); qed_goalw "strict_sfst1" Sprod3.thy [sfst_def,spair_def] - "sfst`(|UU,y|) = UU" + "sfst`(|UU,y|) = UU" (fn prems => - [ - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Isfst 1), - (rtac strict_Isfst1 1) - ]); + [ + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Isfst 1), + (rtac strict_Isfst1 1) + ]); qed_goalw "strict_sfst2" Sprod3.thy [sfst_def,spair_def] - "sfst`(|x,UU|) = UU" + "sfst`(|x,UU|) = UU" (fn prems => - [ - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Isfst 1), - (rtac strict_Isfst2 1) - ]); + [ + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Isfst 1), + (rtac strict_Isfst2 1) + ]); qed_goalw "strict_ssnd" Sprod3.thy [ssnd_def] - "p=UU==>ssnd`p=UU" + "p=UU==>ssnd`p=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Issnd 1), - (rtac strict_Issnd 1), - (rtac (inst_sprod_pcpo RS subst) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Issnd 1), + (rtac strict_Issnd 1), + (rtac (inst_sprod_pcpo RS subst) 1), + (atac 1) + ]); qed_goalw "strict_ssnd1" Sprod3.thy [ssnd_def,spair_def] - "ssnd`(|UU,y|) = UU" + "ssnd`(|UU,y|) = UU" (fn prems => - [ - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Issnd 1), - (rtac strict_Issnd1 1) - ]); + [ + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Issnd 1), + (rtac strict_Issnd1 1) + ]); qed_goalw "strict_ssnd2" Sprod3.thy [ssnd_def,spair_def] - "ssnd`(|x,UU|) = UU" + "ssnd`(|x,UU|) = UU" (fn prems => - [ - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Issnd 1), - (rtac strict_Issnd2 1) - ]); + [ + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Issnd 1), + (rtac strict_Issnd2 1) + ]); qed_goalw "sfst2" Sprod3.thy [sfst_def,spair_def] - "y~=UU ==>sfst`(|x,y|)=x" + "y~=UU ==>sfst`(|x,y|)=x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Isfst 1), - (etac Isfst2 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Isfst 1), + (etac Isfst2 1) + ]); qed_goalw "ssnd2" Sprod3.thy [ssnd_def,spair_def] - "x~=UU ==>ssnd`(|x,y|)=y" + "x~=UU ==>ssnd`(|x,y|)=y" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Issnd 1), - (etac Issnd2 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Issnd 1), + (etac Issnd2 1) + ]); qed_goalw "defined_sfstssnd" Sprod3.thy [sfst_def,ssnd_def,spair_def] - "p~=UU ==> sfst`p ~=UU & ssnd`p ~=UU" + "p~=UU ==> sfst`p ~=UU & ssnd`p ~=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Issnd 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Isfst 1), - (rtac defined_IsfstIssnd 1), - (rtac (inst_sprod_pcpo RS subst) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Issnd 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Isfst 1), + (rtac defined_IsfstIssnd 1), + (rtac (inst_sprod_pcpo RS subst) 1), + (atac 1) + ]); qed_goalw "surjective_pairing_Sprod2" Sprod3.thy - [sfst_def,ssnd_def,spair_def] "(|sfst`p , ssnd`p|) = p" + [sfst_def,ssnd_def,spair_def] "(|sfst`p , ssnd`p|) = p" (fn prems => - [ - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Issnd 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Isfst 1), - (rtac (surjective_pairing_Sprod RS sym) 1) - ]); + [ + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Issnd 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Isfst 1), + (rtac (surjective_pairing_Sprod RS sym) 1) + ]); qed_goalw "less_sprod5b" Sprod3.thy [sfst_def,ssnd_def,spair_def] "p1~=UU ==> (p1< - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Issnd 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Issnd 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Isfst 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Isfst 1), - (rtac less_sprod3b 1), - (rtac (inst_sprod_pcpo RS subst) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Issnd 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Issnd 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Isfst 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Isfst 1), + (rtac less_sprod3b 1), + (rtac (inst_sprod_pcpo RS subst) 1), + (atac 1) + ]); qed_goalw "less_sprod5c" Sprod3.thy [sfst_def,ssnd_def,spair_def] "[|(|xa,ya|) << (|x,y|);xa~=UU;ya~=UU;x~=UU;y~=UU|] ==>xa< - [ - (cut_facts_tac prems 1), - (rtac less_sprod4c 1), - (REPEAT (atac 2)), - (rtac (beta_cfun_sprod RS subst) 1), - (rtac (beta_cfun_sprod RS subst) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac less_sprod4c 1), + (REPEAT (atac 2)), + (rtac (beta_cfun_sprod RS subst) 1), + (rtac (beta_cfun_sprod RS subst) 1), + (atac 1) + ]); qed_goalw "lub_sprod2" Sprod3.thy [sfst_def,ssnd_def,spair_def] "[|is_chain(S)|] ==> range(S) <<| \ \ (| lub(range(%i.sfst`(S i))), lub(range(%i.ssnd`(S i))) |)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun_sprod RS ssubst) 1), - (rtac (beta_cfun RS ext RS ssubst) 1), - (rtac cont_Issnd 1), - (rtac (beta_cfun RS ext RS ssubst) 1), - (rtac cont_Isfst 1), - (rtac lub_sprod 1), - (resolve_tac prems 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun_sprod RS ssubst) 1), + (rtac (beta_cfun RS ext RS ssubst) 1), + (rtac cont_Issnd 1), + (rtac (beta_cfun RS ext RS ssubst) 1), + (rtac cont_Isfst 1), + (rtac lub_sprod 1), + (resolve_tac prems 1) + ]); val thelub_sprod2 = (lub_sprod2 RS thelubI); @@ -628,51 +628,51 @@ *) qed_goalw "ssplit1" Sprod3.thy [ssplit_def] - "ssplit`f`UU=UU" + "ssplit`f`UU=UU" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (cont_tacR 1), - (rtac (strictify1 RS ssubst) 1), - (rtac refl 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (cont_tacR 1), + (rtac (strictify1 RS ssubst) 1), + (rtac refl 1) + ]); qed_goalw "ssplit2" Sprod3.thy [ssplit_def] - "[|x~=UU;y~=UU|] ==> ssplit`f`(|x,y|)= f`x`y" + "[|x~=UU;y~=UU|] ==> ssplit`f`(|x,y|)= f`x`y" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (cont_tacR 1), - (rtac (strictify2 RS ssubst) 1), - (rtac defined_spair 1), - (resolve_tac prems 1), - (resolve_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (cont_tacR 1), - (rtac (sfst2 RS ssubst) 1), - (resolve_tac prems 1), - (rtac (ssnd2 RS ssubst) 1), - (resolve_tac prems 1), - (rtac refl 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (cont_tacR 1), + (rtac (strictify2 RS ssubst) 1), + (rtac defined_spair 1), + (resolve_tac prems 1), + (resolve_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (cont_tacR 1), + (rtac (sfst2 RS ssubst) 1), + (resolve_tac prems 1), + (rtac (ssnd2 RS ssubst) 1), + (resolve_tac prems 1), + (rtac refl 1) + ]); qed_goalw "ssplit3" Sprod3.thy [ssplit_def] "ssplit`spair`z=z" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (cont_tacR 1), - (res_inst_tac [("Q","z=UU")] classical2 1), - (hyp_subst_tac 1), - (rtac strictify1 1), - (rtac trans 1), - (rtac strictify2 1), - (atac 1), - (rtac (beta_cfun RS ssubst) 1), - (cont_tacR 1), - (rtac surjective_pairing_Sprod2 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (cont_tacR 1), + (res_inst_tac [("Q","z=UU")] classical2 1), + (hyp_subst_tac 1), + (rtac strictify1 1), + (rtac trans 1), + (rtac strictify2 1), + (atac 1), + (rtac (beta_cfun RS ssubst) 1), + (cont_tacR 1), + (rtac surjective_pairing_Sprod2 1) + ]); (* ------------------------------------------------------------------------ *) @@ -680,9 +680,9 @@ (* ------------------------------------------------------------------------ *) val Sprod_rews = [strict_spair1,strict_spair2,strict_sfst1,strict_sfst2, - strict_ssnd1,strict_ssnd2,sfst2,ssnd2,defined_spair, - ssplit1,ssplit2]; + strict_ssnd1,strict_ssnd2,sfst2,ssnd2,defined_spair, + ssplit1,ssplit2]; Addsimps [strict_spair1,strict_spair2,strict_sfst1,strict_sfst2, - strict_ssnd1,strict_ssnd2,sfst2,ssnd2,defined_spair, - ssplit1,ssplit2]; + strict_ssnd1,strict_ssnd2,sfst2,ssnd2,defined_spair, + ssplit1,ssplit2]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Ssum0.ML --- a/src/HOLCF/Ssum0.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Ssum0.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/ssum0.ML +(* Title: HOLCF/ssum0.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for theory ssum0.thy @@ -14,28 +14,28 @@ 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) - ]); + [ + (rtac CollectI 1), + (rtac disjI1 1), + (rtac exI 1), + (rtac refl 1) + ]); 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) - ]); + [ + (rtac CollectI 1), + (rtac disjI2 1), + (rtac exI 1), + (rtac refl 1) + ]); qed_goal "inj_onto_Abs_Ssum" Ssum0.thy "inj_onto Abs_Ssum Ssum" (fn prems => - [ - (rtac inj_onto_inverseI 1), - (etac Abs_Ssum_inverse 1) - ]); + [ + (rtac inj_onto_inverseI 1), + (etac Abs_Ssum_inverse 1) + ]); (* ------------------------------------------------------------------------ *) (* Strictness of Sinr_Rep, Sinl_Rep and Isinl, Isinr *) @@ -44,19 +44,19 @@ 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) - ]); + [ + (rtac ext 1), + (rtac ext 1), + (rtac ext 1), + (fast_tac HOL_cs 1) + ]); qed_goalw "strict_IsinlIsinr" Ssum0.thy [Isinl_def,Isinr_def] "Isinl(UU) = Isinr(UU)" (fn prems => - [ - (rtac (strict_SinlSinr_Rep RS arg_cong) 1) - ]); + [ + (rtac (strict_SinlSinr_Rep RS arg_cong) 1) + ]); (* ------------------------------------------------------------------------ *) @@ -64,35 +64,35 @@ (* ------------------------------------------------------------------------ *) qed_goalw "noteq_SinlSinr_Rep" Ssum0.thy [Sinl_Rep_def,Sinr_Rep_def] - "(Sinl_Rep(a) = Sinr_Rep(b)) ==> a=UU & b=UU" + "(Sinl_Rep(a) = Sinr_Rep(b)) ==> a=UU & b=UU" (fn prems => - [ - (rtac conjI 1), - (res_inst_tac [("Q","a=UU")] classical2 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), - (res_inst_tac [("Q","b=UU")] classical2 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) - ]); + [ + (rtac conjI 1), + (res_inst_tac [("Q","a=UU")] classical2 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), + (res_inst_tac [("Q","b=UU")] classical2 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) + ]); qed_goalw "noteq_IsinlIsinr" Ssum0.thy [Isinl_def,Isinr_def] - "Isinl(a)=Isinr(b) ==> a=UU & b=UU" + "Isinl(a)=Isinr(b) ==> a=UU & b=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac noteq_SinlSinr_Rep 1), - (etac (inj_onto_Abs_Ssum RS inj_ontoD) 1), - (rtac SsumIl 1), - (rtac SsumIr 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac noteq_SinlSinr_Rep 1), + (etac (inj_onto_Abs_Ssum RS inj_ontoD) 1), + (rtac SsumIl 1), + (rtac SsumIr 1) + ]); @@ -103,122 +103,122 @@ qed_goalw "inject_Sinl_Rep1" Ssum0.thy [Sinl_Rep_def] "(Sinl_Rep(a) = Sinl_Rep(UU)) ==> a=UU" (fn prems => - [ - (res_inst_tac [("Q","a=UU")] classical2 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) - ]); + [ + (res_inst_tac [("Q","a=UU")] classical2 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) + ]); qed_goalw "inject_Sinr_Rep1" Ssum0.thy [Sinr_Rep_def] "(Sinr_Rep(b) = Sinr_Rep(UU)) ==> b=UU" (fn prems => - [ - (res_inst_tac [("Q","b=UU")] classical2 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) - ]); + [ + (res_inst_tac [("Q","b=UU")] classical2 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) + ]); 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) - ]); + [ + (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) + ]); 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) - ]); + [ + (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) + ]); qed_goal "inject_Sinl_Rep" Ssum0.thy - "Sinl_Rep(a1)=Sinl_Rep(a2) ==> a1=a2" + "Sinl_Rep(a1)=Sinl_Rep(a2) ==> a1=a2" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("Q","a1=UU")] classical2 1), - (hyp_subst_tac 1), - (rtac (inject_Sinl_Rep1 RS sym) 1), - (etac sym 1), - (res_inst_tac [("Q","a2=UU")] classical2 1), - (hyp_subst_tac 1), - (etac inject_Sinl_Rep1 1), - (etac inject_Sinl_Rep2 1), - (atac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("Q","a1=UU")] classical2 1), + (hyp_subst_tac 1), + (rtac (inject_Sinl_Rep1 RS sym) 1), + (etac sym 1), + (res_inst_tac [("Q","a2=UU")] classical2 1), + (hyp_subst_tac 1), + (etac inject_Sinl_Rep1 1), + (etac inject_Sinl_Rep2 1), + (atac 1), + (atac 1) + ]); qed_goal "inject_Sinr_Rep" Ssum0.thy - "Sinr_Rep(b1)=Sinr_Rep(b2) ==> b1=b2" + "Sinr_Rep(b1)=Sinr_Rep(b2) ==> b1=b2" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("Q","b1=UU")] classical2 1), - (hyp_subst_tac 1), - (rtac (inject_Sinr_Rep1 RS sym) 1), - (etac sym 1), - (res_inst_tac [("Q","b2=UU")] classical2 1), - (hyp_subst_tac 1), - (etac inject_Sinr_Rep1 1), - (etac inject_Sinr_Rep2 1), - (atac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("Q","b1=UU")] classical2 1), + (hyp_subst_tac 1), + (rtac (inject_Sinr_Rep1 RS sym) 1), + (etac sym 1), + (res_inst_tac [("Q","b2=UU")] classical2 1), + (hyp_subst_tac 1), + (etac inject_Sinr_Rep1 1), + (etac inject_Sinr_Rep2 1), + (atac 1), + (atac 1) + ]); 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_onto_Abs_Ssum RS inj_ontoD) 1), - (rtac SsumIl 1), - (rtac SsumIl 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac inject_Sinl_Rep 1), + (etac (inj_onto_Abs_Ssum RS inj_ontoD) 1), + (rtac SsumIl 1), + (rtac SsumIl 1) + ]); 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_onto_Abs_Ssum RS inj_ontoD) 1), - (rtac SsumIr 1), - (rtac SsumIr 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac inject_Sinr_Rep 1), + (etac (inj_onto_Abs_Ssum RS inj_ontoD) 1), + (rtac SsumIr 1), + (rtac SsumIr 1) + ]); qed_goal "inject_Isinl_rev" Ssum0.thy "a1~=a2 ==> Isinl(a1) ~= Isinl(a2)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac contrapos 1), - (etac inject_Isinl 2), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac contrapos 1), + (etac inject_Isinl 2), + (atac 1) + ]); qed_goal "inject_Isinr_rev" Ssum0.thy "b1~=b2 ==> Isinr(b1) ~= Isinr(b2)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac contrapos 1), - (etac inject_Isinr 2), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac contrapos 1), + (etac inject_Isinr 2), + (atac 1) + ]); (* ------------------------------------------------------------------------ *) (* Exhaustion of the strict sum ++ *) @@ -226,78 +226,78 @@ (* ------------------------------------------------------------------------ *) 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)" + "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), - (res_inst_tac [("Q","z= Abs_Ssum(Sinl_Rep(UU))")] classical2 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), - (res_inst_tac [("Q","z= Abs_Ssum(Sinl_Rep(UU))")] classical2 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) - ]); + [ + (rtac (rewrite_rule [Ssum_def] Rep_Ssum RS CollectE) 1), + (etac disjE 1), + (etac exE 1), + (res_inst_tac [("Q","z= Abs_Ssum(Sinl_Rep(UU))")] classical2 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), + (res_inst_tac [("Q","z= Abs_Ssum(Sinl_Rep(UU))")] classical2 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) + ]); (* ------------------------------------------------------------------------ *) (* elimination rules for the strict sum ++ *) (* ------------------------------------------------------------------------ *) qed_goal "IssumE" Ssum0.thy - "[|p=Isinl(UU) ==> Q ;\ -\ !!x.[|p=Isinl(x); x~=UU |] ==> Q;\ -\ !!y.[|p=Isinr(y); y~=UU |] ==> Q|] ==> Q" + "[|p=Isinl(UU) ==> Q ;\ +\ !!x.[|p=Isinl(x); x~=UU |] ==> Q;\ +\ !!y.[|p=Isinr(y); y~=UU |] ==> Q|] ==> Q" (fn prems => - [ - (rtac (Exh_Ssum RS disjE) 1), - (etac disjE 2), - (eresolve_tac prems 1), - (etac exE 1), - (etac conjE 1), - (eresolve_tac prems 1), - (atac 1), - (etac exE 1), - (etac conjE 1), - (eresolve_tac prems 1), - (atac 1) - ]); + [ + (rtac (Exh_Ssum RS disjE) 1), + (etac disjE 2), + (eresolve_tac prems 1), + (etac exE 1), + (etac conjE 1), + (eresolve_tac prems 1), + (atac 1), + (etac exE 1), + (etac conjE 1), + (eresolve_tac prems 1), + (atac 1) + ]); qed_goal "IssumE2" Ssum0.thy "[| !!x. [| p = Isinl(x) |] ==> Q; !!y. [| p = Isinr(y) |] ==> Q |] ==>Q" (fn prems => - [ - (rtac IssumE 1), - (eresolve_tac prems 1), - (eresolve_tac prems 1), - (eresolve_tac prems 1) - ]); + [ + (rtac IssumE 1), + (eresolve_tac prems 1), + (eresolve_tac prems 1), + (eresolve_tac prems 1) + ]); @@ -307,82 +307,82 @@ (* ------------------------------------------------------------------------ *) qed_goalw "Iwhen1" Ssum0.thy [Iwhen_def] - "Iwhen f g (Isinl UU) = UU" + "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) - ]); + [ + (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) + ]); qed_goalw "Iwhen2" Ssum0.thy [Iwhen_def] - "x~=UU ==> Iwhen f g (Isinl x) = f`x" + "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) - ]); + [ + (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) + ]); qed_goalw "Iwhen3" Ssum0.thy [Iwhen_def] - "y~=UU ==> Iwhen f g (Isinr y) = g`y" + "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) - ]); + [ + (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) + ]); (* ------------------------------------------------------------------------ *) (* instantiate the simplifier *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Ssum1.ML --- a/src/HOLCF/Ssum1.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Ssum1.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/ssum1.ML +(* Title: HOLCF/ssum1.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for theory ssum1.thy @@ -11,200 +11,200 @@ local fun eq_left s1 s2 = - ( - (res_inst_tac [("s",s1),("t",s2)] (inject_Isinl RS subst) 1) - THEN (rtac trans 1) - THEN (atac 2) - THEN (etac sym 1)); + ( + (res_inst_tac [("s",s1),("t",s2)] (inject_Isinl RS subst) 1) + THEN (rtac trans 1) + THEN (atac 2) + THEN (etac sym 1)); fun eq_right s1 s2 = - ( - (res_inst_tac [("s",s1),("t",s2)] (inject_Isinr RS subst) 1) - THEN (rtac trans 1) - THEN (atac 2) - THEN (etac sym 1)); + ( + (res_inst_tac [("s",s1),("t",s2)] (inject_Isinr RS subst) 1) + THEN (rtac trans 1) + THEN (atac 2) + THEN (etac sym 1)); fun UU_left s1 = - ( - (res_inst_tac [("t",s1)](noteq_IsinlIsinr RS conjunct1 RS ssubst)1) - THEN (rtac trans 1) - THEN (atac 2) - THEN (etac sym 1)); + ( + (res_inst_tac [("t",s1)](noteq_IsinlIsinr RS conjunct1 RS ssubst)1) + THEN (rtac trans 1) + THEN (atac 2) + THEN (etac sym 1)); fun UU_right s1 = - ( - (res_inst_tac [("t",s1)](noteq_IsinlIsinr RS conjunct2 RS ssubst)1) - THEN (rtac trans 1) - THEN (atac 2) - THEN (etac sym 1)) + ( + (res_inst_tac [("t",s1)](noteq_IsinlIsinr RS conjunct2 RS ssubst)1) + THEN (rtac trans 1) + THEN (atac 2) + THEN (etac sym 1)) in val less_ssum1a = prove_goalw Ssum1.thy [less_ssum_def] "[|s1=Isinl(x::'a); s2=Isinl(y::'a)|] ==> less_ssum 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) - ]); + [ + (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 less_ssum1b = prove_goalw Ssum1.thy [less_ssum_def] "[|s1=Isinr(x::'b); s2=Isinr(y::'b)|] ==> less_ssum 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) - ]); + [ + (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 less_ssum1c = prove_goalw Ssum1.thy [less_ssum_def] "[|s1=Isinl(x::'a); s2=Isinr(y::'b)|] ==> less_ssum 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) - ]); + [ + (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 Ssum1.thy [less_ssum_def] "[|s1=Isinr(x); s2=Isinl(y)|] ==> less_ssum 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) - ]) + [ + (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; @@ -213,40 +213,40 @@ (* ------------------------------------------------------------------------ *) qed_goal "less_ssum2a" Ssum1.thy - "less_ssum (Isinl x) (Isinl y) = (x << y)" + "less_ssum (Isinl x) (Isinl y) = (x << y)" (fn prems => - [ - (rtac less_ssum1a 1), - (rtac refl 1), - (rtac refl 1) - ]); + [ + (rtac less_ssum1a 1), + (rtac refl 1), + (rtac refl 1) + ]); qed_goal "less_ssum2b" Ssum1.thy - "less_ssum (Isinr x) (Isinr y) = (x << y)" + "less_ssum (Isinr x) (Isinr y) = (x << y)" (fn prems => - [ - (rtac less_ssum1b 1), - (rtac refl 1), - (rtac refl 1) - ]); + [ + (rtac less_ssum1b 1), + (rtac refl 1), + (rtac refl 1) + ]); qed_goal "less_ssum2c" Ssum1.thy - "less_ssum (Isinl x) (Isinr y) = (x = UU)" + "less_ssum (Isinl x) (Isinr y) = (x = UU)" (fn prems => - [ - (rtac less_ssum1c 1), - (rtac refl 1), - (rtac refl 1) - ]); + [ + (rtac less_ssum1c 1), + (rtac refl 1), + (rtac refl 1) + ]); qed_goal "less_ssum2d" Ssum1.thy - "less_ssum (Isinr x) (Isinl y) = (x = UU)" + "less_ssum (Isinr x) (Isinl y) = (x = UU)" (fn prems => - [ - (rtac less_ssum1d 1), - (rtac refl 1), - (rtac refl 1) - ]); + [ + (rtac less_ssum1d 1), + (rtac refl 1), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) @@ -255,99 +255,99 @@ qed_goal "refl_less_ssum" Ssum1.thy "less_ssum p p" (fn prems => - [ - (res_inst_tac [("p","p")] IssumE2 1), - (hyp_subst_tac 1), - (rtac (less_ssum2a RS iffD2) 1), - (rtac refl_less 1), - (hyp_subst_tac 1), - (rtac (less_ssum2b RS iffD2) 1), - (rtac refl_less 1) - ]); + [ + (res_inst_tac [("p","p")] IssumE2 1), + (hyp_subst_tac 1), + (rtac (less_ssum2a RS iffD2) 1), + (rtac refl_less 1), + (hyp_subst_tac 1), + (rtac (less_ssum2b RS iffD2) 1), + (rtac refl_less 1) + ]); qed_goal "antisym_less_ssum" Ssum1.thy "[|less_ssum p1 p2; less_ssum p2 p1|] ==> p1=p2" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p1")] IssumE2 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p2")] IssumE2 1), - (hyp_subst_tac 1), - (res_inst_tac [("f","Isinl")] arg_cong 1), - (rtac antisym_less 1), - (etac (less_ssum2a RS iffD1) 1), - (etac (less_ssum2a RS iffD1) 1), - (hyp_subst_tac 1), - (etac (less_ssum2d RS iffD1 RS ssubst) 1), - (etac (less_ssum2c RS iffD1 RS ssubst) 1), - (rtac strict_IsinlIsinr 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p2")] IssumE2 1), - (hyp_subst_tac 1), - (etac (less_ssum2c RS iffD1 RS ssubst) 1), - (etac (less_ssum2d RS iffD1 RS ssubst) 1), - (rtac (strict_IsinlIsinr RS sym) 1), - (hyp_subst_tac 1), - (res_inst_tac [("f","Isinr")] arg_cong 1), - (rtac antisym_less 1), - (etac (less_ssum2b RS iffD1) 1), - (etac (less_ssum2b RS iffD1) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","p1")] IssumE2 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p2")] IssumE2 1), + (hyp_subst_tac 1), + (res_inst_tac [("f","Isinl")] arg_cong 1), + (rtac antisym_less 1), + (etac (less_ssum2a RS iffD1) 1), + (etac (less_ssum2a RS iffD1) 1), + (hyp_subst_tac 1), + (etac (less_ssum2d RS iffD1 RS ssubst) 1), + (etac (less_ssum2c RS iffD1 RS ssubst) 1), + (rtac strict_IsinlIsinr 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p2")] IssumE2 1), + (hyp_subst_tac 1), + (etac (less_ssum2c RS iffD1 RS ssubst) 1), + (etac (less_ssum2d RS iffD1 RS ssubst) 1), + (rtac (strict_IsinlIsinr RS sym) 1), + (hyp_subst_tac 1), + (res_inst_tac [("f","Isinr")] arg_cong 1), + (rtac antisym_less 1), + (etac (less_ssum2b RS iffD1) 1), + (etac (less_ssum2b RS iffD1) 1) + ]); qed_goal "trans_less_ssum" Ssum1.thy "[|less_ssum p1 p2; less_ssum p2 p3|] ==> less_ssum p1 p3" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","p1")] IssumE2 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p3")] IssumE2 1), - (hyp_subst_tac 1), - (rtac (less_ssum2a RS iffD2) 1), - (res_inst_tac [("p","p2")] IssumE2 1), - (hyp_subst_tac 1), - (rtac trans_less 1), - (etac (less_ssum2a RS iffD1) 1), - (etac (less_ssum2a RS iffD1) 1), - (hyp_subst_tac 1), - (etac (less_ssum2c RS iffD1 RS ssubst) 1), - (rtac minimal 1), - (hyp_subst_tac 1), - (rtac (less_ssum2c RS iffD2) 1), - (res_inst_tac [("p","p2")] IssumE2 1), - (hyp_subst_tac 1), - (rtac UU_I 1), - (rtac trans_less 1), - (etac (less_ssum2a RS iffD1) 1), - (rtac (antisym_less_inverse RS conjunct1) 1), - (etac (less_ssum2c RS iffD1) 1), - (hyp_subst_tac 1), - (etac (less_ssum2c RS iffD1) 1), - (hyp_subst_tac 1), - (res_inst_tac [("p","p3")] IssumE2 1), - (hyp_subst_tac 1), - (rtac (less_ssum2d RS iffD2) 1), - (res_inst_tac [("p","p2")] IssumE2 1), - (hyp_subst_tac 1), - (etac (less_ssum2d RS iffD1) 1), - (hyp_subst_tac 1), - (rtac UU_I 1), - (rtac trans_less 1), - (etac (less_ssum2b RS iffD1) 1), - (rtac (antisym_less_inverse RS conjunct1) 1), - (etac (less_ssum2d RS iffD1) 1), - (hyp_subst_tac 1), - (rtac (less_ssum2b RS iffD2) 1), - (res_inst_tac [("p","p2")] IssumE2 1), - (hyp_subst_tac 1), - (etac (less_ssum2d RS iffD1 RS ssubst) 1), - (rtac minimal 1), - (hyp_subst_tac 1), - (rtac trans_less 1), - (etac (less_ssum2b RS iffD1) 1), - (etac (less_ssum2b RS iffD1) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","p1")] IssumE2 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p3")] IssumE2 1), + (hyp_subst_tac 1), + (rtac (less_ssum2a RS iffD2) 1), + (res_inst_tac [("p","p2")] IssumE2 1), + (hyp_subst_tac 1), + (rtac trans_less 1), + (etac (less_ssum2a RS iffD1) 1), + (etac (less_ssum2a RS iffD1) 1), + (hyp_subst_tac 1), + (etac (less_ssum2c RS iffD1 RS ssubst) 1), + (rtac minimal 1), + (hyp_subst_tac 1), + (rtac (less_ssum2c RS iffD2) 1), + (res_inst_tac [("p","p2")] IssumE2 1), + (hyp_subst_tac 1), + (rtac UU_I 1), + (rtac trans_less 1), + (etac (less_ssum2a RS iffD1) 1), + (rtac (antisym_less_inverse RS conjunct1) 1), + (etac (less_ssum2c RS iffD1) 1), + (hyp_subst_tac 1), + (etac (less_ssum2c RS iffD1) 1), + (hyp_subst_tac 1), + (res_inst_tac [("p","p3")] IssumE2 1), + (hyp_subst_tac 1), + (rtac (less_ssum2d RS iffD2) 1), + (res_inst_tac [("p","p2")] IssumE2 1), + (hyp_subst_tac 1), + (etac (less_ssum2d RS iffD1) 1), + (hyp_subst_tac 1), + (rtac UU_I 1), + (rtac trans_less 1), + (etac (less_ssum2b RS iffD1) 1), + (rtac (antisym_less_inverse RS conjunct1) 1), + (etac (less_ssum2d RS iffD1) 1), + (hyp_subst_tac 1), + (rtac (less_ssum2b RS iffD2) 1), + (res_inst_tac [("p","p2")] IssumE2 1), + (hyp_subst_tac 1), + (etac (less_ssum2d RS iffD1 RS ssubst) 1), + (rtac minimal 1), + (hyp_subst_tac 1), + (rtac trans_less 1), + (etac (less_ssum2b RS iffD1) 1), + (etac (less_ssum2b RS iffD1) 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Ssum2.ML --- a/src/HOLCF/Ssum2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Ssum2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/ssum2.ML +(* Title: HOLCF/ssum2.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for ssum2.thy @@ -13,36 +13,36 @@ (* ------------------------------------------------------------------------ *) qed_goal "less_ssum3a" Ssum2.thy - "(Isinl(x) << Isinl(y)) = (x << y)" + "(Isinl(x) << Isinl(y)) = (x << y)" (fn prems => - [ - (rtac (inst_ssum_po RS ssubst) 1), - (rtac less_ssum2a 1) - ]); + [ + (rtac (inst_ssum_po RS ssubst) 1), + (rtac less_ssum2a 1) + ]); qed_goal "less_ssum3b" Ssum2.thy - "(Isinr(x) << Isinr(y)) = (x << y)" + "(Isinr(x) << Isinr(y)) = (x << y)" (fn prems => - [ - (rtac (inst_ssum_po RS ssubst) 1), - (rtac less_ssum2b 1) - ]); + [ + (rtac (inst_ssum_po RS ssubst) 1), + (rtac less_ssum2b 1) + ]); qed_goal "less_ssum3c" Ssum2.thy - "(Isinl(x) << Isinr(y)) = (x = UU)" + "(Isinl(x) << Isinr(y)) = (x = UU)" (fn prems => - [ - (rtac (inst_ssum_po RS ssubst) 1), - (rtac less_ssum2c 1) - ]); + [ + (rtac (inst_ssum_po RS ssubst) 1), + (rtac less_ssum2c 1) + ]); qed_goal "less_ssum3d" Ssum2.thy - "(Isinr(x) << Isinl(y)) = (x = UU)" + "(Isinr(x) << Isinl(y)) = (x = UU)" (fn prems => - [ - (rtac (inst_ssum_po RS ssubst) 1), - (rtac less_ssum2d 1) - ]); + [ + (rtac (inst_ssum_po RS ssubst) 1), + (rtac less_ssum2d 1) + ]); (* ------------------------------------------------------------------------ *) @@ -51,16 +51,16 @@ qed_goal "minimal_ssum" Ssum2.thy "Isinl(UU) << s" (fn prems => - [ - (res_inst_tac [("p","s")] IssumE2 1), - (hyp_subst_tac 1), - (rtac (less_ssum3a RS iffD2) 1), - (rtac minimal 1), - (hyp_subst_tac 1), - (rtac (strict_IsinlIsinr RS ssubst) 1), - (rtac (less_ssum3b RS iffD2) 1), - (rtac minimal 1) - ]); + [ + (res_inst_tac [("p","s")] IssumE2 1), + (hyp_subst_tac 1), + (rtac (less_ssum3a RS iffD2) 1), + (rtac minimal 1), + (hyp_subst_tac 1), + (rtac (strict_IsinlIsinr RS ssubst) 1), + (rtac (less_ssum3b RS iffD2) 1), + (rtac minimal 1) + ]); (* ------------------------------------------------------------------------ *) @@ -69,17 +69,17 @@ qed_goalw "monofun_Isinl" Ssum2.thy [monofun] "monofun(Isinl)" (fn prems => - [ - (strip_tac 1), - (etac (less_ssum3a RS iffD2) 1) - ]); + [ + (strip_tac 1), + (etac (less_ssum3a RS iffD2) 1) + ]); qed_goalw "monofun_Isinr" Ssum2.thy [monofun] "monofun(Isinr)" (fn prems => - [ - (strip_tac 1), - (etac (less_ssum3b RS iffD2) 1) - ]); + [ + (strip_tac 1), + (etac (less_ssum3b RS iffD2) 1) + ]); (* ------------------------------------------------------------------------ *) @@ -89,73 +89,73 @@ qed_goalw "monofun_Iwhen1" Ssum2.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), + [ + (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) - ]); + ]); qed_goalw "monofun_Iwhen2" Ssum2.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), + [ + (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) - ]); + (etac monofun_cfun_fun 1) + ]); qed_goalw "monofun_Iwhen3" Ssum2.thy [monofun] "monofun(Iwhen(f)(g))" (fn prems => - [ - (strip_tac 1), - (res_inst_tac [("p","x")] IssumE 1), - (hyp_subst_tac 1), + [ + (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), + (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), + (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), + (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), + (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), + (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), + (hyp_subst_tac 1), (asm_simp_tac Ssum0_ss 1), - (rtac monofun_cfun_arg 1), - (etac (less_ssum3b RS iffD1) 1) - ]); + (rtac monofun_cfun_arg 1), + (etac (less_ssum3b RS iffD1) 1) + ]); @@ -168,72 +168,72 @@ qed_goal "ssum_lemma1" Ssum2.thy "[|~(!i.? x.Y(i::nat)=Isinl(x))|] ==> (? i.! x.Y(i)~=Isinl(x))" (fn prems => - [ - (cut_facts_tac prems 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (fast_tac HOL_cs 1) + ]); qed_goal "ssum_lemma2" Ssum2.thy "[|(? i.!x.(Y::nat => 'a++'b)(i::nat)~=Isinl(x::'a))|] ==>\ \ (? i y. (Y::nat => 'a++'b)(i::nat)=Isinr(y::'b) & y~=UU)" (fn prems => - [ - (cut_facts_tac prems 1), - (etac exE 1), - (res_inst_tac [("p","Y(i)")] IssumE 1), - (dtac spec 1), - (contr_tac 1), - (dtac spec 1), - (contr_tac 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (etac exE 1), + (res_inst_tac [("p","Y(i)")] IssumE 1), + (dtac spec 1), + (contr_tac 1), + (dtac spec 1), + (contr_tac 1), + (fast_tac HOL_cs 1) + ]); qed_goal "ssum_lemma3" Ssum2.thy "[|is_chain(Y);(? i x. Y(i)=Isinr(x::'b) & (x::'b)~=UU)|] ==>\ \ (!i.? y.Y(i)=Isinr(y))" (fn prems => - [ - (cut_facts_tac prems 1), - (etac exE 1), - (etac exE 1), - (rtac allI 1), - (res_inst_tac [("p","Y(ia)")] IssumE 1), - (rtac exI 1), - (rtac trans 1), - (rtac strict_IsinlIsinr 2), - (atac 1), - (etac exI 2), - (etac conjE 1), - (res_inst_tac [("m","i"),("n","ia")] nat_less_cases 1), - (hyp_subst_tac 2), - (etac exI 2), - (eres_inst_tac [("P","x=UU")] notE 1), - (rtac (less_ssum3d RS iffD1) 1), - (eres_inst_tac [("s","Y(i)"),("t","Isinr(x)::'a++'b")] subst 1), - (eres_inst_tac [("s","Y(ia)"),("t","Isinl(xa)::'a++'b")] subst 1), - (etac (chain_mono RS mp) 1), - (atac 1), - (eres_inst_tac [("P","xa=UU")] notE 1), - (rtac (less_ssum3c RS iffD1) 1), - (eres_inst_tac [("s","Y(i)"),("t","Isinr(x)::'a++'b")] subst 1), - (eres_inst_tac [("s","Y(ia)"),("t","Isinl(xa)::'a++'b")] subst 1), - (etac (chain_mono RS mp) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac exE 1), + (etac exE 1), + (rtac allI 1), + (res_inst_tac [("p","Y(ia)")] IssumE 1), + (rtac exI 1), + (rtac trans 1), + (rtac strict_IsinlIsinr 2), + (atac 1), + (etac exI 2), + (etac conjE 1), + (res_inst_tac [("m","i"),("n","ia")] nat_less_cases 1), + (hyp_subst_tac 2), + (etac exI 2), + (eres_inst_tac [("P","x=UU")] notE 1), + (rtac (less_ssum3d RS iffD1) 1), + (eres_inst_tac [("s","Y(i)"),("t","Isinr(x)::'a++'b")] subst 1), + (eres_inst_tac [("s","Y(ia)"),("t","Isinl(xa)::'a++'b")] subst 1), + (etac (chain_mono RS mp) 1), + (atac 1), + (eres_inst_tac [("P","xa=UU")] notE 1), + (rtac (less_ssum3c RS iffD1) 1), + (eres_inst_tac [("s","Y(i)"),("t","Isinr(x)::'a++'b")] subst 1), + (eres_inst_tac [("s","Y(ia)"),("t","Isinl(xa)::'a++'b")] subst 1), + (etac (chain_mono RS mp) 1), + (atac 1) + ]); qed_goal "ssum_lemma4" Ssum2.thy "is_chain(Y) ==> (!i.? x.Y(i)=Isinl(x))|(!i.? y.Y(i)=Isinr(y))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac classical2 1), - (etac disjI1 1), - (rtac disjI2 1), - (etac ssum_lemma3 1), - (rtac ssum_lemma2 1), - (etac ssum_lemma1 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac classical2 1), + (etac disjI1 1), + (rtac disjI2 1), + (etac ssum_lemma3 1), + (rtac ssum_lemma2 1), + (etac ssum_lemma1 1) + ]); (* ------------------------------------------------------------------------ *) @@ -243,13 +243,13 @@ qed_goal "ssum_lemma5" Ssum2.thy "z=Isinl(x)==> Isinl((Iwhen (LAM x.x) (LAM y.UU))(z)) = z" (fn prems => - [ - (cut_facts_tac prems 1), - (hyp_subst_tac 1), - (res_inst_tac [("Q","x=UU")] classical2 1), + [ + (cut_facts_tac prems 1), + (hyp_subst_tac 1), + (res_inst_tac [("Q","x=UU")] classical2 1), (asm_simp_tac Ssum0_ss 1), (asm_simp_tac Ssum0_ss 1) - ]); + ]); (* ------------------------------------------------------------------------ *) (* restricted surjectivity of Isinr *) @@ -258,13 +258,13 @@ qed_goal "ssum_lemma6" Ssum2.thy "z=Isinr(x)==> Isinr((Iwhen (LAM y.UU) (LAM x.x))(z)) = z" (fn prems => - [ - (cut_facts_tac prems 1), - (hyp_subst_tac 1), - (res_inst_tac [("Q","x=UU")] classical2 1), + [ + (cut_facts_tac prems 1), + (hyp_subst_tac 1), + (res_inst_tac [("Q","x=UU")] classical2 1), (asm_simp_tac Ssum0_ss 1), (asm_simp_tac Ssum0_ss 1) - ]); + ]); (* ------------------------------------------------------------------------ *) (* technical lemmas *) @@ -273,36 +273,36 @@ qed_goal "ssum_lemma7" Ssum2.thy "[|Isinl(x) << z; x~=UU|] ==> ? y.z=Isinl(y) & y~=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","z")] IssumE 1), - (hyp_subst_tac 1), - (etac notE 1), - (rtac antisym_less 1), - (etac (less_ssum3a RS iffD1) 1), - (rtac minimal 1), - (fast_tac HOL_cs 1), - (hyp_subst_tac 1), - (rtac notE 1), - (etac (less_ssum3c RS iffD1) 2), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","z")] IssumE 1), + (hyp_subst_tac 1), + (etac notE 1), + (rtac antisym_less 1), + (etac (less_ssum3a RS iffD1) 1), + (rtac minimal 1), + (fast_tac HOL_cs 1), + (hyp_subst_tac 1), + (rtac notE 1), + (etac (less_ssum3c RS iffD1) 2), + (atac 1) + ]); qed_goal "ssum_lemma8" Ssum2.thy "[|Isinr(x) << z; x~=UU|] ==> ? y.z=Isinr(y) & y~=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("p","z")] IssumE 1), - (hyp_subst_tac 1), - (etac notE 1), - (etac (less_ssum3d RS iffD1) 1), - (hyp_subst_tac 1), - (rtac notE 1), - (etac (less_ssum3d RS iffD1) 2), - (atac 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("p","z")] IssumE 1), + (hyp_subst_tac 1), + (etac notE 1), + (etac (less_ssum3d RS iffD1) 1), + (hyp_subst_tac 1), + (rtac notE 1), + (etac (less_ssum3d RS iffD1) 2), + (atac 1), + (fast_tac HOL_cs 1) + ]); (* ------------------------------------------------------------------------ *) (* the type 'a ++ 'b is a cpo in three steps *) @@ -313,40 +313,40 @@ \ range(Y) <<|\ \ Isinl(lub(range(%i.(Iwhen (LAM x.x) (LAM y.UU))(Y i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_lubI 1), - (rtac conjI 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (etac allE 1), - (etac exE 1), - (res_inst_tac [("t","Y(i)")] (ssum_lemma5 RS subst) 1), - (atac 1), - (rtac (monofun_Isinl RS monofunE RS spec RS spec RS mp) 1), - (rtac is_ub_thelub 1), - (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), - (strip_tac 1), - (res_inst_tac [("p","u")] IssumE2 1), - (res_inst_tac [("t","u")] (ssum_lemma5 RS subst) 1), - (atac 1), - (rtac (monofun_Isinl RS monofunE RS spec RS spec RS mp) 1), - (rtac is_lub_thelub 1), - (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), - (etac (monofun_Iwhen3 RS ub2ub_monofun) 1), - (hyp_subst_tac 1), - (rtac (less_ssum3c RS iffD2) 1), - (rtac chain_UU_I_inverse 1), - (rtac allI 1), - (res_inst_tac [("p","Y(i)")] IssumE 1), - (asm_simp_tac Ssum0_ss 1), - (asm_simp_tac Ssum0_ss 2), - (etac notE 1), - (rtac (less_ssum3c RS iffD1) 1), - (res_inst_tac [("t","Isinl(x)")] subst 1), - (atac 1), - (etac (ub_rangeE RS spec) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_lubI 1), + (rtac conjI 1), + (rtac ub_rangeI 1), + (rtac allI 1), + (etac allE 1), + (etac exE 1), + (res_inst_tac [("t","Y(i)")] (ssum_lemma5 RS subst) 1), + (atac 1), + (rtac (monofun_Isinl RS monofunE RS spec RS spec RS mp) 1), + (rtac is_ub_thelub 1), + (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), + (strip_tac 1), + (res_inst_tac [("p","u")] IssumE2 1), + (res_inst_tac [("t","u")] (ssum_lemma5 RS subst) 1), + (atac 1), + (rtac (monofun_Isinl RS monofunE RS spec RS spec RS mp) 1), + (rtac is_lub_thelub 1), + (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), + (etac (monofun_Iwhen3 RS ub2ub_monofun) 1), + (hyp_subst_tac 1), + (rtac (less_ssum3c RS iffD2) 1), + (rtac chain_UU_I_inverse 1), + (rtac allI 1), + (res_inst_tac [("p","Y(i)")] IssumE 1), + (asm_simp_tac Ssum0_ss 1), + (asm_simp_tac Ssum0_ss 2), + (etac notE 1), + (rtac (less_ssum3c RS iffD1) 1), + (res_inst_tac [("t","Isinl(x)")] subst 1), + (atac 1), + (etac (ub_rangeE RS spec) 1) + ]); qed_goal "lub_ssum1b" Ssum2.thy @@ -354,40 +354,40 @@ \ range(Y) <<|\ \ Isinr(lub(range(%i.(Iwhen (LAM y.UU) (LAM x.x))(Y i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac is_lubI 1), - (rtac conjI 1), - (rtac ub_rangeI 1), - (rtac allI 1), - (etac allE 1), - (etac exE 1), - (res_inst_tac [("t","Y(i)")] (ssum_lemma6 RS subst) 1), - (atac 1), - (rtac (monofun_Isinr RS monofunE RS spec RS spec RS mp) 1), - (rtac is_ub_thelub 1), - (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), - (strip_tac 1), - (res_inst_tac [("p","u")] IssumE2 1), - (hyp_subst_tac 1), - (rtac (less_ssum3d RS iffD2) 1), - (rtac chain_UU_I_inverse 1), - (rtac allI 1), - (res_inst_tac [("p","Y(i)")] IssumE 1), - (asm_simp_tac Ssum0_ss 1), - (asm_simp_tac Ssum0_ss 1), - (etac notE 1), - (rtac (less_ssum3d RS iffD1) 1), - (res_inst_tac [("t","Isinr(y)")] subst 1), - (atac 1), - (etac (ub_rangeE RS spec) 1), - (res_inst_tac [("t","u")] (ssum_lemma6 RS subst) 1), - (atac 1), - (rtac (monofun_Isinr RS monofunE RS spec RS spec RS mp) 1), - (rtac is_lub_thelub 1), - (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), - (etac (monofun_Iwhen3 RS ub2ub_monofun) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac is_lubI 1), + (rtac conjI 1), + (rtac ub_rangeI 1), + (rtac allI 1), + (etac allE 1), + (etac exE 1), + (res_inst_tac [("t","Y(i)")] (ssum_lemma6 RS subst) 1), + (atac 1), + (rtac (monofun_Isinr RS monofunE RS spec RS spec RS mp) 1), + (rtac is_ub_thelub 1), + (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), + (strip_tac 1), + (res_inst_tac [("p","u")] IssumE2 1), + (hyp_subst_tac 1), + (rtac (less_ssum3d RS iffD2) 1), + (rtac chain_UU_I_inverse 1), + (rtac allI 1), + (res_inst_tac [("p","Y(i)")] IssumE 1), + (asm_simp_tac Ssum0_ss 1), + (asm_simp_tac Ssum0_ss 1), + (etac notE 1), + (rtac (less_ssum3d RS iffD1) 1), + (res_inst_tac [("t","Isinr(y)")] subst 1), + (atac 1), + (etac (ub_rangeE RS spec) 1), + (res_inst_tac [("t","u")] (ssum_lemma6 RS subst) 1), + (atac 1), + (rtac (monofun_Isinr RS monofunE RS spec RS spec RS mp) 1), + (rtac is_lub_thelub 1), + (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), + (etac (monofun_Iwhen3 RS ub2ub_monofun) 1) + ]); val thelub_ssum1a = lub_ssum1a RS thelubI; @@ -405,17 +405,17 @@ *) qed_goal "cpo_ssum" Ssum2.thy - "is_chain(Y::nat=>'a ++'b) ==> ? x.range(Y) <<|x" + "is_chain(Y::nat=>'a ++'b) ==> ? x.range(Y) <<|x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (ssum_lemma4 RS disjE) 1), - (atac 1), - (rtac exI 1), - (etac lub_ssum1a 1), - (atac 1), - (rtac exI 1), - (etac lub_ssum1b 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (ssum_lemma4 RS disjE) 1), + (atac 1), + (rtac exI 1), + (etac lub_ssum1a 1), + (atac 1), + (rtac exI 1), + (etac lub_ssum1b 1), + (atac 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Ssum3.ML --- a/src/HOLCF/Ssum3.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Ssum3.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/ssum3.ML +(* Title: HOLCF/ssum3.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for ssum3.thy @@ -15,82 +15,82 @@ qed_goal "contlub_Isinl" Ssum3.thy "contlub(Isinl)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac trans 1), - (rtac (thelub_ssum1a RS sym) 2), - (rtac allI 3), - (rtac exI 3), - (rtac refl 3), - (etac (monofun_Isinl RS ch2ch_monofun) 2), - (res_inst_tac [("Q","lub(range(Y))=UU")] classical2 1), - (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1), - (atac 1), - (res_inst_tac [("f","Isinl")] arg_cong 1), - (rtac (chain_UU_I_inverse RS sym) 1), - (rtac allI 1), - (res_inst_tac [("s","UU"),("t","Y(i)")] ssubst 1), - (etac (chain_UU_I RS spec ) 1), - (atac 1), - (rtac Iwhen1 1), - (res_inst_tac [("f","Isinl")] arg_cong 1), - (rtac lub_equal 1), - (atac 1), - (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1), - (etac (monofun_Isinl RS ch2ch_monofun) 1), - (rtac allI 1), - (res_inst_tac [("Q","Y(k)=UU")] classical2 1), - (asm_simp_tac Ssum0_ss 1), - (asm_simp_tac Ssum0_ss 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac trans 1), + (rtac (thelub_ssum1a RS sym) 2), + (rtac allI 3), + (rtac exI 3), + (rtac refl 3), + (etac (monofun_Isinl RS ch2ch_monofun) 2), + (res_inst_tac [("Q","lub(range(Y))=UU")] classical2 1), + (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1), + (atac 1), + (res_inst_tac [("f","Isinl")] arg_cong 1), + (rtac (chain_UU_I_inverse RS sym) 1), + (rtac allI 1), + (res_inst_tac [("s","UU"),("t","Y(i)")] ssubst 1), + (etac (chain_UU_I RS spec ) 1), + (atac 1), + (rtac Iwhen1 1), + (res_inst_tac [("f","Isinl")] arg_cong 1), + (rtac lub_equal 1), + (atac 1), + (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1), + (etac (monofun_Isinl RS ch2ch_monofun) 1), + (rtac allI 1), + (res_inst_tac [("Q","Y(k)=UU")] classical2 1), + (asm_simp_tac Ssum0_ss 1), + (asm_simp_tac Ssum0_ss 1) + ]); qed_goal "contlub_Isinr" Ssum3.thy "contlub(Isinr)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac trans 1), - (rtac (thelub_ssum1b RS sym) 2), - (rtac allI 3), - (rtac exI 3), - (rtac refl 3), - (etac (monofun_Isinr RS ch2ch_monofun) 2), - (res_inst_tac [("Q","lub(range(Y))=UU")] classical2 1), - (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1), - (atac 1), - ((rtac arg_cong 1) THEN (rtac (chain_UU_I_inverse RS sym) 1)), - (rtac allI 1), - (res_inst_tac [("s","UU"),("t","Y(i)")] ssubst 1), - (etac (chain_UU_I RS spec ) 1), - (atac 1), - (rtac (strict_IsinlIsinr RS subst) 1), - (rtac Iwhen1 1), - ((rtac arg_cong 1) THEN (rtac lub_equal 1)), - (atac 1), - (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1), - (etac (monofun_Isinr RS ch2ch_monofun) 1), - (rtac allI 1), - (res_inst_tac [("Q","Y(k)=UU")] classical2 1), - (asm_simp_tac Ssum0_ss 1), - (asm_simp_tac Ssum0_ss 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac trans 1), + (rtac (thelub_ssum1b RS sym) 2), + (rtac allI 3), + (rtac exI 3), + (rtac refl 3), + (etac (monofun_Isinr RS ch2ch_monofun) 2), + (res_inst_tac [("Q","lub(range(Y))=UU")] classical2 1), + (res_inst_tac [("s","UU"),("t","lub(range(Y))")] ssubst 1), + (atac 1), + ((rtac arg_cong 1) THEN (rtac (chain_UU_I_inverse RS sym) 1)), + (rtac allI 1), + (res_inst_tac [("s","UU"),("t","Y(i)")] ssubst 1), + (etac (chain_UU_I RS spec ) 1), + (atac 1), + (rtac (strict_IsinlIsinr RS subst) 1), + (rtac Iwhen1 1), + ((rtac arg_cong 1) THEN (rtac lub_equal 1)), + (atac 1), + (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1), + (etac (monofun_Isinr RS ch2ch_monofun) 1), + (rtac allI 1), + (res_inst_tac [("Q","Y(k)=UU")] classical2 1), + (asm_simp_tac Ssum0_ss 1), + (asm_simp_tac Ssum0_ss 1) + ]); qed_goal "cont_Isinl" Ssum3.thy "cont(Isinl)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Isinl 1), - (rtac contlub_Isinl 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Isinl 1), + (rtac contlub_Isinl 1) + ]); qed_goal "cont_Isinr" Ssum3.thy "cont(Isinr)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Isinr 1), - (rtac contlub_Isinr 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Isinr 1), + (rtac contlub_Isinr 1) + ]); (* ------------------------------------------------------------------------ *) @@ -99,47 +99,47 @@ qed_goal "contlub_Iwhen1" Ssum3.thy "contlub(Iwhen)" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac trans 1), - (rtac (thelub_fun RS sym) 2), - (etac (monofun_Iwhen1 RS ch2ch_monofun) 2), - (rtac (expand_fun_eq RS iffD2) 1), - (strip_tac 1), - (rtac trans 1), - (rtac (thelub_fun RS sym) 2), - (rtac ch2ch_fun 2), - (etac (monofun_Iwhen1 RS ch2ch_monofun) 2), - (rtac (expand_fun_eq RS iffD2) 1), - (strip_tac 1), - (res_inst_tac [("p","xa")] IssumE 1), - (asm_simp_tac Ssum0_ss 1), - (rtac (lub_const RS thelubI RS sym) 1), - (asm_simp_tac Ssum0_ss 1), - (etac contlub_cfun_fun 1), - (asm_simp_tac Ssum0_ss 1), - (rtac (lub_const RS thelubI RS sym) 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac trans 1), + (rtac (thelub_fun RS sym) 2), + (etac (monofun_Iwhen1 RS ch2ch_monofun) 2), + (rtac (expand_fun_eq RS iffD2) 1), + (strip_tac 1), + (rtac trans 1), + (rtac (thelub_fun RS sym) 2), + (rtac ch2ch_fun 2), + (etac (monofun_Iwhen1 RS ch2ch_monofun) 2), + (rtac (expand_fun_eq RS iffD2) 1), + (strip_tac 1), + (res_inst_tac [("p","xa")] IssumE 1), + (asm_simp_tac Ssum0_ss 1), + (rtac (lub_const RS thelubI RS sym) 1), + (asm_simp_tac Ssum0_ss 1), + (etac contlub_cfun_fun 1), + (asm_simp_tac Ssum0_ss 1), + (rtac (lub_const RS thelubI RS sym) 1) + ]); qed_goal "contlub_Iwhen2" Ssum3.thy "contlub(Iwhen(f))" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (rtac trans 1), - (rtac (thelub_fun RS sym) 2), - (etac (monofun_Iwhen2 RS ch2ch_monofun) 2), - (rtac (expand_fun_eq RS iffD2) 1), - (strip_tac 1), - (res_inst_tac [("p","x")] IssumE 1), - (asm_simp_tac Ssum0_ss 1), - (rtac (lub_const RS thelubI RS sym) 1), - (asm_simp_tac Ssum0_ss 1), - (rtac (lub_const RS thelubI RS sym) 1), - (asm_simp_tac Ssum0_ss 1), - (etac contlub_cfun_fun 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (rtac trans 1), + (rtac (thelub_fun RS sym) 2), + (etac (monofun_Iwhen2 RS ch2ch_monofun) 2), + (rtac (expand_fun_eq RS iffD2) 1), + (strip_tac 1), + (res_inst_tac [("p","x")] IssumE 1), + (asm_simp_tac Ssum0_ss 1), + (rtac (lub_const RS thelubI RS sym) 1), + (asm_simp_tac Ssum0_ss 1), + (rtac (lub_const RS thelubI RS sym) 1), + (asm_simp_tac Ssum0_ss 1), + (etac contlub_cfun_fun 1) + ]); (* ------------------------------------------------------------------------ *) (* continuity for Iwhen in its third argument *) @@ -152,219 +152,219 @@ qed_goal "ssum_lemma9" Ssum3.thy "[| is_chain(Y); lub(range(Y)) = Isinl(x)|] ==> !i.? x.Y(i)=Isinl(x)" (fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (res_inst_tac [("p","Y(i)")] IssumE 1), - (etac exI 1), - (etac exI 1), - (res_inst_tac [("P","y=UU")] notE 1), - (atac 1), - (rtac (less_ssum3d RS iffD1) 1), - (etac subst 1), - (etac subst 1), - (etac is_ub_thelub 1) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (res_inst_tac [("p","Y(i)")] IssumE 1), + (etac exI 1), + (etac exI 1), + (res_inst_tac [("P","y=UU")] notE 1), + (atac 1), + (rtac (less_ssum3d RS iffD1) 1), + (etac subst 1), + (etac subst 1), + (etac is_ub_thelub 1) + ]); qed_goal "ssum_lemma10" Ssum3.thy "[| is_chain(Y); lub(range(Y)) = Isinr(x)|] ==> !i.? x.Y(i)=Isinr(x)" (fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (res_inst_tac [("p","Y(i)")] IssumE 1), - (rtac exI 1), - (etac trans 1), - (rtac strict_IsinlIsinr 1), - (etac exI 2), - (res_inst_tac [("P","xa=UU")] notE 1), - (atac 1), - (rtac (less_ssum3c RS iffD1) 1), - (etac subst 1), - (etac subst 1), - (etac is_ub_thelub 1) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (res_inst_tac [("p","Y(i)")] IssumE 1), + (rtac exI 1), + (etac trans 1), + (rtac strict_IsinlIsinr 1), + (etac exI 2), + (res_inst_tac [("P","xa=UU")] notE 1), + (atac 1), + (rtac (less_ssum3c RS iffD1) 1), + (etac subst 1), + (etac subst 1), + (etac is_ub_thelub 1) + ]); qed_goal "ssum_lemma11" Ssum3.thy "[| is_chain(Y); lub(range(Y)) = Isinl(UU) |] ==>\ \ Iwhen f g (lub(range Y)) = lub(range(%i. Iwhen f g (Y i)))" (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac Ssum0_ss 1), - (rtac (chain_UU_I_inverse RS sym) 1), - (rtac allI 1), - (res_inst_tac [("s","Isinl(UU)"),("t","Y(i)")] subst 1), - (rtac (inst_ssum_pcpo RS subst) 1), - (rtac (chain_UU_I RS spec RS sym) 1), - (atac 1), - (etac (inst_ssum_pcpo RS ssubst) 1), - (asm_simp_tac Ssum0_ss 1) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac Ssum0_ss 1), + (rtac (chain_UU_I_inverse RS sym) 1), + (rtac allI 1), + (res_inst_tac [("s","Isinl(UU)"),("t","Y(i)")] subst 1), + (rtac (inst_ssum_pcpo RS subst) 1), + (rtac (chain_UU_I RS spec RS sym) 1), + (atac 1), + (etac (inst_ssum_pcpo RS ssubst) 1), + (asm_simp_tac Ssum0_ss 1) + ]); qed_goal "ssum_lemma12" Ssum3.thy "[| is_chain(Y); lub(range(Y)) = Isinl(x); x ~= UU |] ==>\ \ Iwhen f g (lub(range Y)) = lub(range(%i. Iwhen f g (Y i)))" (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac Ssum0_ss 1), - (res_inst_tac [("t","x")] subst 1), - (rtac inject_Isinl 1), - (rtac trans 1), - (atac 2), - (rtac (thelub_ssum1a RS sym) 1), - (atac 1), - (etac ssum_lemma9 1), - (atac 1), - (rtac trans 1), - (rtac contlub_cfun_arg 1), - (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1), - (atac 1), - (rtac lub_equal2 1), - (rtac (chain_mono2 RS exE) 1), - (atac 2), - (rtac chain_UU_I_inverse2 1), - (rtac (inst_ssum_pcpo RS ssubst) 1), - (etac swap 1), - (rtac inject_Isinl 1), - (rtac trans 1), - (etac sym 1), - (etac notnotD 1), - (rtac exI 1), - (strip_tac 1), - (rtac (ssum_lemma9 RS spec RS exE) 1), - (atac 1), - (atac 1), - (res_inst_tac [("t","Y(i)")] ssubst 1), - (atac 1), - (rtac trans 1), - (rtac cfun_arg_cong 1), - (rtac Iwhen2 1), - (res_inst_tac [("Pa","Y(i)=UU")] swap 1), - (fast_tac HOL_cs 1), - (rtac (inst_ssum_pcpo RS ssubst) 1), - (res_inst_tac [("t","Y(i)")] ssubst 1), - (atac 1), - (fast_tac HOL_cs 1), - (rtac (Iwhen2 RS ssubst) 1), - (res_inst_tac [("Pa","Y(i)=UU")] swap 1), - (fast_tac HOL_cs 1), - (rtac (inst_ssum_pcpo RS ssubst) 1), - (res_inst_tac [("t","Y(i)")] ssubst 1), - (atac 1), - (fast_tac HOL_cs 1), + [ + (cut_facts_tac prems 1), + (asm_simp_tac Ssum0_ss 1), + (res_inst_tac [("t","x")] subst 1), + (rtac inject_Isinl 1), + (rtac trans 1), + (atac 2), + (rtac (thelub_ssum1a RS sym) 1), + (atac 1), + (etac ssum_lemma9 1), + (atac 1), + (rtac trans 1), + (rtac contlub_cfun_arg 1), + (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1), + (atac 1), + (rtac lub_equal2 1), + (rtac (chain_mono2 RS exE) 1), + (atac 2), + (rtac chain_UU_I_inverse2 1), + (rtac (inst_ssum_pcpo RS ssubst) 1), + (etac swap 1), + (rtac inject_Isinl 1), + (rtac trans 1), + (etac sym 1), + (etac notnotD 1), + (rtac exI 1), + (strip_tac 1), + (rtac (ssum_lemma9 RS spec RS exE) 1), + (atac 1), + (atac 1), + (res_inst_tac [("t","Y(i)")] ssubst 1), + (atac 1), + (rtac trans 1), + (rtac cfun_arg_cong 1), + (rtac Iwhen2 1), + (res_inst_tac [("Pa","Y(i)=UU")] swap 1), + (fast_tac HOL_cs 1), + (rtac (inst_ssum_pcpo RS ssubst) 1), + (res_inst_tac [("t","Y(i)")] ssubst 1), + (atac 1), + (fast_tac HOL_cs 1), + (rtac (Iwhen2 RS ssubst) 1), + (res_inst_tac [("Pa","Y(i)=UU")] swap 1), + (fast_tac HOL_cs 1), + (rtac (inst_ssum_pcpo RS ssubst) 1), + (res_inst_tac [("t","Y(i)")] ssubst 1), + (atac 1), + (fast_tac HOL_cs 1), (simp_tac (simpset_of "Cfun3") 1), - (rtac (monofun_fapp2 RS ch2ch_monofun) 1), - (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), - (etac (monofun_Iwhen3 RS ch2ch_monofun) 1) - ]); + (rtac (monofun_fapp2 RS ch2ch_monofun) 1), + (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), + (etac (monofun_Iwhen3 RS ch2ch_monofun) 1) + ]); qed_goal "ssum_lemma13" Ssum3.thy "[| is_chain(Y); lub(range(Y)) = Isinr(x); x ~= UU |] ==>\ \ Iwhen f g (lub(range Y)) = lub(range(%i. Iwhen f g (Y i)))" (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac Ssum0_ss 1), - (res_inst_tac [("t","x")] subst 1), - (rtac inject_Isinr 1), - (rtac trans 1), - (atac 2), - (rtac (thelub_ssum1b RS sym) 1), - (atac 1), - (etac ssum_lemma10 1), - (atac 1), - (rtac trans 1), - (rtac contlub_cfun_arg 1), - (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1), - (atac 1), - (rtac lub_equal2 1), - (rtac (chain_mono2 RS exE) 1), - (atac 2), - (rtac chain_UU_I_inverse2 1), - (rtac (inst_ssum_pcpo RS ssubst) 1), - (etac swap 1), - (rtac inject_Isinr 1), - (rtac trans 1), - (etac sym 1), - (rtac (strict_IsinlIsinr RS subst) 1), - (etac notnotD 1), - (rtac exI 1), - (strip_tac 1), - (rtac (ssum_lemma10 RS spec RS exE) 1), - (atac 1), - (atac 1), - (res_inst_tac [("t","Y(i)")] ssubst 1), - (atac 1), - (rtac trans 1), - (rtac cfun_arg_cong 1), - (rtac Iwhen3 1), - (res_inst_tac [("Pa","Y(i)=UU")] swap 1), - (fast_tac HOL_cs 1), - (dtac notnotD 1), - (rtac (inst_ssum_pcpo RS ssubst) 1), - (rtac (strict_IsinlIsinr RS ssubst) 1), - (res_inst_tac [("t","Y(i)")] ssubst 1), - (atac 1), - (fast_tac HOL_cs 1), - (rtac (Iwhen3 RS ssubst) 1), - (res_inst_tac [("Pa","Y(i)=UU")] swap 1), - (fast_tac HOL_cs 1), - (dtac notnotD 1), - (rtac (inst_ssum_pcpo RS ssubst) 1), - (rtac (strict_IsinlIsinr RS ssubst) 1), - (res_inst_tac [("t","Y(i)")] ssubst 1), - (atac 1), - (fast_tac HOL_cs 1), + [ + (cut_facts_tac prems 1), + (asm_simp_tac Ssum0_ss 1), + (res_inst_tac [("t","x")] subst 1), + (rtac inject_Isinr 1), + (rtac trans 1), + (atac 2), + (rtac (thelub_ssum1b RS sym) 1), + (atac 1), + (etac ssum_lemma10 1), + (atac 1), + (rtac trans 1), + (rtac contlub_cfun_arg 1), + (rtac (monofun_Iwhen3 RS ch2ch_monofun) 1), + (atac 1), + (rtac lub_equal2 1), + (rtac (chain_mono2 RS exE) 1), + (atac 2), + (rtac chain_UU_I_inverse2 1), + (rtac (inst_ssum_pcpo RS ssubst) 1), + (etac swap 1), + (rtac inject_Isinr 1), + (rtac trans 1), + (etac sym 1), + (rtac (strict_IsinlIsinr RS subst) 1), + (etac notnotD 1), + (rtac exI 1), + (strip_tac 1), + (rtac (ssum_lemma10 RS spec RS exE) 1), + (atac 1), + (atac 1), + (res_inst_tac [("t","Y(i)")] ssubst 1), + (atac 1), + (rtac trans 1), + (rtac cfun_arg_cong 1), + (rtac Iwhen3 1), + (res_inst_tac [("Pa","Y(i)=UU")] swap 1), + (fast_tac HOL_cs 1), + (dtac notnotD 1), + (rtac (inst_ssum_pcpo RS ssubst) 1), + (rtac (strict_IsinlIsinr RS ssubst) 1), + (res_inst_tac [("t","Y(i)")] ssubst 1), + (atac 1), + (fast_tac HOL_cs 1), + (rtac (Iwhen3 RS ssubst) 1), + (res_inst_tac [("Pa","Y(i)=UU")] swap 1), + (fast_tac HOL_cs 1), + (dtac notnotD 1), + (rtac (inst_ssum_pcpo RS ssubst) 1), + (rtac (strict_IsinlIsinr RS ssubst) 1), + (res_inst_tac [("t","Y(i)")] ssubst 1), + (atac 1), + (fast_tac HOL_cs 1), (simp_tac (simpset_of "Cfun3") 1), - (rtac (monofun_fapp2 RS ch2ch_monofun) 1), - (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), - (etac (monofun_Iwhen3 RS ch2ch_monofun) 1) - ]); + (rtac (monofun_fapp2 RS ch2ch_monofun) 1), + (etac (monofun_Iwhen3 RS ch2ch_monofun) 1), + (etac (monofun_Iwhen3 RS ch2ch_monofun) 1) + ]); qed_goal "contlub_Iwhen3" Ssum3.thy "contlub(Iwhen(f)(g))" (fn prems => - [ - (rtac contlubI 1), - (strip_tac 1), - (res_inst_tac [("p","lub(range(Y))")] IssumE 1), - (etac ssum_lemma11 1), - (atac 1), - (etac ssum_lemma12 1), - (atac 1), - (atac 1), - (etac ssum_lemma13 1), - (atac 1), - (atac 1) - ]); + [ + (rtac contlubI 1), + (strip_tac 1), + (res_inst_tac [("p","lub(range(Y))")] IssumE 1), + (etac ssum_lemma11 1), + (atac 1), + (etac ssum_lemma12 1), + (atac 1), + (atac 1), + (etac ssum_lemma13 1), + (atac 1), + (atac 1) + ]); qed_goal "cont_Iwhen1" Ssum3.thy "cont(Iwhen)" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Iwhen1 1), - (rtac contlub_Iwhen1 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Iwhen1 1), + (rtac contlub_Iwhen1 1) + ]); qed_goal "cont_Iwhen2" Ssum3.thy "cont(Iwhen(f))" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Iwhen2 1), - (rtac contlub_Iwhen2 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Iwhen2 1), + (rtac contlub_Iwhen2 1) + ]); qed_goal "cont_Iwhen3" Ssum3.thy "cont(Iwhen(f)(g))" (fn prems => - [ - (rtac monocontlub2cont 1), - (rtac monofun_Iwhen3 1), - (rtac contlub_Iwhen3 1) - ]); + [ + (rtac monocontlub2cont 1), + (rtac monofun_Iwhen3 1), + (rtac contlub_Iwhen3 1) + ]); (* ------------------------------------------------------------------------ *) (* continuous versions of lemmas for 'a ++ 'b *) @@ -372,352 +372,352 @@ 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) - ]); + [ + (simp_tac (Ssum0_ss addsimps [cont_Isinl]) 1), + (rtac (inst_ssum_pcpo RS sym) 1) + ]); 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) - ]); + [ + (simp_tac (Ssum0_ss addsimps [cont_Isinr]) 1), + (rtac (inst_ssum_pcpo RS sym) 1) + ]); qed_goalw "noteq_sinlsinr" Ssum3.thy [sinl_def,sinr_def] - "sinl`a=sinr`b ==> a=UU & b=UU" + "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) - ]); + [ + (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) + ]); qed_goalw "inject_sinl" Ssum3.thy [sinl_def,sinr_def] - "sinl`a1=sinl`a2==> a1=a2" + "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) - ]); + [ + (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) + ]); qed_goalw "inject_sinr" Ssum3.thy [sinl_def,sinr_def] - "sinr`a1=sinr`a2==> a1=a2" + "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) - ]); + [ + (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) + ]); qed_goal "defined_sinl" Ssum3.thy - "x~=UU ==> sinl`x ~= UU" + "x~=UU ==> sinl`x ~= UU" (fn prems => - [ - (cut_facts_tac prems 1), - (etac swap 1), - (rtac inject_sinl 1), - (rtac (strict_sinl RS ssubst) 1), - (etac notnotD 1) - ]); + [ + (cut_facts_tac prems 1), + (etac swap 1), + (rtac inject_sinl 1), + (rtac (strict_sinl RS ssubst) 1), + (etac notnotD 1) + ]); qed_goal "defined_sinr" Ssum3.thy - "x~=UU ==> sinr`x ~= UU" + "x~=UU ==> sinr`x ~= UU" (fn prems => - [ - (cut_facts_tac prems 1), - (etac swap 1), - (rtac inject_sinr 1), - (rtac (strict_sinr RS ssubst) 1), - (etac notnotD 1) - ]); + [ + (cut_facts_tac prems 1), + (etac swap 1), + (rtac inject_sinr 1), + (rtac (strict_sinr RS ssubst) 1), + (etac notnotD 1) + ]); qed_goalw "Exh_Ssum1" Ssum3.thy [sinl_def,sinr_def] - "z=UU | (? a. z=sinl`a & a~=UU) | (? b. z=sinr`b & b~=UU)" + "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), - (rtac (inst_ssum_pcpo RS ssubst) 1), - (rtac Exh_Ssum 1) - ]); + [ + (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1), + (rtac (inst_ssum_pcpo RS ssubst) 1), + (rtac Exh_Ssum 1) + ]); qed_goalw "ssumE" Ssum3.thy [sinl_def,sinr_def] - "[|p=UU ==> Q ;\ -\ !!x.[|p=sinl`x; x~=UU |] ==> Q;\ -\ !!y.[|p=sinr`y; y~=UU |] ==> Q|] ==> Q" + "[|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), - (rtac (inst_ssum_pcpo RS ssubst) 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) - ]); + [ + (rtac IssumE 1), + (resolve_tac prems 1), + (rtac (inst_ssum_pcpo RS ssubst) 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) + ]); qed_goalw "ssumE2" Ssum3.thy [sinl_def,sinr_def] "[|!!x.[|p=sinl`x|] ==> Q;\ -\ !!y.[|p=sinr`y|] ==> Q|] ==> Q" +\ !!y.[|p=sinr`y|] ==> Q|] ==> Q" (fn prems => - [ - (rtac IssumE2 1), - (resolve_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Isinl 1), - (atac 1), - (resolve_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_Isinr 1), - (atac 1) - ]); + [ + (rtac IssumE2 1), + (resolve_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Isinl 1), + (atac 1), + (resolve_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_Isinr 1), + (atac 1) + ]); qed_goalw "sswhen1" Ssum3.thy [sswhen_def,sinl_def,sinr_def] - "sswhen`f`g`UU = UU" + "sswhen`f`g`UU = UU" (fn prems => - [ - (rtac (inst_ssum_pcpo RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont2cont_CF1L]) 1)), - (simp_tac Ssum0_ss 1) - ]); + [ + (rtac (inst_ssum_pcpo RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont2cont_CF1L]) 1)), + (simp_tac Ssum0_ss 1) + ]); qed_goalw "sswhen2" Ssum3.thy [sswhen_def,sinl_def,sinr_def] - "x~=UU==> sswhen`f`g`(sinl`x) = f`x" + "x~=UU==> sswhen`f`g`(sinl`x) = f`x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (asm_simp_tac Ssum0_ss 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (asm_simp_tac Ssum0_ss 1) + ]); qed_goalw "sswhen3" Ssum3.thy [sswhen_def,sinl_def,sinr_def] - "x~=UU==> sswhen`f`g`(sinr`x) = g`x" + "x~=UU==> sswhen`f`g`(sinr`x) = g`x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (asm_simp_tac Ssum0_ss 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (asm_simp_tac Ssum0_ss 1) + ]); qed_goalw "less_ssum4a" Ssum3.thy [sinl_def,sinr_def] - "(sinl`x << sinl`y) = (x << y)" + "(sinl`x << sinl`y) = (x << y)" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac less_ssum3a 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac less_ssum3a 1) + ]); qed_goalw "less_ssum4b" Ssum3.thy [sinl_def,sinr_def] - "(sinr`x << sinr`y) = (x << y)" + "(sinr`x << sinr`y) = (x << y)" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac less_ssum3b 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac less_ssum3b 1) + ]); qed_goalw "less_ssum4c" Ssum3.thy [sinl_def,sinr_def] - "(sinl`x << sinr`y) = (x = UU)" + "(sinl`x << sinr`y) = (x = UU)" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac less_ssum3c 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac less_ssum3c 1) + ]); qed_goalw "less_ssum4d" Ssum3.thy [sinl_def,sinr_def] - "(sinr`x << sinl`y) = (x = UU)" + "(sinr`x << sinl`y) = (x = UU)" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac less_ssum3d 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac less_ssum3d 1) + ]); qed_goalw "ssum_chainE" Ssum3.thy [sinl_def,sinr_def] - "is_chain(Y) ==> (!i.? x.(Y i)=sinl`x)|(!i.? y.(Y i)=sinr`y)" + "is_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) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (Ssum0_ss addsimps [cont_Isinr,cont_Isinl]) 1), + (etac ssum_lemma4 1) + ]); qed_goalw "thelub_ssum2a" Ssum3.thy [sinl_def,sinr_def,sswhen_def] "[| is_chain(Y); !i.? x. Y(i) = sinl`x |] ==>\ \ lub(range(Y)) = sinl`(lub(range(%i. sswhen`(LAM x. x)`(LAM y. UU)`(Y i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ext RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (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) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ext RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (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) + ]); qed_goalw "thelub_ssum2b" Ssum3.thy [sinl_def,sinr_def,sswhen_def] "[| is_chain(Y); !i.? x. Y(i) = sinr`x |] ==>\ \ lub(range(Y)) = sinr`(lub(range(%i. sswhen`(LAM y. UU)`(LAM x.x)`(Y i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (rtac (beta_cfun RS ext RS ssubst) 1), - (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, - cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), - (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) - ]); + [ + (cut_facts_tac prems 1), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (rtac (beta_cfun RS ext RS ssubst) 1), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), + (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) + ]); qed_goalw "thelub_ssum2a_rev" Ssum3.thy [sinl_def,sinr_def] - "[| is_chain(Y); lub(range(Y)) = sinl`x|] ==> !i.? x.Y(i)=sinl`x" + "[| is_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) - ]); + [ + (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) + ]); qed_goalw "thelub_ssum2b_rev" Ssum3.thy [sinl_def,sinr_def] - "[| is_chain(Y); lub(range(Y)) = sinr`x|] ==> !i.? x.Y(i)=sinr`x" + "[| is_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) - ]); + [ + (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) + ]); qed_goal "thelub_ssum3" Ssum3.thy "is_chain(Y) ==>\ \ lub(range(Y)) = sinl`(lub(range(%i. sswhen`(LAM x. x)`(LAM y.UU)`(Y i))))\ \ | lub(range(Y)) = sinr`(lub(range(%i. sswhen`(LAM y. UU)`(LAM x.x)`(Y i))))" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (ssum_chainE RS disjE) 1), - (atac 1), - (rtac disjI1 1), - (etac thelub_ssum2a 1), - (atac 1), - (rtac disjI2 1), - (etac thelub_ssum2b 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (ssum_chainE RS disjE) 1), + (atac 1), + (rtac disjI1 1), + (etac thelub_ssum2a 1), + (atac 1), + (rtac disjI2 1), + (etac thelub_ssum2b 1), + (atac 1) + ]); qed_goal "sswhen4" Ssum3.thy - "sswhen`sinl`sinr`z=z" + "sswhen`sinl`sinr`z=z" (fn prems => - [ - (res_inst_tac [("p","z")] ssumE 1), + [ + (res_inst_tac [("p","z")] ssumE 1), (asm_simp_tac ((simpset_of "Cfun3") addsimps [sswhen1,sswhen2,sswhen3]) 1), (asm_simp_tac ((simpset_of "Cfun3") addsimps [sswhen1,sswhen2,sswhen3]) 1), (asm_simp_tac ((simpset_of "Cfun3") addsimps [sswhen1,sswhen2,sswhen3]) 1) - ]); + ]); (* ------------------------------------------------------------------------ *) @@ -725,7 +725,7 @@ (* ------------------------------------------------------------------------ *) val Ssum_rews = [strict_sinl,strict_sinr,defined_sinl,defined_sinr, - sswhen1,sswhen2,sswhen3]; + sswhen1,sswhen2,sswhen3]; Addsimps [strict_sinl,strict_sinr,defined_sinl,defined_sinr, - sswhen1,sswhen2,sswhen3]; + sswhen1,sswhen2,sswhen3]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Tr1.ML --- a/src/HOLCF/Tr1.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Tr1.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/tr1.ML +(* Title: HOLCF/tr1.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for tr1.thy @@ -15,60 +15,60 @@ val dist_less_tr = [ prove_goalw Tr1.thy [TT_def] "~TT << UU" (fn prems => - [ - (rtac classical3 1), - (rtac defined_sinl 1), - (rtac not_less2not_eq 1), - (resolve_tac dist_less_one 1), - (rtac (rep_tr_iso RS subst) 1), - (rtac (rep_tr_iso RS subst) 1), - (rtac cfun_arg_cong 1), - (rtac ((abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) RS iso_strict ) RS conjunct2 RS ssubst) 1), - (etac (eq_UU_iff RS ssubst) 1) - ]), + [ + (rtac classical3 1), + (rtac defined_sinl 1), + (rtac not_less2not_eq 1), + (resolve_tac dist_less_one 1), + (rtac (rep_tr_iso RS subst) 1), + (rtac (rep_tr_iso RS subst) 1), + (rtac cfun_arg_cong 1), + (rtac ((abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) RS iso_strict ) RS conjunct2 RS ssubst) 1), + (etac (eq_UU_iff RS ssubst) 1) + ]), prove_goalw Tr1.thy [FF_def] "~FF << UU" (fn prems => - [ - (rtac classical3 1), - (rtac defined_sinr 1), - (rtac not_less2not_eq 1), - (resolve_tac dist_less_one 1), - (rtac (rep_tr_iso RS subst) 1), - (rtac (rep_tr_iso RS subst) 1), - (rtac cfun_arg_cong 1), - (rtac ((abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) RS iso_strict ) RS conjunct2 RS ssubst) 1), - (etac (eq_UU_iff RS ssubst) 1) - ]), + [ + (rtac classical3 1), + (rtac defined_sinr 1), + (rtac not_less2not_eq 1), + (resolve_tac dist_less_one 1), + (rtac (rep_tr_iso RS subst) 1), + (rtac (rep_tr_iso RS subst) 1), + (rtac cfun_arg_cong 1), + (rtac ((abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) RS iso_strict ) RS conjunct2 RS ssubst) 1), + (etac (eq_UU_iff RS ssubst) 1) + ]), prove_goalw Tr1.thy [FF_def,TT_def] "~TT << FF" (fn prems => - [ - (rtac classical3 1), - (rtac (less_ssum4c RS iffD1) 2), - (rtac not_less2not_eq 1), - (resolve_tac dist_less_one 1), - (rtac (rep_tr_iso RS subst) 1), - (rtac (rep_tr_iso RS subst) 1), - (etac monofun_cfun_arg 1) - ]), + [ + (rtac classical3 1), + (rtac (less_ssum4c RS iffD1) 2), + (rtac not_less2not_eq 1), + (resolve_tac dist_less_one 1), + (rtac (rep_tr_iso RS subst) 1), + (rtac (rep_tr_iso RS subst) 1), + (etac monofun_cfun_arg 1) + ]), prove_goalw Tr1.thy [FF_def,TT_def] "~FF << TT" (fn prems => - [ - (rtac classical3 1), - (rtac (less_ssum4d RS iffD1) 2), - (rtac not_less2not_eq 1), - (resolve_tac dist_less_one 1), - (rtac (rep_tr_iso RS subst) 1), - (rtac (rep_tr_iso RS subst) 1), - (etac monofun_cfun_arg 1) - ]) + [ + (rtac classical3 1), + (rtac (less_ssum4d RS iffD1) 2), + (rtac not_less2not_eq 1), + (resolve_tac dist_less_one 1), + (rtac (rep_tr_iso RS subst) 1), + (rtac (rep_tr_iso RS subst) 1), + (etac monofun_cfun_arg 1) + ]) ]; fun prover s = prove_goal Tr1.thy s (fn prems => - [ - (rtac not_less2not_eq 1), - (resolve_tac dist_less_tr 1) - ]); + [ + (rtac not_less2not_eq 1), + (resolve_tac dist_less_tr 1) + ]); val dist_eq_tr = map prover ["TT~=UU","FF~=UU","TT~=FF"]; val dist_eq_tr = dist_eq_tr @ (map (fn thm => (thm RS not_sym)) dist_eq_tr); @@ -79,44 +79,44 @@ qed_goalw "Exh_tr" Tr1.thy [FF_def,TT_def] "t=UU | t = TT | t = FF" (fn prems => - [ - (res_inst_tac [("p","rep_tr`t")] ssumE 1), - (rtac disjI1 1), - (rtac ((abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) RS iso_strict ) - RS conjunct2 RS subst) 1), - (rtac (abs_tr_iso RS subst) 1), - (etac cfun_arg_cong 1), - (rtac disjI2 1), - (rtac disjI1 1), - (rtac (abs_tr_iso RS subst) 1), - (rtac cfun_arg_cong 1), - (etac trans 1), - (rtac cfun_arg_cong 1), - (rtac (Exh_one RS disjE) 1), - (contr_tac 1), - (atac 1), - (rtac disjI2 1), - (rtac disjI2 1), - (rtac (abs_tr_iso RS subst) 1), - (rtac cfun_arg_cong 1), - (etac trans 1), - (rtac cfun_arg_cong 1), - (rtac (Exh_one RS disjE) 1), - (contr_tac 1), - (atac 1) - ]); + [ + (res_inst_tac [("p","rep_tr`t")] ssumE 1), + (rtac disjI1 1), + (rtac ((abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) RS iso_strict ) + RS conjunct2 RS subst) 1), + (rtac (abs_tr_iso RS subst) 1), + (etac cfun_arg_cong 1), + (rtac disjI2 1), + (rtac disjI1 1), + (rtac (abs_tr_iso RS subst) 1), + (rtac cfun_arg_cong 1), + (etac trans 1), + (rtac cfun_arg_cong 1), + (rtac (Exh_one RS disjE) 1), + (contr_tac 1), + (atac 1), + (rtac disjI2 1), + (rtac disjI2 1), + (rtac (abs_tr_iso RS subst) 1), + (rtac cfun_arg_cong 1), + (etac trans 1), + (rtac cfun_arg_cong 1), + (rtac (Exh_one RS disjE) 1), + (contr_tac 1), + (atac 1) + ]); qed_goal "trE" Tr1.thy - "[| p=UU ==> Q; p = TT ==>Q; p = FF ==>Q|] ==>Q" + "[| p=UU ==> Q; p = TT ==>Q; p = FF ==>Q|] ==>Q" (fn prems => - [ - (rtac (Exh_tr RS disjE) 1), - (eresolve_tac prems 1), - (etac disjE 1), - (eresolve_tac prems 1), - (eresolve_tac prems 1) - ]); + [ + (rtac (Exh_tr RS disjE) 1), + (eresolve_tac prems 1), + (etac disjE 1), + (eresolve_tac prems 1), + (eresolve_tac prems 1) + ]); (* ------------------------------------------------------------------------ *) @@ -125,20 +125,20 @@ qed_goalw "flat_tr" Tr1.thy [is_flat_def] "is_flat(TT)" (fn prems => - [ - (rtac allI 1), - (rtac allI 1), - (res_inst_tac [("p","x")] trE 1), - (Asm_simp_tac 1), - (res_inst_tac [("p","y")] trE 1), - (asm_simp_tac (!simpset addsimps dist_less_tr) 1), - (asm_simp_tac (!simpset addsimps dist_less_tr) 1), - (asm_simp_tac (!simpset addsimps dist_less_tr) 1), - (res_inst_tac [("p","y")] trE 1), - (asm_simp_tac (!simpset addsimps dist_less_tr) 1), - (asm_simp_tac (!simpset addsimps dist_less_tr) 1), - (asm_simp_tac (!simpset addsimps dist_less_tr) 1) - ]); + [ + (rtac allI 1), + (rtac allI 1), + (res_inst_tac [("p","x")] trE 1), + (Asm_simp_tac 1), + (res_inst_tac [("p","y")] trE 1), + (asm_simp_tac (!simpset addsimps dist_less_tr) 1), + (asm_simp_tac (!simpset addsimps dist_less_tr) 1), + (asm_simp_tac (!simpset addsimps dist_less_tr) 1), + (res_inst_tac [("p","y")] trE 1), + (asm_simp_tac (!simpset addsimps dist_less_tr) 1), + (asm_simp_tac (!simpset addsimps dist_less_tr) 1), + (asm_simp_tac (!simpset addsimps dist_less_tr) 1) + ]); (* ------------------------------------------------------------------------ *) @@ -147,16 +147,16 @@ fun prover s = prove_goalw Tr1.thy [tr_when_def,TT_def,FF_def] s (fn prems => - [ - (Simp_tac 1), - (simp_tac (!simpset addsimps [(rep_tr_iso ), - (abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) - RS iso_strict) RS conjunct1]@dist_eq_one)1) - ]); + [ + (Simp_tac 1), + (simp_tac (!simpset addsimps [(rep_tr_iso ), + (abs_tr_iso RS allI) RS ((rep_tr_iso RS allI) + RS iso_strict) RS conjunct1]@dist_eq_one)1) + ]); val tr_when = map prover [ - "tr_when`x`y`UU = UU", - "tr_when`x`y`TT = x", - "tr_when`x`y`FF = y" - ]; + "tr_when`x`y`UU = UU", + "tr_when`x`y`TT = x", + "tr_when`x`y`FF = y" + ]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Tr2.ML --- a/src/HOLCF/Tr2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Tr2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/tr2.ML +(* Title: HOLCF/tr2.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for tr2.thy @@ -14,25 +14,25 @@ fun prover s = prove_goalw Tr2.thy [andalso_def] s (fn prems => - [ - (simp_tac (!simpset addsimps tr_when) 1) - ]); + [ + (simp_tac (!simpset addsimps tr_when) 1) + ]); val andalso_thms = map prover [ - "(TT andalso y) = y", - "(FF andalso y) = FF", - "(UU andalso y) = UU" - ]; + "(TT andalso y) = y", + "(FF andalso y) = FF", + "(UU andalso y) = UU" + ]; val andalso_thms = andalso_thms @ [prove_goalw Tr2.thy [andalso_def] "(x andalso TT) = x" (fn prems => - [ - (res_inst_tac [("p","x")] trE 1), - (asm_simp_tac (!simpset addsimps tr_when) 1), - (asm_simp_tac (!simpset addsimps tr_when) 1), - (asm_simp_tac (!simpset addsimps tr_when) 1) - ])]; + [ + (res_inst_tac [("p","x")] trE 1), + (asm_simp_tac (!simpset addsimps tr_when) 1), + (asm_simp_tac (!simpset addsimps tr_when) 1), + (asm_simp_tac (!simpset addsimps tr_when) 1) + ])]; (* ------------------------------------------------------------------------ *) (* lemmas about orelse *) @@ -40,25 +40,25 @@ fun prover s = prove_goalw Tr2.thy [orelse_def] s (fn prems => - [ - (simp_tac (!simpset addsimps tr_when) 1) - ]); + [ + (simp_tac (!simpset addsimps tr_when) 1) + ]); val orelse_thms = map prover [ - "(TT orelse y) = TT", - "(FF orelse y) = y", - "(UU orelse y) = UU" - ]; + "(TT orelse y) = TT", + "(FF orelse y) = y", + "(UU orelse y) = UU" + ]; val orelse_thms = orelse_thms @ [prove_goalw Tr2.thy [orelse_def] "(x orelse FF) = x" (fn prems => - [ - (res_inst_tac [("p","x")] trE 1), - (asm_simp_tac (!simpset addsimps tr_when) 1), - (asm_simp_tac (!simpset addsimps tr_when) 1), - (asm_simp_tac (!simpset addsimps tr_when) 1) - ])]; + [ + (res_inst_tac [("p","x")] trE 1), + (asm_simp_tac (!simpset addsimps tr_when) 1), + (asm_simp_tac (!simpset addsimps tr_when) 1), + (asm_simp_tac (!simpset addsimps tr_when) 1) + ])]; (* ------------------------------------------------------------------------ *) @@ -67,15 +67,15 @@ fun prover s = prove_goalw Tr2.thy [neg_def] s (fn prems => - [ - (simp_tac (!simpset addsimps tr_when) 1) - ]); + [ + (simp_tac (!simpset addsimps tr_when) 1) + ]); val neg_thms = map prover [ - "neg`TT = FF", - "neg`FF = TT", - "neg`UU = UU" - ]; + "neg`TT = FF", + "neg`FF = TT", + "neg`UU = UU" + ]; (* ------------------------------------------------------------------------ *) (* lemmas about If_then_else_fi *) @@ -83,16 +83,16 @@ fun prover s = prove_goalw Tr2.thy [ifte_def] s (fn prems => - [ - (simp_tac (!simpset addsimps tr_when) 1) - ]); + [ + (simp_tac (!simpset addsimps tr_when) 1) + ]); val ifte_thms = map prover [ - "If UU then e1 else e2 fi = UU", - "If FF then e1 else e2 fi = e2", - "If TT then e1 else e2 fi = e1"]; + "If UU then e1 else e2 fi = UU", + "If FF then e1 else e2 fi = e2", + "If TT then e1 else e2 fi = e1"]; - + diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/Void.ML --- a/src/HOLCF/Void.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/Void.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/void.ML +(* Title: HOLCF/void.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for void.thy. @@ -18,10 +18,10 @@ qed_goalw "VoidI" Void.thy [UU_void_Rep_def,Void_def] " UU_void_Rep : Void" (fn prems => - [ - (rtac (mem_Collect_eq RS ssubst) 1), - (rtac refl 1) - ]); + [ + (rtac (mem_Collect_eq RS ssubst) 1), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) (* less_void is a partial ordering on void *) @@ -29,27 +29,27 @@ qed_goalw "refl_less_void" Void.thy [ less_void_def ] "less_void x x" (fn prems => - [ - (fast_tac HOL_cs 1) - ]); + [ + (fast_tac HOL_cs 1) + ]); qed_goalw "antisym_less_void" Void.thy [ less_void_def ] - "[|less_void x y; less_void y x|] ==> x = y" + "[|less_void x y; less_void y x|] ==> x = y" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (Rep_Void_inverse RS subst) 1), - (etac subst 1), - (rtac (Rep_Void_inverse RS sym) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (Rep_Void_inverse RS subst) 1), + (etac subst 1), + (rtac (Rep_Void_inverse RS sym) 1) + ]); qed_goalw "trans_less_void" Void.thy [ less_void_def ] - "[|less_void x y; less_void y z|] ==> less_void x z" + "[|less_void x y; less_void y z|] ==> less_void x z" (fn prems => - [ - (cut_facts_tac prems 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (fast_tac HOL_cs 1) + ]); (* ------------------------------------------------------------------------ *) (* a technical lemma about void: *) @@ -58,12 +58,12 @@ qed_goal "unique_void" Void.thy "Rep_Void(x) = UU_void_Rep" (fn prems => - [ - (rtac (mem_Collect_eq RS subst) 1), - (fold_goals_tac [Void_def]), - (rtac Rep_Void 1) - ]); + [ + (rtac (mem_Collect_eq RS subst) 1), + (fold_goals_tac [Void_def]), + (rtac Rep_Void 1) + ]); - + diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/ax_ops/thy_ops.ML --- a/src/HOLCF/ax_ops/thy_ops.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/ax_ops/thy_ops.ML Tue Jan 30 13:42:57 1996 +0100 @@ -111,11 +111,11 @@ (* Changing a->b->c res. a*b->c to a=>b=>c. Called by syn_decls. *) (*####*) fun op_to_fun true sign (Type("->" ,[larg,t]))= - Type("fun",[larg,op_to_fun true sign t]) + Type("fun",[larg,op_to_fun true sign t]) | op_to_fun false sign (Type("->",[args,res])) = let - fun otf (Type("*",[larg,rargs])) = Type("fun",[larg,otf rargs]) - | otf t = Type("fun",[t,res]); - in otf args end + fun otf (Type("*",[larg,rargs])) = Type("fun",[larg,otf rargs]) + | otf t = Type("fun",[t,res]); + in otf args end | op_to_fun _ sign t = t(*error("Wrong type for cinfix/cmixfix : "^ (Sign.string_of_typ sign t))*); (*####*) @@ -167,12 +167,12 @@ ::xrules_of true tail (*####*) | xrules_of true ((name,typ,CMixfix(_))::tail) = let - fun argnames n (Type("->" ,[_,t]))= chr n :: argnames (n+1) t - | argnames _ _ = []; - val names = argnames (ord"A") typ; - in if names = [] then [] else [mk_appl (Constant name) (map Variable names)<-> - foldl (fn (t,arg) => (mk_appl (Constant "@fapp") [t,Variable arg])) - (Constant name,names)] end + fun argnames n (Type("->" ,[_,t]))= chr n :: argnames (n+1) t + | argnames _ _ = []; + val names = argnames (ord"A") typ; + in if names = [] then [] else [mk_appl (Constant name) (map Variable names)<-> + foldl (fn (t,arg) => (mk_appl (Constant "@fapp") [t,Variable arg])) + (Constant name,names)] end @xrules_of true tail (*####*) | xrules_of false ((name,typ,CInfixl(i))::tail) = @@ -187,20 +187,20 @@ ::xrules_of false tail (*####*) | xrules_of false ((name,typ,CMixfix(_))::tail) = let - fun foldr' f l = - let fun itr [] = raise LIST "foldr'" - | itr [a] = a - | itr (a::l) = f(a, itr l) - in itr l end; - fun argnames n (Type("*" ,[_,t]))= chr n :: argnames (n+1) t - | argnames n _ = [chr n]; - val vars = map Variable (case typ of (Type("->" ,[t,_])) =>argnames (ord"A") t - | _ => []); - in if vars = [] then [] else [mk_appl (Constant name) vars <-> - (mk_appl (Constant "@fapp") [Constant name, case vars of [v] => v - | args => mk_appl (Constant "@ctuple") [hd args,foldr' (fn (t,arg) => - mk_appl (Constant "_args") [t,arg]) (tl args)]])] - end + fun foldr' f l = + let fun itr [] = raise LIST "foldr'" + | itr [a] = a + | itr (a::l) = f(a, itr l) + in itr l end; + fun argnames n (Type("*" ,[_,t]))= chr n :: argnames (n+1) t + | argnames n _ = [chr n]; + val vars = map Variable (case typ of (Type("->" ,[t,_])) =>argnames (ord"A") t + | _ => []); + in if vars = [] then [] else [mk_appl (Constant name) vars <-> + (mk_appl (Constant "@fapp") [Constant name, case vars of [v] => v + | args => mk_appl (Constant "@ctuple") [hd args,foldr' (fn (t,arg) => + mk_appl (Constant "_args") [t,arg]) (tl args)]])] + end @xrules_of false tail (*####*) | xrules_of c ((name,typ,Mixfix(_))::tail) = xrules_of c tail diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/ax_ops/thy_syntax.ML --- a/src/HOLCF/ax_ops/thy_syntax.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/ax_ops/thy_syntax.ML Tue Jan 30 13:42:57 1996 +0100 @@ -21,13 +21,13 @@ open HOLCFlogic; val user_keywords = (*####*)filter_out (fn s => s mem (ThyAxioms.axioms_keywords@ - ThyOps.ops_keywords)) (*####*)ThySynData.user_keywords @ + ThyOps.ops_keywords)) (*####*)ThySynData.user_keywords @ ThyAxioms.axioms_keywords @ ThyOps.ops_keywords; val user_sections = (*####*)filter_out (fn (s,_) => s mem (map fst ( - ThyAxioms.axioms_sections@ ThyOps.ops_sections))) (*####*) - ThySynData.user_sections @ + ThyAxioms.axioms_sections@ ThyOps.ops_sections))) (*####*) + ThySynData.user_sections @ ThyAxioms.axioms_sections @ ThyOps.ops_sections; end; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/ccc1.ML --- a/src/HOLCF/ccc1.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/ccc1.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,7 +1,7 @@ -(* Title: HOLCF/ccc1.ML +(* Title: HOLCF/ccc1.ML ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen Lemmas for ccc1.thy *) @@ -16,30 +16,30 @@ qed_goalw "ID1" ccc1.thy [ID_def] "ID`x=x" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (rtac cont_id 1), - (rtac refl 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (rtac cont_id 1), + (rtac refl 1) + ]); qed_goalw "cfcomp1" ccc1.thy [oo_def] "(f oo g)=(LAM x.f`(g`x))" (fn prems => - [ - (rtac (beta_cfun RS ssubst) 1), - (cont_tacR 1), - (rtac (beta_cfun RS ssubst) 1), - (cont_tacR 1), - (rtac refl 1) - ]); + [ + (rtac (beta_cfun RS ssubst) 1), + (cont_tacR 1), + (rtac (beta_cfun RS ssubst) 1), + (cont_tacR 1), + (rtac refl 1) + ]); qed_goal "cfcomp2" ccc1.thy "(f oo g)`x=f`(g`x)" (fn prems => - [ - (rtac (cfcomp1 RS ssubst) 1), - (rtac (beta_cfun RS ssubst) 1), - (cont_tacR 1), - (rtac refl 1) - ]); + [ + (rtac (cfcomp1 RS ssubst) 1), + (rtac (beta_cfun RS ssubst) 1), + (cont_tacR 1), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) @@ -53,35 +53,35 @@ qed_goal "ID2" ccc1.thy "f oo ID = f " (fn prems => - [ - (rtac ext_cfun 1), - (rtac (cfcomp2 RS ssubst) 1), - (rtac (ID1 RS ssubst) 1), - (rtac refl 1) - ]); + [ + (rtac ext_cfun 1), + (rtac (cfcomp2 RS ssubst) 1), + (rtac (ID1 RS ssubst) 1), + (rtac refl 1) + ]); qed_goal "ID3" ccc1.thy "ID oo f = f " (fn prems => - [ - (rtac ext_cfun 1), - (rtac (cfcomp2 RS ssubst) 1), - (rtac (ID1 RS ssubst) 1), - (rtac refl 1) - ]); + [ + (rtac ext_cfun 1), + (rtac (cfcomp2 RS ssubst) 1), + (rtac (ID1 RS ssubst) 1), + (rtac refl 1) + ]); qed_goal "assoc_oo" ccc1.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), - (rtac (cfcomp2 RS ssubst) 1), - (rtac (cfcomp2 RS ssubst) 1), - (rtac refl 1), - (rtac (cfcomp2 RS ssubst) 1), - (rtac (cfcomp2 RS ssubst) 1), - (rtac refl 1) - ]); + [ + (rtac ext_cfun 1), + (res_inst_tac [("s","f`(g`(h`x))")] trans 1), + (rtac (cfcomp2 RS ssubst) 1), + (rtac (cfcomp2 RS ssubst) 1), + (rtac refl 1), + (rtac (cfcomp2 RS ssubst) 1), + (rtac (cfcomp2 RS ssubst) 1), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------ *) (* Merge the different rewrite rules for the simplifier *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/domain/axioms.ML --- a/src/HOLCF/domain/axioms.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/domain/axioms.ML Tue Jan 30 13:42:57 1996 +0100 @@ -35,63 +35,63 @@ val ax_rep_iso = (dname^"_rep_iso",mk_trp(dc_abs`(dc_rep`%x_name') === %x_name')); val ax_when_def = (dname^"_when_def",%%(dname^"_when") == - foldr (uncurry /\ ) (when_funs cons, /\x_name'((when_body cons (fn (x,y) => - Bound(1+length cons+x-y)))`(dc_rep`Bound 0)))); + foldr (uncurry /\ ) (when_funs cons, /\x_name'((when_body cons (fn (x,y) => + Bound(1+length cons+x-y)))`(dc_rep`Bound 0)))); val ax_copy_def = let fun simp_oo (Const ("fapp", _) $ (Const ("fapp", _) $ - Const ("cfcomp", _) $ fc) $ Const ("ID", _)) = fc + Const ("cfcomp", _) $ fc) $ Const ("ID", _)) = fc | simp_oo t = t; fun simp_app (Const ("fapp", _) $ Const ("ID", _) $ t) = t | simp_app t = t; - fun mk_arg m n arg = (if is_lazy arg - then fn t => %%"lift"`(simp_oo (%%"up" oo t)) else Id) - (if_rec arg (cproj (Bound (2*min[n,m])) eqs) (%%"ID")); - fun mk_prod (t1,t2) = %%"ssplit"`(/\ "x" (/\ "y" (%%"spair"` - simp_app(t1`Bound 1)`simp_app(t2`Bound 0)))); - fun one_con (_,args) = if args = [] then %%"ID" else - foldr' mk_prod (mapn (mk_arg (length args-1)) 1 args); - fun mk_sum (t1,t2) = %%"sswhen"`(simp_oo (%%"sinl" oo t1)) - `(simp_oo (%%"sinr" oo t2)); - in (dname^"_copy_def", %%(dname^"_copy") == /\"f" - (dc_abs oo foldr' mk_sum (map one_con cons) oo dc_rep)) end; + fun mk_arg m n arg = (if is_lazy arg + then fn t => %%"lift"`(simp_oo (%%"up" oo t)) else Id) + (if_rec arg (cproj (Bound (2*min[n,m])) eqs) (%%"ID")); + fun mk_prod (t1,t2) = %%"ssplit"`(/\ "x" (/\ "y" (%%"spair"` + simp_app(t1`Bound 1)`simp_app(t2`Bound 0)))); + fun one_con (_,args) = if args = [] then %%"ID" else + foldr' mk_prod (mapn (mk_arg (length args-1)) 1 args); + fun mk_sum (t1,t2) = %%"sswhen"`(simp_oo (%%"sinl" oo t1)) + `(simp_oo (%%"sinr" oo t2)); + in (dname^"_copy_def", %%(dname^"_copy") == /\"f" + (dc_abs oo foldr' mk_sum (map one_con cons) oo dc_rep)) end; (* ----- definitions concerning the constructors, discriminators and selectors ---- *) val axs_con_def = let - fun idxs z x arg = (if is_lazy arg then fn x => %%"up"`x else Id)(Bound(z-x)); - fun prms [] = %%"one" - | prms vs = foldr' (fn(x,t)=> %%"spair"`x`t) (mapn (idxs (length vs)) 1 vs); - val injs = bin_branchr (fn l=> l@["l"]) (fn l=> l@["r"]); - fun cdef ((con,args),injs) = (extern_name con ^"_def",%%con == - foldr /\# (args,dc_abs` - (foldr (fn (i,t) => %%("sin"^i)`t ) (injs, prms args)))); - in map cdef (cons~~(mapn (fn n => K(injs [] cons n)) 0 cons)) end; + fun idxs z x arg = (if is_lazy arg then fn x => %%"up"`x else Id)(Bound(z-x)); + fun prms [] = %%"one" + | prms vs = foldr' (fn(x,t)=> %%"spair"`x`t) (mapn (idxs (length vs)) 1 vs); + val injs = bin_branchr (fn l=> l@["l"]) (fn l=> l@["r"]); + fun cdef ((con,args),injs) = (extern_name con ^"_def",%%con == + foldr /\# (args,dc_abs` + (foldr (fn (i,t) => %%("sin"^i)`t ) (injs, prms args)))); + in map cdef (cons~~(mapn (fn n => K(injs [] cons n)) 0 cons)) end; val axs_dis_def = let - fun ddef (con,_) = (dis_name con ^"_def",%%(dis_name con) == - mk_cfapp(%%(dname^"_when"),map - (fn (con',args) => (foldr /\# - (args,if con'=con then %%"TT" else %%"FF"))) cons)) - in map ddef cons end; + fun ddef (con,_) = (dis_name con ^"_def",%%(dis_name con) == + mk_cfapp(%%(dname^"_when"),map + (fn (con',args) => (foldr /\# + (args,if con'=con then %%"TT" else %%"FF"))) cons)) + in map ddef cons end; val axs_sel_def = let - fun sdef con n arg = (sel_of arg^"_def",%%(sel_of arg) == - mk_cfapp(%%(dname^"_when"),map - (fn (con',args) => if con'<>con then %%"UU" else - foldr /\# (args,Bound (length args - n))) cons)); - in flat(map (fn (con,args) => mapn (sdef con) 1 args) cons) end; + fun sdef con n arg = (sel_of arg^"_def",%%(sel_of arg) == + mk_cfapp(%%(dname^"_when"),map + (fn (con',args) => if con'<>con then %%"UU" else + foldr /\# (args,Bound (length args - n))) cons)); + in flat(map (fn (con,args) => mapn (sdef con) 1 args) cons) end; (* ----- axiom and definitions concerning induction ------------------------------- *) fun cproj' T = cproj T eqs n; val ax_reach = (dname^"_reach" , mk_trp(cproj'(%%"fix"`%%(comp_dname^"_copy")) - `%x_name === %x_name)); + `%x_name === %x_name)); val ax_take_def = (dname^"_take_def",%%(dname^"_take") == mk_lam("n", - cproj'(%%"iterate" $ Bound 0 $ %%(comp_dname^"_copy") $ %%"UU"))); + cproj'(%%"iterate" $ Bound 0 $ %%(comp_dname^"_copy") $ %%"UU"))); val ax_finite_def = (dname^"_finite_def",%%(dname^"_finite") == mk_lam(x_name, - mk_ex("n",(%%(dname^"_take") $ Bound 0)`Bound 1 === Bound 1))); + mk_ex("n",(%%(dname^"_take") $ Bound 0)`Bound 1 === Bound 1))); in [ax_abs_iso, ax_rep_iso, ax_when_def, ax_copy_def] @ axs_con_def @ axs_dis_def @ axs_sel_def @ @@ -105,56 +105,56 @@ val x_name = idx_name dnames "x"; fun copy_app dname = %%(dname^"_copy")`Bound 0; val ax_copy_def = (comp_dname^"_copy_def" , %%(comp_dname^"_copy") == - /\"f"(foldr' cpair (map copy_app dnames))); + /\"f"(foldr' cpair (map copy_app dnames))); val ax_bisim_def = (comp_dname^"_bisim_def",%%(comp_dname^"_bisim") == mk_lam("R", let fun one_con (con,args) = let - val nonrec_args = filter_out is_rec args; - val rec_args = filter is_rec args; - val nonrecs_cnt = length nonrec_args; - val recs_cnt = length rec_args; - val allargs = nonrec_args @ rec_args - @ map (upd_vname (fn s=> s^"'")) rec_args; - val allvns = map vname allargs; - fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg; - val vns1 = map (vname_arg "" ) args; - val vns2 = map (vname_arg "'") args; - val allargs_cnt = nonrecs_cnt + 2*recs_cnt; - val rec_idxs = (recs_cnt-1) downto 0; - val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg) - (allargs~~((allargs_cnt-1) downto 0))); - fun rel_app i ra = proj (Bound(allargs_cnt+2)) dnames (rec_of ra) $ - Bound (2*recs_cnt-i) $ Bound (recs_cnt-i); - val capps = foldr mk_conj (mapn rel_app 1 rec_args, - mk_conj(Bound(allargs_cnt+1)===mk_cfapp(%%con,map (bound_arg allvns) vns1), - Bound(allargs_cnt+0)===mk_cfapp(%%con,map (bound_arg allvns) vns2))); + val nonrec_args = filter_out is_rec args; + val rec_args = filter is_rec args; + val nonrecs_cnt = length nonrec_args; + val recs_cnt = length rec_args; + val allargs = nonrec_args @ rec_args + @ map (upd_vname (fn s=> s^"'")) rec_args; + val allvns = map vname allargs; + fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg; + val vns1 = map (vname_arg "" ) args; + val vns2 = map (vname_arg "'") args; + val allargs_cnt = nonrecs_cnt + 2*recs_cnt; + val rec_idxs = (recs_cnt-1) downto 0; + val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg) + (allargs~~((allargs_cnt-1) downto 0))); + fun rel_app i ra = proj (Bound(allargs_cnt+2)) dnames (rec_of ra) $ + Bound (2*recs_cnt-i) $ Bound (recs_cnt-i); + val capps = foldr mk_conj (mapn rel_app 1 rec_args, + mk_conj(Bound(allargs_cnt+1)===mk_cfapp(%%con,map (bound_arg allvns) vns1), + Bound(allargs_cnt+0)===mk_cfapp(%%con,map (bound_arg allvns) vns2))); in foldr mk_ex (allvns, foldr mk_conj - (map (defined o Bound) nonlazy_idxs,capps)) end; + (map (defined o Bound) nonlazy_idxs,capps)) end; fun one_comp n (_,cons) = mk_all(x_name (n+1), mk_all(x_name (n+1)^"'", mk_imp( - proj (Bound 2) dnames n $ Bound 1 $ Bound 0, + proj (Bound 2) dnames n $ Bound 1 $ Bound 0, foldr' mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU)::map one_con cons)))); in foldr' mk_conj (mapn one_comp 0 eqs)end )); val thy_axs = flat (mapn (calc_axioms comp_dname eqs) 0 eqs) @ - (if length eqs>1 then [ax_copy_def] else []) @ [ax_bisim_def]; + (if length eqs>1 then [ax_copy_def] else []) @ [ax_bisim_def]; in thy' |> add_axioms_i (infer_types thy' thy_axs) end; fun add_gen_by ((tname,finite),(typs,cnstrs)) thy' = let fun pred_name typ ="P"^(if typs=[typ] then "" else string_of_int(1+find(typ,typs))); fun lift_adm t = lift (fn typ => %%"adm" $ %(pred_name typ)) - (if finite then [] else typs,t); + (if finite then [] else typs,t); fun lift_pred_UU t = lift (fn typ => %(pred_name typ) $ UU) (typs,t); fun one_cnstr (cnstr,vns,(args,res)) = let - val rec_args = filter (fn (_,typ) => typ mem typs)(vns~~args); - val app = mk_cfapp(%%cnstr,map (bound_arg vns) vns); - in foldr mk_All (vns, - lift (fn (vn,typ) => %(pred_name typ) $ bound_arg vns vn) - (rec_args,defined app ==> %(pred_name res)$app)) end; + val rec_args = filter (fn (_,typ) => typ mem typs)(vns~~args); + val app = mk_cfapp(%%cnstr,map (bound_arg vns) vns); + in foldr mk_All (vns, + lift (fn (vn,typ) => %(pred_name typ) $ bound_arg vns vn) + (rec_args,defined app ==> %(pred_name res)$app)) end; fun one_conc typ = let val pn = pred_name typ in - %pn $ %("x"^implode(tl(explode pn))) end; + %pn $ %("x"^implode(tl(explode pn))) end; val concl = mk_trp(foldr' mk_conj (map one_conc typs)); val induct =(tname^"_induct",lift_adm(lift_pred_UU( - foldr (op ===>) (map one_cnstr cnstrs,concl)))); + foldr (op ===>) (map one_cnstr cnstrs,concl)))); in thy' |> add_axioms_i (infer_types thy' [induct]) end; end; (* local *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/domain/extender.ML --- a/src/HOLCF/domain/extender.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/domain/extender.ML Tue Jan 30 13:42:57 1996 +0100 @@ -20,78 +20,78 @@ (* ----- general testing and preprocessing of constructor list -------------------- *) fun check_domain(eqs':((string * typ list) * - (string * ThyOps.cmixfix * (bool*string*typ) list) list) list) = let + (string * ThyOps.cmixfix * (bool*string*typ) list) list) list) = let val dtnvs = map fst eqs'; val cons' = flat (map snd eqs'); val test_dupl_typs = (case duplicates (map fst dtnvs) of - [] => false | dups => error ("Duplicate types: " ^ commas_quote dups)); + [] => false | dups => error ("Duplicate types: " ^ commas_quote dups)); val test_dupl_cons = (case duplicates (map first cons') of - [] => false | dups => error ("Duplicate constructors: " ^ commas_quote dups)); + [] => false | dups => error ("Duplicate constructors: " ^ commas_quote dups)); val test_dupl_sels = (case duplicates(map second (flat(map third cons'))) of [] => false | dups => error ("Duplicate selectors: "^commas_quote dups)); val test_dupl_tvars = let fun vname (TFree(v,_)) = v - | vname _ = Imposs "extender:vname"; - in exists (fn tvars => case duplicates (map vname tvars) of - [] => false | dups => error ("Duplicate type arguments: " ^commas_quote dups)) - (map snd dtnvs) end; + | vname _ = Imposs "extender:vname"; + in exists (fn tvars => case duplicates (map vname tvars) of + [] => false | dups => error ("Duplicate type arguments: " ^commas_quote dups)) + (map snd dtnvs) end; (*test for free type variables and invalid use of recursive type*) val analyse_types = forall (fn ((_,typevars),cons') => - forall (fn con' => let - val types = map third (third con'); + forall (fn con' => let + val types = map third (third con'); fun analyse(t as TFree(v,_)) = t mem typevars orelse - error ("Free type variable " ^ v ^ " on rhs.") - | analyse(Type(s,typl)) = (case assoc (dtnvs,s) of None => analyses typl - | Some tvs => tvs = typl orelse - error ("Recursion of type " ^ s ^ " with different arguments")) - | analyse(TVar _) = Imposs "extender:analyse" - and analyses ts = forall analyse ts; - in analyses types end) cons' - ) eqs'; + error ("Free type variable " ^ v ^ " on rhs.") + | analyse(Type(s,typl)) = (case assoc (dtnvs,s) of None => analyses typl + | Some tvs => tvs = typl orelse + error ("Recursion of type " ^ s ^ " with different arguments")) + | analyse(TVar _) = Imposs "extender:analyse" + and analyses ts = forall analyse ts; + in analyses types end) cons' + ) eqs'; in true end; (* let *) fun check_gen_by thy' (typs': string list,cnstrss': string list list) = let val test_dupl_typs = (case duplicates typs' of [] => false - | dups => error ("Duplicate types: " ^ commas_quote dups)); + | dups => error ("Duplicate types: " ^ commas_quote dups)); val test_dupl_cnstrs = map (fn cnstrs' => (case duplicates cnstrs' of [] => false - | dups => error ("Duplicate constructors: " ^ commas_quote dups))) cnstrss'; + | dups => error ("Duplicate constructors: " ^ commas_quote dups))) cnstrss'; val tsig = #tsig(Sign.rep_sg(sign_of thy')); val tycons = map fst (#tycons(Type.rep_tsig tsig)); val test_types = forall(fn t=>t mem tycons orelse error("Unknown type: "^t))typs'; val cnstrss = let - fun type_of c = case (Sign.const_type(sign_of thy') c) of Some t => t - | None => error ("Unknown constructor: "^c); - fun args_result_type (t as (Type(tn,[arg,rest]))) = - if tn = "->" orelse tn = "=>" - then let val (ts,r) = args_result_type rest in (arg::ts,r) end - else ([],t) - | args_result_type t = ([],t); + fun type_of c = case (Sign.const_type(sign_of thy') c) of Some t => t + | None => error ("Unknown constructor: "^c); + fun args_result_type (t as (Type(tn,[arg,rest]))) = + if tn = "->" orelse tn = "=>" + then let val (ts,r) = args_result_type rest in (arg::ts,r) end + else ([],t) + | args_result_type t = ([],t); in map (map (fn cn => let val (args,res) = args_result_type (type_of cn) in - (cn,mk_var_names args,(args,res)) end)) cnstrss' - : (string * (* operator name of constr *) - string list * (* argument name list *) - (typ list * (* argument types *) - typ)) (* result type *) - list list end; + (cn,mk_var_names args,(args,res)) end)) cnstrss' + : (string * (* operator name of constr *) + string list * (* argument name list *) + (typ list * (* argument types *) + typ)) (* result type *) + list list end; fun test_equal_type tn (cn,_,(_,rt)) = fst (type_name_vars rt) = tn orelse - error("Inappropriate result type for constructor "^cn); + error("Inappropriate result type for constructor "^cn); val typs = map (fn (tn, cnstrs) => - (map (test_equal_type tn) cnstrs; snd(third(hd(cnstrs))))) - (typs'~~cnstrss); + (map (test_equal_type tn) cnstrs; snd(third(hd(cnstrs))))) + (typs'~~cnstrss); val test_typs = map (fn (typ,cnstrs) => - if not (Type.typ_instance(tsig,typ,TVar(("'a",0),["pcpo"]))) - then error("Not a pcpo type: "^fst(type_name_vars typ)) - else map (fn (cn,_,(_,rt)) => rt=typ - orelse error("Non-identical result types for constructors "^ - first(hd cnstrs)^" and "^ cn )) cnstrs) (typs~~cnstrss); + if not (Type.typ_instance(tsig,typ,TVar(("'a",0),["pcpo"]))) + then error("Not a pcpo type: "^fst(type_name_vars typ)) + else map (fn (cn,_,(_,rt)) => rt=typ + orelse error("Non-identical result types for constructors "^ + first(hd cnstrs)^" and "^ cn )) cnstrs) (typs~~cnstrss); val proper_args = let - fun occurs tn (Type(tn',ts)) = (tn'=tn) orelse exists (occurs tn) ts - | occurs _ _ = false; - fun proper_arg cn atyp = forall (fn typ => let - val tn = fst(type_name_vars typ) - in atyp=typ orelse not (occurs tn atyp) orelse - error("Illegal use of type "^ tn ^ - " as argument of constructor " ^cn)end )typs; - fun proper_curry (cn,_,(args,_)) = forall (proper_arg cn) args; + fun occurs tn (Type(tn',ts)) = (tn'=tn) orelse exists (occurs tn) ts + | occurs _ _ = false; + fun proper_arg cn atyp = forall (fn typ => let + val tn = fst(type_name_vars typ) + in atyp=typ orelse not (occurs tn atyp) orelse + error("Illegal use of type "^ tn ^ + " as argument of constructor " ^cn)end )typs; + fun proper_curry (cn,_,(args,_)) = forall (proper_arg cn) args; in map (map proper_curry) cnstrss end; in (typs, flat cnstrss) end; @@ -104,12 +104,12 @@ val thy' = thy'' |> Domain_Syntax.add_syntax (comp_dname,eqs'); val dts = map (Type o fst) eqs'; fun cons cons' = (map (fn (con,syn,args) => - (ThyOps.const_name con syn, - map (fn ((lazy,sel,tp),vn) => ((lazy, - find (tp,dts) handle LIST "find" => ~1), - sel,vn)) - (args~~(mk_var_names(map third args))) - )) cons') : cons list; + (ThyOps.const_name con syn, + map (fn ((lazy,sel,tp),vn) => ((lazy, + find (tp,dts) handle LIST "find" => ~1), + sel,vn)) + (args~~(mk_var_names(map third args))) + )) cons') : cons list; val eqs = map (fn (dtnvs,cons') => (dtnvs,cons cons')) eqs' : eq list; val thy = thy' |> Domain_Axioms.add_axioms (comp_dname,eqs); in (thy,eqs) end; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/domain/interface.ML --- a/src/HOLCF/domain/interface.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/domain/interface.ML Tue Jan 30 13:42:57 1996 +0100 @@ -26,24 +26,24 @@ fun gen_val dname name = "val "^name^" = " ^gt_ax (dname^"_"^name)^";"; fun gen_vall name l = "val "^name^" = " ^mk_list l^";"; val rews1 = "iso_rews @ when_rews @\n\ - \con_rews @ sel_rews @ dis_rews @ dists_eq @ dists_le @\n\ - \copy_rews"; + \con_rews @ sel_rews @ dis_rews @ dists_eq @ dists_le @\n\ + \copy_rews"; fun gen_domain eqname num ((dname,_), cons') = let val axioms1 = ["abs_iso", "rep_iso", "when_def"]; - (* con_defs , sel_defs, dis_defs *) + (* con_defs , sel_defs, dis_defs *) val axioms2 = ["copy_def"]; val theorems = - "iso_rews, exhaust, cases, when_rews, con_rews, sel_rews, dis_rews,\n \ - \dists_eq, dists_le, inverts, injects, copy_rews"; + "iso_rews, exhaust, cases, when_rews, con_rews, sel_rews, dis_rews,\n \ + \dists_eq, dists_le, inverts, injects, copy_rews"; in "structure "^dname^" = struct\n"^ cat_lines(map (gen_val dname) axioms1)^"\n"^ gen_vall"con_defs"(map(fn (con,_,_) => gt_ax(strip_esc con^"_def")) cons')^"\n"^ gen_vall"sel_defs"(flat(map (fn (_,_,args) => map (fn (_,sel,_) => - gt_ax(sel^"_def")) args) cons'))^"\n"^ + gt_ax(sel^"_def")) args) cons'))^"\n"^ gen_vall"dis_defs"(map(fn (con,_,_) => gt_ax(dis_name_ con^"_def")) - cons')^"\n"^ + cons')^"\n"^ cat_lines(map (gen_val dname) axioms2)^"\n"^ "val ("^ theorems ^") =\n\ \Domain_Theorems.theorems thy "^eqname^";\n" ^ @@ -57,10 +57,10 @@ val comp_dname = implode (separate "_" dnames); val conss' = map (fn (dname,cons'') => let - fun sel n m = upd_second (fn None => dname^"_sel_"^(string_of_int n)^ - "_"^(string_of_int m) - | Some s => s) - fun fill_sels n con = upd_third (mapn (sel n) 1) con; + fun sel n m = upd_second (fn None => dname^"_sel_"^(string_of_int n)^ + "_"^(string_of_int m) + | Some s => s) + fun fill_sels n con = upd_third (mapn (sel n) 1) con; in mapn fill_sels 1 cons'' end) (dnames~~(map snd eqs'')); val eqs' = dtnvs~~conss'; @@ -69,50 +69,50 @@ fun string_of_sort_emb [] = "" | string_of_sort_emb [x] = "\"" ^x^ "\"" - | string_of_sort_emb (x::xs) = "\"" ^x^ "\"" ^ ", "^(string_of_sort_emb xs); + | string_of_sort_emb (x::xs) = "\"" ^x^ "\"" ^ ", "^(string_of_sort_emb xs); - fun string_of_sort l = "["^ (string_of_sort_emb l)^"]"; + fun string_of_sort l = "["^ (string_of_sort_emb l)^"]"; fun mk_tnv (n,v) = mk_pair(quote n,mk_list(map mk_typ v)) and mk_typ (TFree(name,sort)) = "TFree"^mk_pair(quote name,string_of_sort sort) | mk_typ (Type (name,args)) = "Type" ^mk_tnv(name,args) | mk_typ _ = Imposs "interface:mk_typ"; fun mk_conslist cons' = mk_list (map - (fn (c,syn,ts)=>mk_triple(quote c,syn,mk_list + (fn (c,syn,ts)=>mk_triple(quote c,syn,mk_list (map (fn (b,s ,tp) =>mk_triple(makestring(b:bool),quote s, - mk_typ tp)) ts))) cons'); + mk_typ tp)) ts))) cons'); in ("val (thy, "^comp_dname^"_equations) = thy |> Extender.add_domain \n" ^ mk_pair(quote comp_dname, - mk_list(map (fn (t,cs)=> mk_pair (mk_tnv t,mk_conslist cs)) eqs')) + mk_list(map (fn (t,cs)=> mk_pair (mk_tnv t,mk_conslist cs)) eqs')) ^ ";\nval thy = thy", let - fun plural s = if num > 1 then s^"s" else "["^s^"]"; - val comp_axioms = [(* copy, *) "take_def", "finite_def", "reach" - (*, "bisim_def" *)]; - val comp_theorems = "take_rews, " ^ implode (separate ", " (map plural - ["take_lemma","finite"]))^", finite_ind, ind, coind"; - fun eqname n = "(hd(funpow "^string_of_int n^" tl "^comp_dname^"_equations), " - ^comp_dname^"_equations)"; - fun collect sep name = if num = 1 then name else - implode (separate sep (map (fn s => s^"."^name) dnames)); + fun plural s = if num > 1 then s^"s" else "["^s^"]"; + val comp_axioms = [(* copy, *) "take_def", "finite_def", "reach" + (*, "bisim_def" *)]; + val comp_theorems = "take_rews, " ^ implode (separate ", " (map plural + ["take_lemma","finite"]))^", finite_ind, ind, coind"; + fun eqname n = "(hd(funpow "^string_of_int n^" tl "^comp_dname^"_equations), " + ^comp_dname^"_equations)"; + fun collect sep name = if num = 1 then name else + implode (separate sep (map (fn s => s^"."^name) dnames)); in - implode (separate "end; (* struct *)\n\n" - (mapn (fn n => gen_domain (eqname n) num) 0 eqs'))^(if num > 1 then - "end; (* struct *)\n\n\ - \structure "^comp_dname^" = struct\n" else "") ^ - (if num > 1 then gen_val comp_dname "copy_def" ^"\n" else "") ^ - implode ((map (fn s => gen_vall (plural s) - (map (fn dn => gt_ax(dn ^ "_" ^ s)) dnames) ^"\n") comp_axioms)) ^ - gen_val comp_dname "bisim_def" ^"\n\ + implode (separate "end; (* struct *)\n\n" + (mapn (fn n => gen_domain (eqname n) num) 0 eqs'))^(if num > 1 then + "end; (* struct *)\n\n\ + \structure "^comp_dname^" = struct\n" else "") ^ + (if num > 1 then gen_val comp_dname "copy_def" ^"\n" else "") ^ + implode ((map (fn s => gen_vall (plural s) + (map (fn dn => gt_ax(dn ^ "_" ^ s)) dnames) ^"\n") comp_axioms)) ^ + gen_val comp_dname "bisim_def" ^"\n\ \val ("^ comp_theorems ^") =\n\ - \Domain_Theorems.comp_theorems thy \ - \(" ^ quote comp_dname ^ ","^ comp_dname ^"_equations,\n\ - \ ["^collect "," "cases" ^"],\n\ - \ "^ collect "@" "con_rews " ^",\n\ - \ "^ collect "@" "copy_rews" ^");\n\ - \val rews = "^(if num>1 then collect" @ " "rews"else rews1)^ " @ take_rews;\n\ - \end; (* struct *)" + \Domain_Theorems.comp_theorems thy \ + \(" ^ quote comp_dname ^ ","^ comp_dname ^"_equations,\n\ + \ ["^collect "," "cases" ^"],\n\ + \ "^ collect "@" "con_rews " ^",\n\ + \ "^ collect "@" "copy_rews" ^");\n\ + \val rews = "^(if num>1 then collect" @ " "rews"else rews1)^ " @ take_rews;\n\ + \end; (* struct *)" end ) end; @@ -122,47 +122,47 @@ val tname = implode (separate "_" typs) in ("|> Extender.add_gen_by " ^ mk_pair(mk_pair(quote tname,makestring (finite:bool)), - mk_pair(mk_list(map quote typs), - mk_list (map (fn cns => mk_list(map quote cns)) cnstrss))), + mk_pair(mk_list(map quote typs), + mk_list (map (fn cns => mk_list(map quote cns)) cnstrss))), "val "^tname^"_induct = " ^gt_ax (tname^"_induct")^";") end; (* ----- parser for domain declaration equation ----------------------------------- *) (** val sort = name >> (fn s => [strip_quotes s]) - || "{" $$-- !! (list (name >> strip_quotes) --$$ "}"); + || "{" $$-- !! (list (name >> strip_quotes) --$$ "}"); val tvar = (type_var -- (optional ("::" $$-- !! sort) ["pcpo"])) >> TFree **) val tvar = type_var >> (fn tv => TFree(strip_quotes tv,["pcpo"])); val type_args = "(" $$-- !! (list1 tvar --$$ ")") - || tvar >> (fn x => [x]) - || empty >> K []; + || tvar >> (fn x => [x]) + || empty >> K []; val con_typ = type_args -- ident >> (fn (x,y) => Type(y,x)); val typ = con_typ - || tvar; + || tvar; val domain_arg = "(" $$-- (optional ($$ "lazy" >> K true) false) - -- (optional ((ident >> Some) --$$ "::") None) - -- typ --$$ ")" >> (fn ((lazy,sel),tp) => (lazy,sel,tp)) - || ident >> (fn x => (false,None,Type(x,[]))) - || tvar >> (fn x => (false,None,x)); + -- (optional ((ident >> Some) --$$ "::") None) + -- typ --$$ ")" >> (fn ((lazy,sel),tp) => (lazy,sel,tp)) + || ident >> (fn x => (false,None,Type(x,[]))) + || tvar >> (fn x => (false,None,x)); val domain_cons = (name >> strip_quotes) -- !! (repeat domain_arg) - -- ThyOps.opt_cmixfix - >> (fn ((con,args),syn) => (con,syn,args)); + -- ThyOps.opt_cmixfix + >> (fn ((con,args),syn) => (con,syn,args)); in val domain_decl = (enum1 "," (con_typ --$$ "=" -- !! - (enum1 "|" domain_cons))) >> mk_domain; + (enum1 "|" domain_cons))) >> mk_domain; val gen_by_decl = (optional ($$ "finite" >> K true) false) -- - (enum1 "," (ident --$$ "by" -- !! - (enum1 "|" (name >> strip_quotes)))) >> mk_gen_by; + (enum1 "," (ident --$$ "by" -- !! + (enum1 "|" (name >> strip_quotes)))) >> mk_gen_by; end; val user_keywords = "lazy"::"by"::"finite":: - (**)filter_out (fn s=>s="lazy" orelse s="by" orelse s="finite")(**) - ThySynData.user_keywords; + (**)filter_out (fn s=>s="lazy" orelse s="by" orelse s="finite")(**) + ThySynData.user_keywords; val user_sections = ("domain", domain_decl)::("generated", gen_by_decl):: - (**)filter_out (fn (s,_)=>s="domain" orelse s="generated")(**) - ThySynData.user_sections; + (**)filter_out (fn (s,_)=>s="domain" orelse s="generated")(**) + ThySynData.user_sections; end; in diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/domain/library.ML --- a/src/HOLCF/domain/library.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/domain/library.ML Tue Jan 30 13:42:57 1996 +0100 @@ -20,7 +20,7 @@ in itr l end; fun foldr' f l = foldr'' f (l,Id); fun map_cumulr f start xs = foldr (fn (x,(ys,res)) => case f(x,res) of (y,res2) => - (y::ys,res2)) (xs,([],start)); + (y::ys,res2)) (xs,([],start)); fun first (x,_,_) = x; fun second (_,x,_) = x; fun third (_,_,x) = x; @@ -37,10 +37,10 @@ fun atomize thm = case concl_of thm of _ $ (Const("op &",_) $ _ $ _) => atomize (thm RS conjunct1) @ - atomize (thm RS conjunct2) + atomize (thm RS conjunct2) | _ $ (Const("All" ,_) $ Abs (s,_,_)) => atomize (thm RS - (read_instantiate [("x","?"^s)] spec)) - | _ => [thm]; + (read_instantiate [("x","?"^s)] spec)) + | _ => [thm]; (* ----- specific support for domain ---------------------------------------------- *) @@ -59,8 +59,8 @@ in implode o strip o explode end; fun extern_name con = case explode con of - ("o"::"p"::" "::rest) => implode rest - | _ => con; + ("o"::"p"::" "::rest) => implode rest + | _ => con; fun dis_name con = "is_"^ (extern_name con); fun dis_name_ con = "is_"^ (strip_esc con); @@ -72,8 +72,8 @@ | typid (TFree (id,_) ) = hd (tl (explode id)) | typid (TVar ((id,_),_)) = hd (tl (explode id)); fun nonreserved id = let val cs = explode id in - if not(hd cs mem ["x","f","P"]) then id - else implode(chr(1+ord (hd cs))::tl cs) end; + if not(hd cs mem ["x","f","P"]) then id + else implode(chr(1+ord (hd cs))::tl cs) end; fun index_vnames(vn::vns,tab) = (case assoc(tab,vn) of None => if vn mem vns @@ -96,14 +96,14 @@ (* ----- constructor list handling ----- *) -type cons = (string * (* operator name of constr *) - ((bool*int)* (* (lazy,recursive element or ~1) *) - string* (* selector name *) - string) (* argument name *) - list); (* argument list *) -type eq = (string * (* name of abstracted type *) - typ list) * (* arguments of abstracted type *) - cons list; (* represented type, as a constructor list *) +type cons = (string * (* operator name of constr *) + ((bool*int)* (* (lazy,recursive element or ~1) *) + string* (* selector name *) + string) (* argument name *) + list); (* argument list *) +type eq = (string * (* name of abstracted type *) + typ list) * (* arguments of abstracted type *) + cons list; (* represented type, as a constructor list *) val rec_of = snd o first; val is_lazy = fst o first; @@ -114,7 +114,7 @@ fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); fun nonlazy args = map vname (filter_out is_lazy args); fun is_one_con_one_arg p cons = length cons = 1 andalso let val args = snd(hd cons) in - length args = 1 andalso p (hd args) end; + length args = 1 andalso p (hd args) end; (* ----- support for term expressions ----- *) @@ -130,14 +130,14 @@ fun mk_lam (x,T) = Abs(x,dummyT,T); fun mk_all (x,P) = HOLogic.mk_all (x,dummyT,P); local - fun tf (Type (s,args)) = foldl (op $) (%s,map tf args) - | tf (TFree(s,_ )) = %s - | tf _ = Imposs "mk_constrainall"; + fun tf (Type (s,args)) = foldl (op $) (%s,map tf args) + | tf (TFree(s,_ )) = %s + | tf _ = Imposs "mk_constrainall"; in fun mk_constrain (typ,T) = %%"_constrain" $ T $ tf typ; fun mk_constrainall (x,typ,P) = %%"All" $ (%%"_constrainAbs"$Abs(x,dummyT,P)$tf typ); end; - + fun mk_ex (x,P) = mk_exists (x,dummyT,P); fun mk_not P = Const("not" ,boolT --> boolT) $ P; end; @@ -177,26 +177,26 @@ fun cont_eta_contract (Const("fabs",TT) $ Abs(a,T,body)) = (case cont_eta_contract body of body' as (Const("fapp",Ta) $ f $ Bound 0) => - if not (0 mem loose_bnos f) then incr_boundvars ~1 f - else Const("fabs",TT) $ Abs(a,T,body') + if not (0 mem loose_bnos f) then incr_boundvars ~1 f + else Const("fabs",TT) $ Abs(a,T,body') | body' => Const("fabs",TT) $ Abs(a,T,body')) | cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t | cont_eta_contract t = t; fun idx_name dnames s n = s ^ (if length dnames = 1 then "" else string_of_int n); fun when_funs cons = if length cons = 1 then ["f"] - else mapn (fn n => K("f"^(string_of_int n))) 1 cons; + else mapn (fn n => K("f"^(string_of_int n))) 1 cons; fun when_body cons funarg = let - fun one_fun n (_,[] ) = /\ "dummy" (funarg(1,n)) - | one_fun n (_,args) = let - val l2 = length args; - fun idxs m arg = (if is_lazy arg then fn x=> %%"lift"`%%"ID"`x - else Id) (Bound(l2-m)); - in cont_eta_contract (foldr'' - (fn (a,t) => %%"ssplit"`(/\# (a,t))) - (args, - fn a => /\# (a,(mk_cfapp(funarg (l2,n),mapn idxs 1 args)))) - ) end; + fun one_fun n (_,[] ) = /\ "dummy" (funarg(1,n)) + | one_fun n (_,args) = let + val l2 = length args; + fun idxs m arg = (if is_lazy arg then fn x=> %%"lift"`%%"ID"`x + else Id) (Bound(l2-m)); + in cont_eta_contract (foldr'' + (fn (a,t) => %%"ssplit"`(/\# (a,t))) + (args, + fn a => /\# (a,(mk_cfapp(funarg (l2,n),mapn idxs 1 args)))) + ) end; in foldr' (fn (x,y)=> %%"sswhen"`x`y) (mapn one_fun 1 cons) end; end; (* struct *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/domain/syntax.ML --- a/src/HOLCF/domain/syntax.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/domain/syntax.ML Tue Jan 30 13:42:57 1996 +0100 @@ -15,14 +15,14 @@ open Domain_Library; infixr 5 -->; infixr 6 ~>; fun calc_syntax dtypeprod ((dname,typevars), - (cons':(string*ThyOps.cmixfix*(bool*string*typ) list) list))= + (cons':(string*ThyOps.cmixfix*(bool*string*typ) list) list))= let (* ----- constants concerning the isomorphism ------------------------------------- *) local fun opt_lazy (lazy,_,t) = if lazy then Type("u",[t]) else t fun prod (_,_,args) = if args = [] then Type("one",[]) - else foldr' (mk_typ "**") (map opt_lazy args); + else foldr' (mk_typ "**") (map opt_lazy args); in val dtype = Type(dname,typevars); @@ -42,20 +42,20 @@ local val escape = let - fun esc (c :: cs) = if c mem ["'","_","(",")","/"] then "'" :: c :: esc cs - else c :: esc cs - | esc [] = [] - in implode o esc o explode end; + fun esc (c :: cs) = if c mem ["'","_","(",")","/"] then "'" :: c :: esc cs + else c :: esc cs + | esc [] = [] + in implode o esc o explode end; fun freetvar s = if (mk_tvar s) mem typevars then freetvar ("t"^s) else mk_tvar s; fun when_type (_ ,_,args) = foldr (op ~>) (map third args,freetvar "t"); fun con (name,s,args) = (name,foldr (op ~>) (map third args,dtype),s); fun dis (con ,s,_ ) = (dis_name_ con, dtype~>Type("tr",[]), - ThyOps.Mixfix(Mixfix("is'_"^ - (if is_infix s then Id else escape)con,[],max_pri))); + ThyOps.Mixfix(Mixfix("is'_"^ + (if is_infix s then Id else escape)con,[],max_pri))); fun sel (_ ,_,args) = map (fn(_,sel,typ)=>(sel,dtype ~> typ,NoSyn'))args; in val const_when = (dname^"_when", foldr (op ~>) ((map when_type cons'), - dtype ~> freetvar "t"), NoSyn'); + dtype ~> freetvar "t"), NoSyn'); val consts_con = map con cons'; val consts_dis = map dis cons'; val consts_sel = flat(map sel cons'); @@ -64,31 +64,31 @@ (* ----- constants concerning induction ------------------------------------------- *) val const_take = (dname^"_take" ,Type("nat",[]) --> dtype ~> dtype ,NoSyn'); - val const_finite = (dname^"_finite",dtype-->HOLogic.boolT ,NoSyn'); + val const_finite = (dname^"_finite",dtype-->HOLogic.boolT ,NoSyn'); (* ----- case translation --------------------------------------------------------- *) local open Syntax in val case_trans = let - fun c_ast con syn = Constant (ThyOps.const_name con syn); - fun expvar n = Variable ("e"^(string_of_int n)); - fun argvar n m _ = Variable ("a"^(string_of_int n)^"_"^(string_of_int m)); - fun app s (l,r) = mk_appl (Constant s) [l,r]; - fun case1 n (con,syn,args) = mk_appl (Constant "@case1") - [if is_infix syn - then mk_appl (c_ast con syn) (mapn (argvar n) 1 args) - else foldl (app "@fapp") (c_ast con syn, (mapn (argvar n) 1 args)), - expvar n]; - fun arg1 n (con,_,args) = if args = [] then expvar n - else mk_appl (Constant "LAM ") - [foldr' (app "_idts") (mapn (argvar n) 1 args) , expvar n]; + fun c_ast con syn = Constant (ThyOps.const_name con syn); + fun expvar n = Variable ("e"^(string_of_int n)); + fun argvar n m _ = Variable ("a"^(string_of_int n)^"_"^(string_of_int m)); + fun app s (l,r) = mk_appl (Constant s) [l,r]; + fun case1 n (con,syn,args) = mk_appl (Constant "@case1") + [if is_infix syn + then mk_appl (c_ast con syn) (mapn (argvar n) 1 args) + else foldl (app "@fapp") (c_ast con syn, (mapn (argvar n) 1 args)), + expvar n]; + fun arg1 n (con,_,args) = if args = [] then expvar n + else mk_appl (Constant "LAM ") + [foldr' (app "_idts") (mapn (argvar n) 1 args) , expvar n]; in mk_appl (Constant "@case") [Variable "x", foldr' - (fn (c,cs) => mk_appl (Constant "@case2") [c,cs]) - (mapn case1 1 cons')] <-> + (fn (c,cs) => mk_appl (Constant "@case2") [c,cs]) + (mapn case1 1 cons')] <-> mk_appl (Constant "@fapp") [foldl - (fn (w,a ) => mk_appl (Constant "@fapp" ) [w,a ]) - (Constant (dname^"_when"),mapn arg1 1 cons'), - Variable "x"] + (fn (w,a ) => mk_appl (Constant "@fapp" ) [w,a ]) + (Constant (dname^"_when"),mapn arg1 1 cons'), + Variable "x"] end; end; @@ -103,7 +103,7 @@ in (* local *) fun add_syntax (comp_dname,eqs': ((string * typ list) * - (string * ThyOps.cmixfix * (bool*string*typ) list) list) list) thy'' = + (string * ThyOps.cmixfix * (bool*string*typ) list) list) list) thy'' = let fun thy_type (dname,typevars) = (dname, length typevars, NoSyn); fun thy_arity (dname,typevars) = (dname, map (K ["pcpo"]) typevars, ["pcpo"]); @@ -118,10 +118,10 @@ val ctt = map (calc_syntax funprod) eqs'; val add_cur_ops_i = ThyOps.add_ops_i {curried=true, strict=false, total=false}; in thy'' |> add_types thy_types - |> add_arities thy_arities - |> add_cur_ops_i (flat(map fst ctt)) - |> add_cur_ops_i ((if length eqs'>1 then [const_copy] else[])@[const_bisim]) - |> add_trrules_i (flat(map snd ctt)) + |> add_arities thy_arities + |> add_cur_ops_i (flat(map fst ctt)) + |> add_cur_ops_i ((if length eqs'>1 then [const_copy] else[])@[const_bisim]) + |> add_trrules_i (flat(map snd ctt)) end; (* let *) end; (* local *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/domain/theorems.ML --- a/src/HOLCF/domain/theorems.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/domain/theorems.ML Tue Jan 30 13:42:57 1996 +0100 @@ -34,31 +34,31 @@ val b=0; fun _ y t = by t; fun g defs t = let val sg = sign_of thy; - val ct = Thm.cterm_of sg (inferT sg t); - in goalw_cterm defs ct end; + val ct = Thm.cterm_of sg (inferT sg t); + in goalw_cterm defs ct end; *) fun pg'' thy defs t = let val sg = sign_of thy; - val ct = Thm.cterm_of sg (inferT sg t); - in prove_goalw_cterm defs ct end; + val ct = Thm.cterm_of sg (inferT sg t); + in prove_goalw_cterm defs ct end; fun pg' thy defs t tacsf=pg'' thy defs t (fn [] => tacsf - | prems=> (cut_facts_tac prems 1)::tacsf); + | prems=> (cut_facts_tac prems 1)::tacsf); fun REPEAT_DETERM_UNTIL p tac = let fun drep st = if p st then Sequence.single st - else (case Sequence.pull(tapply(tac,st)) of - None => Sequence.null - | Some(st',_) => drep st') + else (case Sequence.pull(tapply(tac,st)) of + None => Sequence.null + | Some(st',_) => drep st') in Tactic drep end; val UNTIL_SOLVED = REPEAT_DETERM_UNTIL (has_fewer_prems 1); local val trueI2 = prove_goal HOL.thy "f~=x ==> True" (fn prems => [rtac TrueI 1]) in val kill_neq_tac = dtac trueI2 end; -fun case_UU_tac rews i v = res_inst_tac [("Q",v^"=UU")] classical2 i THEN - asm_simp_tac (HOLCF_ss addsimps rews) i; +fun case_UU_tac rews i v = res_inst_tac [("Q",v^"=UU")] classical2 i THEN + asm_simp_tac (HOLCF_ss addsimps rews) i; val chain_tac = REPEAT_DETERM o resolve_tac - [is_chain_iterate, ch2ch_fappR, ch2ch_fappL]; + [is_chain_iterate, ch2ch_fappR, ch2ch_fappL]; (* ----- general proofs ----------------------------------------------------------- *) @@ -66,17 +66,17 @@ cut_facts_tac prems 1, etac swap 1, dtac notnotD 1, - etac (hd prems) 1]); + etac (hd prems) 1]); val dist_eqI = prove_goal Porder0.thy "~ x << y ==> x ~= y" (fn prems => [ - cut_facts_tac prems 1, - etac swap 1, - dtac notnotD 1, - asm_simp_tac HOLCF_ss 1]); + cut_facts_tac prems 1, + etac swap 1, + dtac notnotD 1, + asm_simp_tac HOLCF_ss 1]); val cfst_strict = prove_goal Cprod3.thy "cfst`UU = UU" (fn _ => [ - (simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]); + (simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]); val csnd_strict = prove_goal Cprod3.thy "csnd`UU = UU" (fn _ => [ - (simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]); + (simp_tac (HOLCF_ss addsimps [inst_cprod_pcpo2]) 1)]); in @@ -96,7 +96,7 @@ val axs_con_def = map (fn (con,_) => ga (extern_name con ^"_def")) cons; val axs_dis_def = map (fn (con,_) => ga ( dis_name con ^"_def")) cons; val axs_sel_def = flat(map (fn (_,args) => - map (fn arg => ga (sel_of arg ^"_def")) args) cons); + map (fn arg => ga (sel_of arg ^"_def")) args) cons); val ax_copy_def = ga (dname^"_copy_def" ); end; (* local *) @@ -108,238 +108,238 @@ val x_name = "x"; val (rep_strict, abs_strict) = let - val r = ax_rep_iso RS (ax_abs_iso RS (allI RSN(2,allI RS iso_strict))) - in (r RS conjunct1, r RS conjunct2) end; + val r = ax_rep_iso RS (ax_abs_iso RS (allI RSN(2,allI RS iso_strict))) + in (r RS conjunct1, r RS conjunct2) end; val abs_defin' = pg [] ((dc_abs`%x_name === UU) ==> (%x_name === UU)) [ - res_inst_tac [("t",x_name)] (ax_abs_iso RS subst) 1, - etac ssubst 1, - rtac rep_strict 1]; + res_inst_tac [("t",x_name)] (ax_abs_iso RS subst) 1, + etac ssubst 1, + rtac rep_strict 1]; val rep_defin' = pg [] ((dc_rep`%x_name === UU) ==> (%x_name === UU)) [ - res_inst_tac [("t",x_name)] (ax_rep_iso RS subst) 1, - etac ssubst 1, - rtac abs_strict 1]; + res_inst_tac [("t",x_name)] (ax_rep_iso RS subst) 1, + etac ssubst 1, + rtac abs_strict 1]; val iso_rews = [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict]; local val iso_swap = pg [] (dc_rep`%"x" === %"y" ==> %"x" === dc_abs`%"y") [ - dres_inst_tac [("f",dname^"_abs")] cfun_arg_cong 1, - etac (ax_rep_iso RS subst) 1]; + dres_inst_tac [("f",dname^"_abs")] cfun_arg_cong 1, + etac (ax_rep_iso RS subst) 1]; fun exh foldr1 cn quant foldr2 var = let fun one_con (con,args) = let val vns = map vname args in foldr quant (vns, foldr2 ((%x_name === con_app2 con (var vns) vns):: - map (defined o (var vns)) (nonlazy args))) end + map (defined o (var vns)) (nonlazy args))) end in foldr1 ((cn(%x_name===UU))::map one_con cons) end; in val cases = let - fun common_tac thm = rtac thm 1 THEN contr_tac 1; - fun unit_tac true = common_tac liftE1 - | unit_tac _ = all_tac; - fun prod_tac [] = common_tac oneE - | prod_tac [arg] = unit_tac (is_lazy arg) - | prod_tac (arg::args) = - common_tac sprodE THEN - kill_neq_tac 1 THEN - unit_tac (is_lazy arg) THEN - prod_tac args; - fun sum_one_tac p = SELECT_GOAL(EVERY[ - rtac p 1, - rewrite_goals_tac axs_con_def, - dtac iso_swap 1, - simp_tac HOLCF_ss 1, - UNTIL_SOLVED(fast_tac HOL_cs 1)]) 1; - fun sum_tac [(_,args)] [p] = - prod_tac args THEN sum_one_tac p - | sum_tac ((_,args)::cons') (p::prems) = DETERM( - common_tac ssumE THEN - kill_neq_tac 1 THEN kill_neq_tac 2 THEN - prod_tac args THEN sum_one_tac p) THEN - sum_tac cons' prems - | sum_tac _ _ = Imposs "theorems:sum_tac"; - in pg'' thy [] (exh (fn l => foldr (op ===>) (l,mk_trp(%"P"))) - (fn T => T ==> %"P") mk_All - (fn l => foldr (op ===>) (map mk_trp l,mk_trp(%"P"))) - bound_arg) - (fn prems => [ - cut_facts_tac [excluded_middle] 1, - etac disjE 1, - rtac (hd prems) 2, - etac rep_defin' 2, - if is_one_con_one_arg (not o is_lazy) cons - then rtac (hd (tl prems)) 1 THEN atac 2 THEN - rewrite_goals_tac axs_con_def THEN - simp_tac (HOLCF_ss addsimps [ax_rep_iso]) 1 - else sum_tac cons (tl prems)])end; + fun common_tac thm = rtac thm 1 THEN contr_tac 1; + fun unit_tac true = common_tac liftE1 + | unit_tac _ = all_tac; + fun prod_tac [] = common_tac oneE + | prod_tac [arg] = unit_tac (is_lazy arg) + | prod_tac (arg::args) = + common_tac sprodE THEN + kill_neq_tac 1 THEN + unit_tac (is_lazy arg) THEN + prod_tac args; + fun sum_one_tac p = SELECT_GOAL(EVERY[ + rtac p 1, + rewrite_goals_tac axs_con_def, + dtac iso_swap 1, + simp_tac HOLCF_ss 1, + UNTIL_SOLVED(fast_tac HOL_cs 1)]) 1; + fun sum_tac [(_,args)] [p] = + prod_tac args THEN sum_one_tac p + | sum_tac ((_,args)::cons') (p::prems) = DETERM( + common_tac ssumE THEN + kill_neq_tac 1 THEN kill_neq_tac 2 THEN + prod_tac args THEN sum_one_tac p) THEN + sum_tac cons' prems + | sum_tac _ _ = Imposs "theorems:sum_tac"; + in pg'' thy [] (exh (fn l => foldr (op ===>) (l,mk_trp(%"P"))) + (fn T => T ==> %"P") mk_All + (fn l => foldr (op ===>) (map mk_trp l,mk_trp(%"P"))) + bound_arg) + (fn prems => [ + cut_facts_tac [excluded_middle] 1, + etac disjE 1, + rtac (hd prems) 2, + etac rep_defin' 2, + if is_one_con_one_arg (not o is_lazy) cons + then rtac (hd (tl prems)) 1 THEN atac 2 THEN + rewrite_goals_tac axs_con_def THEN + simp_tac (HOLCF_ss addsimps [ax_rep_iso]) 1 + else sum_tac cons (tl prems)])end; val exhaust = pg [] (mk_trp(exh (foldr' mk_disj) Id mk_ex (foldr' mk_conj) (K %))) [ - rtac cases 1, - UNTIL_SOLVED(fast_tac HOL_cs 1)]; + rtac cases 1, + UNTIL_SOLVED(fast_tac HOL_cs 1)]; end; local val when_app = foldl (op `) (%%(dname^"_when"), map % (when_funs cons)); val when_appl = pg [ax_when_def] (mk_trp(when_app`%x_name===when_body cons - (fn (_,n) => %(nth_elem(n-1,when_funs cons)))`(dc_rep`%x_name))) [ - simp_tac HOLCF_ss 1]; + (fn (_,n) => %(nth_elem(n-1,when_funs cons)))`(dc_rep`%x_name))) [ + simp_tac HOLCF_ss 1]; in val when_strict = pg [] ((if is_one_con_one_arg (K true) cons - then fn t => mk_trp(strict(%"f")) ===> t else Id)(mk_trp(strict when_app))) [ - simp_tac(HOLCF_ss addsimps [when_appl,rep_strict]) 1]; + then fn t => mk_trp(strict(%"f")) ===> t else Id)(mk_trp(strict when_app))) [ + simp_tac(HOLCF_ss addsimps [when_appl,rep_strict]) 1]; val when_apps = let fun one_when n (con,args) = pg axs_con_def - (lift_defined % (nonlazy args, mk_trp(when_app`(con_app con args) === - mk_cfapp(%(nth_elem(n,when_funs cons)),map %# args))))[ - asm_simp_tac (HOLCF_ss addsimps [when_appl,ax_abs_iso]) 1]; - in mapn one_when 0 cons end; + (lift_defined % (nonlazy args, mk_trp(when_app`(con_app con args) === + mk_cfapp(%(nth_elem(n,when_funs cons)),map %# args))))[ + asm_simp_tac (HOLCF_ss addsimps [when_appl,ax_abs_iso]) 1]; + in mapn one_when 0 cons end; end; val when_rews = when_strict::when_apps; (* ----- theorems concerning the constructors, discriminators and selectors ------- *) val dis_stricts = map (fn (con,_) => pg axs_dis_def (mk_trp( - (if is_one_con_one_arg (K true) cons then mk_not else Id) - (strict(%%(dis_name con))))) [ - simp_tac (HOLCF_ss addsimps (if is_one_con_one_arg (K true) cons - then [ax_when_def] else when_rews)) 1]) cons; + (if is_one_con_one_arg (K true) cons then mk_not else Id) + (strict(%%(dis_name con))))) [ + simp_tac (HOLCF_ss addsimps (if is_one_con_one_arg (K true) cons + then [ax_when_def] else when_rews)) 1]) cons; val dis_apps = let fun one_dis c (con,args)= pg (axs_dis_def) - (lift_defined % (nonlazy args, (*(if is_one_con_one_arg is_lazy cons - then curry (lift_defined %#) args else Id) + (lift_defined % (nonlazy args, (*(if is_one_con_one_arg is_lazy cons + then curry (lift_defined %#) args else Id) #################*) - (mk_trp((%%(dis_name c))`(con_app con args) === - %%(if con=c then "TT" else "FF"))))) [ - asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; - in flat(map (fn (c,_) => map (one_dis c) cons) cons) end; + (mk_trp((%%(dis_name c))`(con_app con args) === + %%(if con=c then "TT" else "FF"))))) [ + asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; + in flat(map (fn (c,_) => map (one_dis c) cons) cons) end; val dis_defins = map (fn (con,args) => pg [] (defined(%x_name)==> - defined(%%(dis_name con)`%x_name)) [ - rtac cases 1, - contr_tac 1, - UNTIL_SOLVED (CHANGED(asm_simp_tac - (HOLCF_ss addsimps dis_apps) 1))]) cons; + defined(%%(dis_name con)`%x_name)) [ + rtac cases 1, + contr_tac 1, + UNTIL_SOLVED (CHANGED(asm_simp_tac + (HOLCF_ss addsimps dis_apps) 1))]) cons; val dis_rews = dis_stricts @ dis_defins @ dis_apps; val con_stricts = flat(map (fn (con,args) => map (fn vn => - pg (axs_con_def) - (mk_trp(con_app2 con (fn arg => if vname arg = vn - then UU else %# arg) args === UU))[ - asm_simp_tac (HOLCF_ss addsimps [abs_strict]) 1] - ) (nonlazy args)) cons); + pg (axs_con_def) + (mk_trp(con_app2 con (fn arg => if vname arg = vn + then UU else %# arg) args === UU))[ + asm_simp_tac (HOLCF_ss addsimps [abs_strict]) 1] + ) (nonlazy args)) cons); val con_defins = map (fn (con,args) => pg [] - (lift_defined % (nonlazy args, - mk_trp(defined(con_app con args)))) ([ - rtac swap3 1] @ (if is_one_con_one_arg (K true) cons - then [ - if is_lazy (hd args) then rtac defined_up 2 - else atac 2, - rtac abs_defin' 1, - asm_full_simp_tac (HOLCF_ss addsimps axs_con_def) 1] - else [ - eres_inst_tac [("f",dis_name con)] cfun_arg_cong 1, - asm_simp_tac (HOLCF_ss addsimps dis_rews) 1])))cons; + (lift_defined % (nonlazy args, + mk_trp(defined(con_app con args)))) ([ + rtac swap3 1] @ (if is_one_con_one_arg (K true) cons + then [ + if is_lazy (hd args) then rtac defined_up 2 + else atac 2, + rtac abs_defin' 1, + asm_full_simp_tac (HOLCF_ss addsimps axs_con_def) 1] + else [ + eres_inst_tac [("f",dis_name con)] cfun_arg_cong 1, + asm_simp_tac (HOLCF_ss addsimps dis_rews) 1])))cons; val con_rews = con_stricts @ con_defins; val sel_stricts = let fun one_sel sel = pg axs_sel_def (mk_trp(strict(%%sel))) [ - simp_tac (HOLCF_ss addsimps when_rews) 1]; + simp_tac (HOLCF_ss addsimps when_rews) 1]; in flat(map (fn (_,args) => map (fn arg => one_sel (sel_of arg)) args) cons) end; val sel_apps = let fun one_sel c n sel = map (fn (con,args) => - let val nlas = nonlazy args; - val vns = map vname args; - in pg axs_sel_def (lift_defined % - (filter (fn v => con=c andalso (v<>nth_elem(n,vns))) nlas, + let val nlas = nonlazy args; + val vns = map vname args; + in pg axs_sel_def (lift_defined % + (filter (fn v => con=c andalso (v<>nth_elem(n,vns))) nlas, mk_trp((%%sel)`(con_app con args) === (if con=c then %(nth_elem(n,vns)) else UU)))) - ( (if con=c then [] - else map(case_UU_tac(when_rews@con_stricts)1) nlas) - @(if con=c andalso ((nth_elem(n,vns)) mem nlas) - then[case_UU_tac (when_rews @ con_stricts) 1 - (nth_elem(n,vns))] else []) - @ [asm_simp_tac(HOLCF_ss addsimps when_rews)1])end) cons; + ( (if con=c then [] + else map(case_UU_tac(when_rews@con_stricts)1) nlas) + @(if con=c andalso ((nth_elem(n,vns)) mem nlas) + then[case_UU_tac (when_rews @ con_stricts) 1 + (nth_elem(n,vns))] else []) + @ [asm_simp_tac(HOLCF_ss addsimps when_rews)1])end) cons; in flat(map (fn (c,args) => - flat(mapn (fn n => fn arg => one_sel c n (sel_of arg)) 0 args)) cons) end; + flat(mapn (fn n => fn arg => one_sel c n (sel_of arg)) 0 args)) cons) end; val sel_defins = if length cons = 1 then map (fn arg => pg [] (defined(%x_name) ==> - defined(%%(sel_of arg)`%x_name)) [ - rtac cases 1, - contr_tac 1, - UNTIL_SOLVED (CHANGED(asm_simp_tac - (HOLCF_ss addsimps sel_apps) 1))]) - (filter_out is_lazy (snd(hd cons))) else []; + defined(%%(sel_of arg)`%x_name)) [ + rtac cases 1, + contr_tac 1, + UNTIL_SOLVED (CHANGED(asm_simp_tac + (HOLCF_ss addsimps sel_apps) 1))]) + (filter_out is_lazy (snd(hd cons))) else []; val sel_rews = sel_stricts @ sel_defins @ sel_apps; val distincts_le = let fun dist (con1, args1) (con2, args2) = pg [] - (lift_defined % ((nonlazy args1), - (mk_trp (con_app con1 args1 ~<< con_app con2 args2))))([ - rtac swap3 1, - eres_inst_tac [("fo5",dis_name con1)] monofun_cfun_arg 1] - @ map (case_UU_tac (con_stricts @ dis_rews) 1) (nonlazy args2) - @[asm_simp_tac (HOLCF_ss addsimps dis_rews) 1]); + (lift_defined % ((nonlazy args1), + (mk_trp (con_app con1 args1 ~<< con_app con2 args2))))([ + rtac swap3 1, + eres_inst_tac [("fo5",dis_name con1)] monofun_cfun_arg 1] + @ map (case_UU_tac (con_stricts @ dis_rews) 1) (nonlazy args2) + @[asm_simp_tac (HOLCF_ss addsimps dis_rews) 1]); fun distinct (con1,args1) (con2,args2) = - let val arg1 = (con1, args1); - val arg2 = (con2, (map (fn (arg,vn) => upd_vname (K vn) arg) - (args2~~variantlist(map vname args2,map vname args1)))); - in [dist arg1 arg2, dist arg2 arg1] end; + let val arg1 = (con1, args1); + val arg2 = (con2, (map (fn (arg,vn) => upd_vname (K vn) arg) + (args2~~variantlist(map vname args2,map vname args1)))); + in [dist arg1 arg2, dist arg2 arg1] end; fun distincts [] = [] | distincts (c::cs) = (map (distinct c) cs) :: distincts cs; in distincts cons end; val dists_le = flat (flat distincts_le); val dists_eq = let fun distinct (_,args1) ((_,args2),leqs) = let - val (le1,le2) = (hd leqs, hd(tl leqs)); - val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI) in - if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else - if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else - [eq1, eq2] end; + val (le1,le2) = (hd leqs, hd(tl leqs)); + val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI) in + if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else + if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else + [eq1, eq2] end; fun distincts [] = [] | distincts ((c,leqs)::cs) = flat(map (distinct c) ((map fst cs)~~leqs)) @ - distincts cs; + distincts cs; in distincts (cons~~distincts_le) end; local fun pgterm rel con args = let - fun append s = upd_vname(fn v => v^s); - val (largs,rargs) = (args, map (append "'") args); - in pg [] (mk_trp (rel(con_app con largs,con_app con rargs)) ===> - lift_defined % ((nonlazy largs),lift_defined % ((nonlazy rargs), - mk_trp (foldr' mk_conj - (map rel (map %# largs ~~ map %# rargs)))))) end; + fun append s = upd_vname(fn v => v^s); + val (largs,rargs) = (args, map (append "'") args); + in pg [] (mk_trp (rel(con_app con largs,con_app con rargs)) ===> + lift_defined % ((nonlazy largs),lift_defined % ((nonlazy rargs), + mk_trp (foldr' mk_conj + (map rel (map %# largs ~~ map %# rargs)))))) end; val cons' = filter (fn (_,args) => args<>[]) cons; in val inverts = map (fn (con,args) => - pgterm (op <<) con args (flat(map (fn arg => [ - TRY(rtac conjI 1), - dres_inst_tac [("fo5",sel_of arg)] monofun_cfun_arg 1, - asm_full_simp_tac (HOLCF_ss addsimps sel_apps) 1] - ) args))) cons'; + pgterm (op <<) con args (flat(map (fn arg => [ + TRY(rtac conjI 1), + dres_inst_tac [("fo5",sel_of arg)] monofun_cfun_arg 1, + asm_full_simp_tac (HOLCF_ss addsimps sel_apps) 1] + ) args))) cons'; val injects = map (fn ((con,args),inv_thm) => - pgterm (op ===) con args [ - etac (antisym_less_inverse RS conjE) 1, - dtac inv_thm 1, REPEAT(atac 1), - dtac inv_thm 1, REPEAT(atac 1), - TRY(safe_tac HOL_cs), - REPEAT(rtac antisym_less 1 ORELSE atac 1)] ) - (cons'~~inverts); + pgterm (op ===) con args [ + etac (antisym_less_inverse RS conjE) 1, + dtac inv_thm 1, REPEAT(atac 1), + dtac inv_thm 1, REPEAT(atac 1), + TRY(safe_tac HOL_cs), + REPEAT(rtac antisym_less 1 ORELSE atac 1)] ) + (cons'~~inverts); end; (* ----- theorems concerning one induction step ----------------------------------- *) val copy_strict = pg [ax_copy_def] ((if is_one_con_one_arg (K true) cons then fn t => - mk_trp(strict(cproj (%"f") eqs (rec_of (hd(snd(hd cons)))))) ===> t - else Id) (mk_trp(strict(dc_copy`%"f")))) [ - asm_simp_tac(HOLCF_ss addsimps [abs_strict,rep_strict, - cfst_strict,csnd_strict]) 1]; + mk_trp(strict(cproj (%"f") eqs (rec_of (hd(snd(hd cons)))))) ===> t + else Id) (mk_trp(strict(dc_copy`%"f")))) [ + asm_simp_tac(HOLCF_ss addsimps [abs_strict,rep_strict, + cfst_strict,csnd_strict]) 1]; val copy_apps = map (fn (con,args) => pg (ax_copy_def::axs_con_def) - (lift_defined %# (filter is_nonlazy_rec args, - mk_trp(dc_copy`%"f"`(con_app con args) === - (con_app2 con (app_rec_arg (cproj (%"f") eqs)) args)))) - (map (case_UU_tac [ax_abs_iso] 1 o vname) - (filter(fn a=>not(is_rec a orelse is_lazy a))args)@ - [asm_simp_tac (HOLCF_ss addsimps [ax_abs_iso]) 1]) - )cons; + (lift_defined %# (filter is_nonlazy_rec args, + mk_trp(dc_copy`%"f"`(con_app con args) === + (con_app2 con (app_rec_arg (cproj (%"f") eqs)) args)))) + (map (case_UU_tac [ax_abs_iso] 1 o vname) + (filter(fn a=>not(is_rec a orelse is_lazy a))args)@ + [asm_simp_tac (HOLCF_ss addsimps [ax_abs_iso]) 1]) + )cons; val copy_stricts = map(fn(con,args)=>pg[](mk_trp(dc_copy`UU`(con_app con args) ===UU)) - (let val rews = cfst_strict::csnd_strict::copy_strict::copy_apps@con_rews - in map (case_UU_tac rews 1) (nonlazy args) @ [ - asm_simp_tac (HOLCF_ss addsimps rews) 1] end)) - (filter (fn (_,args)=>exists is_nonlazy_rec args) cons); + (let val rews = cfst_strict::csnd_strict::copy_strict::copy_apps@con_rews + in map (case_UU_tac rews 1) (nonlazy args) @ [ + asm_simp_tac (HOLCF_ss addsimps rews) 1] end)) + (filter (fn (_,args)=>exists is_nonlazy_rec args) cons); val copy_rews = copy_strict::copy_apps @ copy_stricts; in (iso_rews, exhaust, cases, when_rews, - con_rews, sel_rews, dis_rews, dists_eq, dists_le, inverts, injects, - copy_rews) + con_rews, sel_rews, dis_rews, dists_eq, dists_le, inverts, injects, + copy_rews) end; (* let *) @@ -369,186 +369,186 @@ val P_name = idx_name dnames "P"; local - val iterate_ss = simpset_of "Fix"; + val iterate_ss = simpset_of "Fix"; val iterate_Cprod_strict_ss = iterate_ss addsimps [cfst_strict, csnd_strict]; val iterate_Cprod_ss = iterate_ss addsimps [cfst2,csnd2,csplit2]; val copy_con_rews = copy_rews @ con_rews; val copy_take_defs = (if length dnames=1 then [] else [ax_copy2_def]) @axs_take_def; val take_stricts = pg copy_take_defs (mk_trp(foldr' mk_conj (map (fn ((dn,args),_)=> - (dc_take dn $ %"n")`UU === mk_constrain(Type(dn,args),UU)) eqs)))([ - nat_ind_tac "n" 1, - simp_tac iterate_ss 1, - simp_tac iterate_Cprod_strict_ss 1, - asm_simp_tac iterate_Cprod_ss 1, - TRY(safe_tac HOL_cs)] @ - map(K(asm_simp_tac (HOL_ss addsimps copy_rews)1))dnames); + (dc_take dn $ %"n")`UU === mk_constrain(Type(dn,args),UU)) eqs)))([ + nat_ind_tac "n" 1, + simp_tac iterate_ss 1, + simp_tac iterate_Cprod_strict_ss 1, + asm_simp_tac iterate_Cprod_ss 1, + TRY(safe_tac HOL_cs)] @ + map(K(asm_simp_tac (HOL_ss addsimps copy_rews)1))dnames); val take_stricts' = rewrite_rule copy_take_defs take_stricts; val take_0s = mapn (fn n => fn dn => pg axs_take_def(mk_trp((dc_take dn $ %%"0") - `%x_name n === UU))[ - simp_tac iterate_Cprod_strict_ss 1]) 1 dnames; + `%x_name n === UU))[ + simp_tac iterate_Cprod_strict_ss 1]) 1 dnames; val take_apps = pg copy_take_defs (mk_trp(foldr' mk_conj - (flat(map (fn ((dn,_),cons) => map (fn (con,args) => foldr mk_all - (map vname args,(dc_take dn $ (%%"Suc" $ %"n"))`(con_app con args) === - con_app2 con (app_rec_arg (fn n=>dc_take (nth_elem(n,dnames))$ %"n")) - args)) cons) eqs)))) ([ - nat_ind_tac "n" 1, - simp_tac iterate_Cprod_strict_ss 1, - simp_tac (HOLCF_ss addsimps copy_con_rews) 1, - TRY(safe_tac HOL_cs)] @ - (flat(map (fn ((dn,_),cons) => map (fn (con,args) => EVERY ( - asm_full_simp_tac iterate_Cprod_ss 1:: - map (case_UU_tac (take_stricts'::copy_con_rews) 1) - (nonlazy args) @[ - asm_full_simp_tac (HOLCF_ss addsimps copy_rews) 1]) - ) cons) eqs))); + (flat(map (fn ((dn,_),cons) => map (fn (con,args) => foldr mk_all + (map vname args,(dc_take dn $ (%%"Suc" $ %"n"))`(con_app con args) === + con_app2 con (app_rec_arg (fn n=>dc_take (nth_elem(n,dnames))$ %"n")) + args)) cons) eqs)))) ([ + nat_ind_tac "n" 1, + simp_tac iterate_Cprod_strict_ss 1, + simp_tac (HOLCF_ss addsimps copy_con_rews) 1, + TRY(safe_tac HOL_cs)] @ + (flat(map (fn ((dn,_),cons) => map (fn (con,args) => EVERY ( + asm_full_simp_tac iterate_Cprod_ss 1:: + map (case_UU_tac (take_stricts'::copy_con_rews) 1) + (nonlazy args) @[ + asm_full_simp_tac (HOLCF_ss addsimps copy_rews) 1]) + ) cons) eqs))); in val take_rews = atomize take_stricts @ take_0s @ atomize take_apps; end; (* local *) val take_lemmas = mapn (fn n => fn(dn,ax_reach) => pg'' thy axs_take_def (mk_All("n", - mk_trp(dc_take dn $ Bound 0 `%(x_name n) === - dc_take dn $ Bound 0 `%(x_name n^"'"))) - ===> mk_trp(%(x_name n) === %(x_name n^"'"))) (fn prems => [ - res_inst_tac[("t",x_name n )](ax_reach RS subst) 1, - res_inst_tac[("t",x_name n^"'")](ax_reach RS subst) 1, - rtac (fix_def2 RS ssubst) 1, - REPEAT(CHANGED(rtac (contlub_cfun_arg RS ssubst) 1 - THEN chain_tac 1)), - rtac (contlub_cfun_fun RS ssubst) 1, - rtac (contlub_cfun_fun RS ssubst) 2, - rtac lub_equal 3, - chain_tac 1, - rtac allI 1, - resolve_tac prems 1])) 1 (dnames~~axs_reach); + mk_trp(dc_take dn $ Bound 0 `%(x_name n) === + dc_take dn $ Bound 0 `%(x_name n^"'"))) + ===> mk_trp(%(x_name n) === %(x_name n^"'"))) (fn prems => [ + res_inst_tac[("t",x_name n )](ax_reach RS subst) 1, + res_inst_tac[("t",x_name n^"'")](ax_reach RS subst) 1, + rtac (fix_def2 RS ssubst) 1, + REPEAT(CHANGED(rtac (contlub_cfun_arg RS ssubst) 1 + THEN chain_tac 1)), + rtac (contlub_cfun_fun RS ssubst) 1, + rtac (contlub_cfun_fun RS ssubst) 2, + rtac lub_equal 3, + chain_tac 1, + rtac allI 1, + resolve_tac prems 1])) 1 (dnames~~axs_reach); local fun one_con p (con,args) = foldr mk_All (map vname args, - lift_defined (bound_arg (map vname args)) (nonlazy args, - lift (fn arg => %(P_name (1+rec_of arg)) $ bound_arg args arg) - (filter is_rec args,mk_trp(%p $ con_app2 con (bound_arg args) args)))); + lift_defined (bound_arg (map vname args)) (nonlazy args, + lift (fn arg => %(P_name (1+rec_of arg)) $ bound_arg args arg) + (filter is_rec args,mk_trp(%p $ con_app2 con (bound_arg args) args)))); fun one_eq ((p,cons),concl) = (mk_trp(%p $ UU) ===> - foldr (op ===>) (map (one_con p) cons,concl)); + foldr (op ===>) (map (one_con p) cons,concl)); fun ind_term concf = foldr one_eq (mapn (fn n => fn x => (P_name n, x)) 1 conss, - mk_trp(foldr' mk_conj (mapn (fn n => concf (P_name n,x_name n)) 1 dnames))); + mk_trp(foldr' mk_conj (mapn (fn n => concf (P_name n,x_name n)) 1 dnames))); val take_ss = HOL_ss addsimps take_rews; fun ind_tacs tacsf thms1 thms2 prems = TRY(safe_tac HOL_cs):: - flat (mapn (fn n => fn (thm1,thm2) => - tacsf (n,prems) (thm1,thm2) @ - flat (map (fn cons => - (resolve_tac prems 1 :: - flat (map (fn (_,args) => - resolve_tac prems 1:: - map (K(atac 1)) (nonlazy args) @ - map (K(atac 1)) (filter is_rec args)) - cons))) - conss)) - 0 (thms1~~thms2)); + flat (mapn (fn n => fn (thm1,thm2) => + tacsf (n,prems) (thm1,thm2) @ + flat (map (fn cons => + (resolve_tac prems 1 :: + flat (map (fn (_,args) => + resolve_tac prems 1:: + map (K(atac 1)) (nonlazy args) @ + map (K(atac 1)) (filter is_rec args)) + cons))) + conss)) + 0 (thms1~~thms2)); local fun all_rec_to ns lazy_rec (n,cons) = forall (exists (fn arg => - is_rec arg andalso not(rec_of arg mem ns) andalso - ((rec_of arg = n andalso not(lazy_rec orelse is_lazy arg)) orelse - rec_of arg <> n andalso all_rec_to (rec_of arg::ns) - (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss)))) - ) o snd) cons; + is_rec arg andalso not(rec_of arg mem ns) andalso + ((rec_of arg = n andalso not(lazy_rec orelse is_lazy arg)) orelse + rec_of arg <> n andalso all_rec_to (rec_of arg::ns) + (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss)))) + ) o snd) cons; fun warn (n,cons) = if all_rec_to [] false (n,cons) then (writeln - ("WARNING: domain "^nth_elem(n,dnames)^" is empty!"); true) - else false; + ("WARNING: domain "^nth_elem(n,dnames)^" is empty!"); true) + else false; fun lazy_rec_to ns lazy_rec (n,cons) = exists (exists (fn arg => - is_rec arg andalso not(rec_of arg mem ns) andalso - ((rec_of arg = n andalso (lazy_rec orelse is_lazy arg)) orelse - rec_of arg <> n andalso lazy_rec_to (rec_of arg::ns) - (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss)))) - ) o snd) cons; + is_rec arg andalso not(rec_of arg mem ns) andalso + ((rec_of arg = n andalso (lazy_rec orelse is_lazy arg)) orelse + rec_of arg <> n andalso lazy_rec_to (rec_of arg::ns) + (lazy_rec orelse is_lazy arg) (n, (nth_elem(rec_of arg,conss)))) + ) o snd) cons; in val is_emptys = map warn (mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs); val is_finite = forall (not o lazy_rec_to [] false) - (mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs) + (mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs) end; in val finite_ind = pg'' thy [] (ind_term (fn (P,x) => fn dn => - mk_all(x,%P $ (dc_take dn $ %"n" `Bound 0)))) (fn prems=> [ - nat_ind_tac "n" 1, - simp_tac (take_ss addsimps prems) 1, - TRY(safe_tac HOL_cs)] - @ flat(mapn (fn n => fn (cons,cases) => [ - res_inst_tac [("x",x_name n)] cases 1, - asm_simp_tac (take_ss addsimps prems) 1] - @ flat(map (fn (con,args) => - asm_simp_tac take_ss 1 :: - map (fn arg => - case_UU_tac (prems@con_rews) 1 ( - nth_elem(rec_of arg,dnames)^"_take n1`"^vname arg)) - (filter is_nonlazy_rec args) @ [ - resolve_tac prems 1] @ - map (K (atac 1)) (nonlazy args) @ - map (K (etac spec 1)) (filter is_rec args)) - cons)) - 1 (conss~~casess))); + mk_all(x,%P $ (dc_take dn $ %"n" `Bound 0)))) (fn prems=> [ + nat_ind_tac "n" 1, + simp_tac (take_ss addsimps prems) 1, + TRY(safe_tac HOL_cs)] + @ flat(mapn (fn n => fn (cons,cases) => [ + res_inst_tac [("x",x_name n)] cases 1, + asm_simp_tac (take_ss addsimps prems) 1] + @ flat(map (fn (con,args) => + asm_simp_tac take_ss 1 :: + map (fn arg => + case_UU_tac (prems@con_rews) 1 ( + nth_elem(rec_of arg,dnames)^"_take n1`"^vname arg)) + (filter is_nonlazy_rec args) @ [ + resolve_tac prems 1] @ + map (K (atac 1)) (nonlazy args) @ + map (K (etac spec 1)) (filter is_rec args)) + cons)) + 1 (conss~~casess))); val (finites,ind) = if is_finite then let fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %"x" === %"x"); val finite_lemmas1a = map (fn dn => pg [] (mk_trp(defined (%"x")) ===> - mk_trp(mk_disj(mk_all("n",dc_take dn $ Bound 0 ` %"x" === UU), - take_enough dn)) ===> mk_trp(take_enough dn)) [ - etac disjE 1, - etac notE 1, - resolve_tac take_lemmas 1, - asm_simp_tac take_ss 1, - atac 1]) dnames; + mk_trp(mk_disj(mk_all("n",dc_take dn $ Bound 0 ` %"x" === UU), + take_enough dn)) ===> mk_trp(take_enough dn)) [ + etac disjE 1, + etac notE 1, + resolve_tac take_lemmas 1, + asm_simp_tac take_ss 1, + atac 1]) dnames; val finite_lemma1b = pg [] (mk_trp (mk_all("n",foldr' mk_conj (mapn - (fn n => fn ((dn,args),_) => mk_constrainall(x_name n,Type(dn,args), - mk_disj(dc_take dn $ Bound 1 ` Bound 0 === UU, - dc_take dn $ Bound 1 ` Bound 0 === Bound 0))) 1 eqs)))) ([ - rtac allI 1, - nat_ind_tac "n" 1, - simp_tac take_ss 1, - TRY(safe_tac(empty_cs addSEs[conjE] addSIs[conjI]))] @ - flat(mapn (fn n => fn (cons,cases) => [ - simp_tac take_ss 1, - rtac allI 1, - res_inst_tac [("x",x_name n)] cases 1, - asm_simp_tac take_ss 1] @ - flat(map (fn (con,args) => - asm_simp_tac take_ss 1 :: - flat(map (fn arg => [ - eres_inst_tac [("x",vname arg)] all_dupE 1, - etac disjE 1, - asm_simp_tac (HOL_ss addsimps con_rews) 1, - asm_simp_tac take_ss 1]) - (filter is_nonlazy_rec args))) - cons)) - 1 (conss~~casess))) handle ERROR => raise ERROR; + (fn n => fn ((dn,args),_) => mk_constrainall(x_name n,Type(dn,args), + mk_disj(dc_take dn $ Bound 1 ` Bound 0 === UU, + dc_take dn $ Bound 1 ` Bound 0 === Bound 0))) 1 eqs)))) ([ + rtac allI 1, + nat_ind_tac "n" 1, + simp_tac take_ss 1, + TRY(safe_tac(empty_cs addSEs[conjE] addSIs[conjI]))] @ + flat(mapn (fn n => fn (cons,cases) => [ + simp_tac take_ss 1, + rtac allI 1, + res_inst_tac [("x",x_name n)] cases 1, + asm_simp_tac take_ss 1] @ + flat(map (fn (con,args) => + asm_simp_tac take_ss 1 :: + flat(map (fn arg => [ + eres_inst_tac [("x",vname arg)] all_dupE 1, + etac disjE 1, + asm_simp_tac (HOL_ss addsimps con_rews) 1, + asm_simp_tac take_ss 1]) + (filter is_nonlazy_rec args))) + cons)) + 1 (conss~~casess))) handle ERROR => raise ERROR; val all_finite=map (fn(dn,l1b)=>pg axs_finite_def (mk_trp(%%(dn^"_finite") $ %"x"))[ - case_UU_tac take_rews 1 "x", - eresolve_tac finite_lemmas1a 1, - step_tac HOL_cs 1, - step_tac HOL_cs 1, - cut_facts_tac [l1b] 1, - fast_tac HOL_cs 1]) (dnames~~atomize finite_lemma1b); + case_UU_tac take_rews 1 "x", + eresolve_tac finite_lemmas1a 1, + step_tac HOL_cs 1, + step_tac HOL_cs 1, + cut_facts_tac [l1b] 1, + fast_tac HOL_cs 1]) (dnames~~atomize finite_lemma1b); in (all_finite, pg'' thy [] (ind_term (fn (P,x) => fn dn => %P $ %x)) - (ind_tacs (fn _ => fn (all_fin,finite_ind) => [ - rtac (rewrite_rule axs_finite_def all_fin RS exE) 1, - etac subst 1, - rtac finite_ind 1]) all_finite (atomize finite_ind)) + (ind_tacs (fn _ => fn (all_fin,finite_ind) => [ + rtac (rewrite_rule axs_finite_def all_fin RS exE) 1, + etac subst 1, + rtac finite_ind 1]) all_finite (atomize finite_ind)) ) end (* let *) else (mapn (fn n => fn dn => read_instantiate_sg (sign_of thy) - [("P",dn^"_finite "^x_name n)] excluded_middle) 1 dnames, + [("P",dn^"_finite "^x_name n)] excluded_middle) 1 dnames, pg'' thy [] (foldr (op ===>) (mapn (fn n =>K(mk_trp(%%"adm" $ %(P_name n))))1 - dnames,ind_term (fn(P,x)=>fn dn=> %P $ %x))) - (ind_tacs (fn (n,prems) => fn (ax_reach,finite_ind) =>[ - rtac (ax_reach RS subst) 1, - res_inst_tac [("x",x_name n)] spec 1, - rtac wfix_ind 1, - rtac adm_impl_admw 1, - resolve_tac adm_thms 1, - rtac adm_subst 1, - cont_tacR 1, - resolve_tac prems 1, - strip_tac 1, - rtac(rewrite_rule axs_take_def finite_ind) 1]) - axs_reach (atomize finite_ind)) + dnames,ind_term (fn(P,x)=>fn dn=> %P $ %x))) + (ind_tacs (fn (n,prems) => fn (ax_reach,finite_ind) =>[ + rtac (ax_reach RS subst) 1, + res_inst_tac [("x",x_name n)] spec 1, + rtac wfix_ind 1, + rtac adm_impl_admw 1, + resolve_tac adm_thms 1, + rtac adm_subst 1, + cont_tacR 1, + resolve_tac prems 1, + strip_tac 1, + rtac(rewrite_rule axs_take_def finite_ind) 1]) + axs_reach (atomize finite_ind)) ) end; (* local *) @@ -558,34 +558,34 @@ val take_ss = HOL_ss addsimps take_rews; val sproj = bin_branchr (fn s => "fst("^s^")") (fn s => "snd("^s^")"); val coind_lemma = pg [ax_bisim_def] (mk_trp(mk_imp(%%(comp_dname^"_bisim") $ %"R", - foldr (fn (x,t)=> mk_all(x,mk_all(x^"'",t))) (xs, - foldr mk_imp (mapn (fn n => K(proj (%"R") dnames n $ - bnd_arg n 0 $ bnd_arg n 1)) 0 dnames, - foldr' mk_conj (mapn (fn n => fn dn => - (dc_take dn $ %"n" `bnd_arg n 0 === - (dc_take dn $ %"n" `bnd_arg n 1))) 0 dnames)))))) ([ - rtac impI 1, - nat_ind_tac "n" 1, - simp_tac take_ss 1, - safe_tac HOL_cs] @ - flat(mapn (fn n => fn x => [ - etac allE 1, etac allE 1, - eres_inst_tac [("P1",sproj "R" dnames n^ - " "^x^" "^x^"'")](mp RS disjE) 1, - TRY(safe_tac HOL_cs), - REPEAT(CHANGED(asm_simp_tac take_ss 1))]) - 0 xs)); + foldr (fn (x,t)=> mk_all(x,mk_all(x^"'",t))) (xs, + foldr mk_imp (mapn (fn n => K(proj (%"R") dnames n $ + bnd_arg n 0 $ bnd_arg n 1)) 0 dnames, + foldr' mk_conj (mapn (fn n => fn dn => + (dc_take dn $ %"n" `bnd_arg n 0 === + (dc_take dn $ %"n" `bnd_arg n 1))) 0 dnames)))))) ([ + rtac impI 1, + nat_ind_tac "n" 1, + simp_tac take_ss 1, + safe_tac HOL_cs] @ + flat(mapn (fn n => fn x => [ + etac allE 1, etac allE 1, + eres_inst_tac [("P1",sproj "R" dnames n^ + " "^x^" "^x^"'")](mp RS disjE) 1, + TRY(safe_tac HOL_cs), + REPEAT(CHANGED(asm_simp_tac take_ss 1))]) + 0 xs)); in val coind = pg [] (mk_trp(%%(comp_dname^"_bisim") $ %"R") ===> - foldr (op ===>) (mapn (fn n => fn x => - mk_trp(proj (%"R") dnames n $ %x $ %(x^"'"))) 0 xs, - mk_trp(foldr' mk_conj (map (fn x => %x === %(x^"'")) xs)))) ([ - TRY(safe_tac HOL_cs)] @ - flat(map (fn take_lemma => [ - rtac take_lemma 1, - cut_facts_tac [coind_lemma] 1, - fast_tac HOL_cs 1]) - take_lemmas)); + foldr (op ===>) (mapn (fn n => fn x => + mk_trp(proj (%"R") dnames n $ %x $ %(x^"'"))) 0 xs, + mk_trp(foldr' mk_conj (map (fn x => %x === %(x^"'")) xs)))) ([ + TRY(safe_tac HOL_cs)] @ + flat(map (fn take_lemma => [ + rtac take_lemma 1, + cut_facts_tac [coind_lemma] 1, + fast_tac HOL_cs 1]) + take_lemmas)); end; (* local *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/ex/Fix2.ML --- a/src/HOLCF/ex/Fix2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/ex/Fix2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,28 +1,28 @@ -(* Title: HOLCF/ex/Fix2.ML +(* Title: HOLCF/ex/Fix2.ML ID: $Id$ - Author: Franz Regensburger - Copyright 1995 Technische Universitaet Muenchen + Author: Franz Regensburger + Copyright 1995 Technische Universitaet Muenchen *) open Fix2; val lemma1 = prove_goal Fix2.thy "fix = gix" (fn prems => - [ - (rtac ext_cfun 1), - (rtac antisym_less 1), - (rtac fix_least 1), - (rtac gix1_def 1), - (rtac gix2_def 1), - (rtac (fix_eq RS sym) 1) - ]); + [ + (rtac ext_cfun 1), + (rtac antisym_less 1), + (rtac fix_least 1), + (rtac gix1_def 1), + (rtac gix2_def 1), + (rtac (fix_eq RS sym) 1) + ]); val lemma2 = prove_goal Fix2.thy "gix`F=lub(range(%i. iterate i F UU))" (fn prems => - [ - (rtac (lemma1 RS subst) 1), - (rtac fix_def2 1) - ]); + [ + (rtac (lemma1 RS subst) 1), + (rtac fix_def2 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/ex/Hoare.ML --- a/src/HOLCF/ex/Hoare.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/ex/Hoare.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,7 +1,7 @@ -(* Title: HOLCF/ex/hoare.ML +(* Title: HOLCF/ex/hoare.ML ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen *) open Hoare; @@ -10,95 +10,95 @@ val hoare_lemma2 = prove_goal HOLCF.thy "b~=TT ==> b=FF | b=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (Exh_tr RS disjE) 1), - (fast_tac HOL_cs 1), - (etac disjE 1), - (contr_tac 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (Exh_tr RS disjE) 1), + (fast_tac HOL_cs 1), + (etac disjE 1), + (contr_tac 1), + (fast_tac HOL_cs 1) + ]); val hoare_lemma3 = prove_goal HOLCF.thy " (!k.b1`(iterate k g x) = TT) | (? k. b1`(iterate k g x)~=TT)" (fn prems => - [ - (fast_tac HOL_cs 1) - ]); + [ + (fast_tac HOL_cs 1) + ]); val hoare_lemma4 = prove_goal HOLCF.thy "(? k. b1`(iterate k g x) ~= TT) ==> \ \ ? k. b1`(iterate k g x) = FF | b1`(iterate k g x) = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (etac exE 1), - (rtac exI 1), - (rtac hoare_lemma2 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac exE 1), + (rtac exI 1), + (rtac hoare_lemma2 1), + (atac 1) + ]); val hoare_lemma5 = prove_goal HOLCF.thy "[|(? k. b1`(iterate k g x) ~= TT);\ \ k=theleast(%n. b1`(iterate n g x) ~= TT)|] ==> \ \ b1`(iterate k g x)=FF | b1`(iterate k g x)=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (hyp_subst_tac 1), - (rtac hoare_lemma2 1), - (etac exE 1), - (etac theleast1 1) - ]); + [ + (cut_facts_tac prems 1), + (hyp_subst_tac 1), + (rtac hoare_lemma2 1), + (etac exE 1), + (etac theleast1 1) + ]); val hoare_lemma6 = prove_goal HOLCF.thy "b=UU ==> b~=TT" (fn prems => - [ - (cut_facts_tac prems 1), - (hyp_subst_tac 1), - (resolve_tac dist_eq_tr 1) - ]); + [ + (cut_facts_tac prems 1), + (hyp_subst_tac 1), + (resolve_tac dist_eq_tr 1) + ]); val hoare_lemma7 = prove_goal HOLCF.thy "b=FF ==> b~=TT" (fn prems => - [ - (cut_facts_tac prems 1), - (hyp_subst_tac 1), - (resolve_tac dist_eq_tr 1) - ]); + [ + (cut_facts_tac prems 1), + (hyp_subst_tac 1), + (resolve_tac dist_eq_tr 1) + ]); val hoare_lemma8 = prove_goal HOLCF.thy "[|(? k. b1`(iterate k g x) ~= TT);\ \ k=theleast(%n. b1`(iterate n g x) ~= TT)|] ==> \ \ !m. m < k --> b1`(iterate m g x)=TT" (fn prems => - [ - (cut_facts_tac prems 1), - (hyp_subst_tac 1), - (etac exE 1), - (strip_tac 1), - (res_inst_tac [("p","b1`(iterate m g x)")] trE 1), - (atac 2), - (rtac (le_less_trans RS less_anti_refl) 1), - (atac 2), - (rtac theleast2 1), - (etac hoare_lemma6 1), - (rtac (le_less_trans RS less_anti_refl) 1), - (atac 2), - (rtac theleast2 1), - (etac hoare_lemma7 1) - ]); + [ + (cut_facts_tac prems 1), + (hyp_subst_tac 1), + (etac exE 1), + (strip_tac 1), + (res_inst_tac [("p","b1`(iterate m g x)")] trE 1), + (atac 2), + (rtac (le_less_trans RS less_anti_refl) 1), + (atac 2), + (rtac theleast2 1), + (etac hoare_lemma6 1), + (rtac (le_less_trans RS less_anti_refl) 1), + (atac 2), + (rtac theleast2 1), + (etac hoare_lemma7 1) + ]); val hoare_lemma28 = prove_goal HOLCF.thy "b1`(y::'a)=(UU::tr) ==> b1`UU = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (etac (flat_tr RS flat_codom RS disjE) 1), - (atac 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (etac (flat_tr RS flat_codom RS disjE) 1), + (atac 1), + (etac spec 1) + ]); (* ----- access to definitions ----- *) @@ -106,34 +106,34 @@ val p_def2 = prove_goalw Hoare.thy [p_def] "p = fix`(LAM f x. If b1`x then f`(g`x) else x fi)" (fn prems => - [ - (rtac refl 1) - ]); + [ + (rtac refl 1) + ]); val q_def2 = prove_goalw Hoare.thy [q_def] "q = fix`(LAM f x. If b1`x orelse b2`x then \ \ f`(g`x) else x fi)" (fn prems => - [ - (rtac refl 1) - ]); + [ + (rtac refl 1) + ]); val p_def3 = prove_goal Hoare.thy "p`x = If b1`x then p`(g`x) else x fi" (fn prems => - [ - (fix_tac3 p_def 1), - (Simp_tac 1) - ]); + [ + (fix_tac3 p_def 1), + (Simp_tac 1) + ]); val q_def3 = prove_goal Hoare.thy "q`x = If b1`x orelse b2`x then q`(g`x) else x fi" (fn prems => - [ - (fix_tac3 q_def 1), - (Simp_tac 1) - ]); + [ + (fix_tac3 q_def 1), + (Simp_tac 1) + ]); (** --------- proves about iterations of p and q ---------- **) @@ -141,51 +141,51 @@ "(! m. m< Suc k --> b1`(iterate m g x)=TT) -->\ \ p`(iterate k g x)=p`x" (fn prems => - [ - (nat_ind_tac "k" 1), - (simp_tac (!simpset delsimps [less_Suc_eq]) 1), - (simp_tac (!simpset delsimps [less_Suc_eq]) 1), - (strip_tac 1), - (res_inst_tac [("s","p`(iterate k1 g x)")] trans 1), - (rtac trans 1), - (rtac (p_def3 RS sym) 2), - (res_inst_tac [("s","TT"),("t","b1`(iterate k1 g x)")] ssubst 1), - (rtac mp 1), - (etac spec 1), - (Simp_tac 1), - (simp_tac HOLCF_ss 1), - (etac mp 1), - (strip_tac 1), - (rtac mp 1), - (etac spec 1), - (etac less_trans 1), - (Simp_tac 1) - ]); + [ + (nat_ind_tac "k" 1), + (simp_tac (!simpset delsimps [less_Suc_eq]) 1), + (simp_tac (!simpset delsimps [less_Suc_eq]) 1), + (strip_tac 1), + (res_inst_tac [("s","p`(iterate k1 g x)")] trans 1), + (rtac trans 1), + (rtac (p_def3 RS sym) 2), + (res_inst_tac [("s","TT"),("t","b1`(iterate k1 g x)")] ssubst 1), + (rtac mp 1), + (etac spec 1), + (Simp_tac 1), + (simp_tac HOLCF_ss 1), + (etac mp 1), + (strip_tac 1), + (rtac mp 1), + (etac spec 1), + (etac less_trans 1), + (Simp_tac 1) + ]); trace_simp := false; val hoare_lemma24 = prove_goal Hoare.thy "(! m. m< Suc k --> b1`(iterate m g x)=TT) --> \ \ q`(iterate k g x)=q`x" (fn prems => - [ - (nat_ind_tac "k" 1), - (simp_tac (!simpset delsimps [less_Suc_eq]) 1), - (simp_tac (!simpset delsimps [less_Suc_eq]) 1), - (strip_tac 1), - (res_inst_tac [("s","q`(iterate k1 g x)")] trans 1), - (rtac trans 1), - (rtac (q_def3 RS sym) 2), - (res_inst_tac [("s","TT"),("t","b1`(iterate k1 g x)")] ssubst 1), - (rtac mp 1), - (etac spec 1), - (Simp_tac 1), - (simp_tac HOLCF_ss 1), - (etac mp 1), - (strip_tac 1), - (rtac mp 1), - (etac spec 1), - (etac less_trans 1), - (Simp_tac 1) - ]); + [ + (nat_ind_tac "k" 1), + (simp_tac (!simpset delsimps [less_Suc_eq]) 1), + (simp_tac (!simpset delsimps [less_Suc_eq]) 1), + (strip_tac 1), + (res_inst_tac [("s","q`(iterate k1 g x)")] trans 1), + (rtac trans 1), + (rtac (q_def3 RS sym) 2), + (res_inst_tac [("s","TT"),("t","b1`(iterate k1 g x)")] ssubst 1), + (rtac mp 1), + (etac spec 1), + (Simp_tac 1), + (simp_tac HOLCF_ss 1), + (etac mp 1), + (strip_tac 1), + (rtac mp 1), + (etac spec 1), + (etac less_trans 1), + (Simp_tac 1) + ]); (* -------- results about p for case (? k. b1`(iterate k g x)~=TT) ------- *) @@ -203,195 +203,195 @@ \ k=theleast(%n.b1`(iterate n g x) ~= TT) & b1`(iterate k g x)=FF \ \ --> p`x = iterate k g x" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("n","k")] natE 1), - (hyp_subst_tac 1), - (Simp_tac 1), - (strip_tac 1), - (etac conjE 1), - (rtac trans 1), - (rtac p_def3 1), - (asm_simp_tac HOLCF_ss 1), - (eres_inst_tac + [ + (cut_facts_tac prems 1), + (res_inst_tac [("n","k")] natE 1), + (hyp_subst_tac 1), + (Simp_tac 1), + (strip_tac 1), + (etac conjE 1), + (rtac trans 1), + (rtac p_def3 1), + (asm_simp_tac HOLCF_ss 1), + (eres_inst_tac [("s","0"),("t","theleast(%n. b1`(iterate n g x) ~= TT)")] subst 1), - (Simp_tac 1), - (hyp_subst_tac 1), - (strip_tac 1), - (etac conjE 1), - (rtac trans 1), - (etac (hoare_lemma10 RS sym) 1), - (atac 1), - (rtac trans 1), - (rtac p_def3 1), - (res_inst_tac [("s","TT"),("t","b1`(iterate xa g x)")] ssubst 1), - (rtac (hoare_lemma8 RS spec RS mp) 1), - (atac 1), - (atac 1), - (Simp_tac 1), - (simp_tac HOLCF_ss 1), - (rtac trans 1), - (rtac p_def3 1), - (simp_tac (!simpset delsimps [iterate_Suc] + (Simp_tac 1), + (hyp_subst_tac 1), + (strip_tac 1), + (etac conjE 1), + (rtac trans 1), + (etac (hoare_lemma10 RS sym) 1), + (atac 1), + (rtac trans 1), + (rtac p_def3 1), + (res_inst_tac [("s","TT"),("t","b1`(iterate xa g x)")] ssubst 1), + (rtac (hoare_lemma8 RS spec RS mp) 1), + (atac 1), + (atac 1), + (Simp_tac 1), + (simp_tac HOLCF_ss 1), + (rtac trans 1), + (rtac p_def3 1), + (simp_tac (!simpset delsimps [iterate_Suc] addsimps [iterate_Suc RS sym]) 1), - (eres_inst_tac [("s","FF")] ssubst 1), - (simp_tac HOLCF_ss 1) - ]); + (eres_inst_tac [("s","FF")] ssubst 1), + (simp_tac HOLCF_ss 1) + ]); val hoare_lemma12 = prove_goal Hoare.thy "(? n. b1`(iterate n g x) ~= TT) ==>\ \ k=theleast(%n. b1`(iterate n g x)~=TT) & b1`(iterate k g x)=UU \ \ --> p`x = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("n","k")] natE 1), - (hyp_subst_tac 1), - (Simp_tac 1), - (strip_tac 1), - (etac conjE 1), - (rtac trans 1), - (rtac p_def3 1), - (asm_simp_tac HOLCF_ss 1), - (hyp_subst_tac 1), - (Simp_tac 1), - (strip_tac 1), - (etac conjE 1), - (rtac trans 1), - (rtac (hoare_lemma10 RS sym) 1), - (atac 1), - (atac 1), - (rtac trans 1), - (rtac p_def3 1), - (res_inst_tac [("s","TT"),("t","b1`(iterate xa g x)")] ssubst 1), - (rtac (hoare_lemma8 RS spec RS mp) 1), - (atac 1), - (atac 1), - (Simp_tac 1), - (asm_simp_tac HOLCF_ss 1), - (rtac trans 1), - (rtac p_def3 1), - (asm_simp_tac HOLCF_ss 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("n","k")] natE 1), + (hyp_subst_tac 1), + (Simp_tac 1), + (strip_tac 1), + (etac conjE 1), + (rtac trans 1), + (rtac p_def3 1), + (asm_simp_tac HOLCF_ss 1), + (hyp_subst_tac 1), + (Simp_tac 1), + (strip_tac 1), + (etac conjE 1), + (rtac trans 1), + (rtac (hoare_lemma10 RS sym) 1), + (atac 1), + (atac 1), + (rtac trans 1), + (rtac p_def3 1), + (res_inst_tac [("s","TT"),("t","b1`(iterate xa g x)")] ssubst 1), + (rtac (hoare_lemma8 RS spec RS mp) 1), + (atac 1), + (atac 1), + (Simp_tac 1), + (asm_simp_tac HOLCF_ss 1), + (rtac trans 1), + (rtac p_def3 1), + (asm_simp_tac HOLCF_ss 1) + ]); (* -------- results about p for case (! k. b1`(iterate k g x)=TT) ------- *) val fernpass_lemma = prove_goal Hoare.thy "(! k. b1`(iterate k g x)=TT) ==> !k.p`(iterate k g x) = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (p_def2 RS ssubst) 1), - (rtac fix_ind 1), - (rtac adm_all 1), - (rtac allI 1), - (rtac adm_eq 1), - (cont_tacR 1), - (rtac allI 1), - (rtac (strict_fapp1 RS ssubst) 1), - (rtac refl 1), - (Simp_tac 1), - (rtac allI 1), - (res_inst_tac [("s","TT"),("t","b1`(iterate k g x)")] ssubst 1), - (etac spec 1), - (asm_simp_tac HOLCF_ss 1), - (rtac (iterate_Suc RS subst) 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (p_def2 RS ssubst) 1), + (rtac fix_ind 1), + (rtac adm_all 1), + (rtac allI 1), + (rtac adm_eq 1), + (cont_tacR 1), + (rtac allI 1), + (rtac (strict_fapp1 RS ssubst) 1), + (rtac refl 1), + (Simp_tac 1), + (rtac allI 1), + (res_inst_tac [("s","TT"),("t","b1`(iterate k g x)")] ssubst 1), + (etac spec 1), + (asm_simp_tac HOLCF_ss 1), + (rtac (iterate_Suc RS subst) 1), + (etac spec 1) + ]); val hoare_lemma16 = prove_goal Hoare.thy "(! k. b1`(iterate k g x)=TT) ==> p`x = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("F1","g"),("t","x")] (iterate_0 RS subst) 1), - (etac (fernpass_lemma RS spec) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("F1","g"),("t","x")] (iterate_0 RS subst) 1), + (etac (fernpass_lemma RS spec) 1) + ]); (* -------- results about q for case (! k. b1`(iterate k g x)=TT) ------- *) val hoare_lemma17 = prove_goal Hoare.thy "(! k. b1`(iterate k g x)=TT) ==> !k.q`(iterate k g x) = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (q_def2 RS ssubst) 1), - (rtac fix_ind 1), - (rtac adm_all 1), - (rtac allI 1), - (rtac adm_eq 1), - (cont_tacR 1), - (rtac allI 1), - (rtac (strict_fapp1 RS ssubst) 1), - (rtac refl 1), - (rtac allI 1), - (Simp_tac 1), - (res_inst_tac [("s","TT"),("t","b1`(iterate k g x)")] ssubst 1), - (etac spec 1), - (asm_simp_tac HOLCF_ss 1), - (rtac (iterate_Suc RS subst) 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (q_def2 RS ssubst) 1), + (rtac fix_ind 1), + (rtac adm_all 1), + (rtac allI 1), + (rtac adm_eq 1), + (cont_tacR 1), + (rtac allI 1), + (rtac (strict_fapp1 RS ssubst) 1), + (rtac refl 1), + (rtac allI 1), + (Simp_tac 1), + (res_inst_tac [("s","TT"),("t","b1`(iterate k g x)")] ssubst 1), + (etac spec 1), + (asm_simp_tac HOLCF_ss 1), + (rtac (iterate_Suc RS subst) 1), + (etac spec 1) + ]); val hoare_lemma18 = prove_goal Hoare.thy "(! k. b1`(iterate k g x)=TT) ==> q`x = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("F1","g"),("t","x")] (iterate_0 RS subst) 1), - (etac (hoare_lemma17 RS spec) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("F1","g"),("t","x")] (iterate_0 RS subst) 1), + (etac (hoare_lemma17 RS spec) 1) + ]); val hoare_lemma19 = prove_goal Hoare.thy "(!k. (b1::'a->tr)`(iterate k g x)=TT) ==> b1`(UU::'a) = UU | (!y.b1`(y::'a)=TT)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (flat_tr RS flat_codom) 1), - (res_inst_tac [("t","x1")] (iterate_0 RS subst) 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (flat_tr RS flat_codom) 1), + (res_inst_tac [("t","x1")] (iterate_0 RS subst) 1), + (etac spec 1) + ]); val hoare_lemma20 = prove_goal Hoare.thy "(! y. b1`(y::'a)=TT) ==> !k.q`(iterate k g (x::'a)) = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (q_def2 RS ssubst) 1), - (rtac fix_ind 1), - (rtac adm_all 1), - (rtac allI 1), - (rtac adm_eq 1), - (cont_tacR 1), - (rtac allI 1), - (rtac (strict_fapp1 RS ssubst) 1), - (rtac refl 1), - (rtac allI 1), - (Simp_tac 1), - (res_inst_tac [("s","TT"),("t","b1`(iterate k g (x::'a))")] ssubst 1), - (etac spec 1), - (asm_simp_tac HOLCF_ss 1), - (rtac (iterate_Suc RS subst) 1), - (etac spec 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (q_def2 RS ssubst) 1), + (rtac fix_ind 1), + (rtac adm_all 1), + (rtac allI 1), + (rtac adm_eq 1), + (cont_tacR 1), + (rtac allI 1), + (rtac (strict_fapp1 RS ssubst) 1), + (rtac refl 1), + (rtac allI 1), + (Simp_tac 1), + (res_inst_tac [("s","TT"),("t","b1`(iterate k g (x::'a))")] ssubst 1), + (etac spec 1), + (asm_simp_tac HOLCF_ss 1), + (rtac (iterate_Suc RS subst) 1), + (etac spec 1) + ]); val hoare_lemma21 = prove_goal Hoare.thy "(! y. b1`(y::'a)=TT) ==> q`(x::'a) = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("F1","g"),("t","x")] (iterate_0 RS subst) 1), - (etac (hoare_lemma20 RS spec) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("F1","g"),("t","x")] (iterate_0 RS subst) 1), + (etac (hoare_lemma20 RS spec) 1) + ]); val hoare_lemma22 = prove_goal Hoare.thy "b1`(UU::'a)=UU ==> q`(UU::'a) = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (q_def3 RS ssubst) 1), - (asm_simp_tac HOLCF_ss 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (q_def3 RS ssubst) 1), + (asm_simp_tac HOLCF_ss 1) + ]); (* -------- results about q for case (? k. b1`(iterate k g x) ~= TT) ------- *) @@ -407,28 +407,28 @@ \ k=theleast(%n. b1`(iterate n g x) ~= TT) & b1`(iterate k g x) =FF \ \ --> q`x = q`(iterate k g x)" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("n","k")] natE 1), - (hyp_subst_tac 1), - (strip_tac 1), - (Simp_tac 1), - (hyp_subst_tac 1), - (strip_tac 1), - (etac conjE 1), - (rtac trans 1), - (rtac (hoare_lemma25 RS sym) 1), - (atac 1), - (atac 1), - (rtac trans 1), - (rtac q_def3 1), - (res_inst_tac [("s","TT"),("t","b1`(iterate xa g x)")] ssubst 1), - (rtac (hoare_lemma8 RS spec RS mp) 1), - (atac 1), - (atac 1), - (Simp_tac 1), - (simp_tac HOLCF_ss 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("n","k")] natE 1), + (hyp_subst_tac 1), + (strip_tac 1), + (Simp_tac 1), + (hyp_subst_tac 1), + (strip_tac 1), + (etac conjE 1), + (rtac trans 1), + (rtac (hoare_lemma25 RS sym) 1), + (atac 1), + (atac 1), + (rtac trans 1), + (rtac q_def3 1), + (res_inst_tac [("s","TT"),("t","b1`(iterate xa g x)")] ssubst 1), + (rtac (hoare_lemma8 RS spec RS mp) 1), + (atac 1), + (atac 1), + (Simp_tac 1), + (simp_tac HOLCF_ss 1) + ]); val hoare_lemma27 = prove_goal Hoare.thy @@ -436,108 +436,108 @@ \ k=theleast(%n. b1`(iterate n g x)~=TT) & b1`(iterate k g x)=UU \ \ --> q`x = UU" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("n","k")] natE 1), - (hyp_subst_tac 1), - (Simp_tac 1), - (strip_tac 1), - (etac conjE 1), - (rtac (q_def3 RS ssubst) 1), - (asm_simp_tac HOLCF_ss 1), - (hyp_subst_tac 1), - (Simp_tac 1), - (strip_tac 1), - (etac conjE 1), - (rtac trans 1), - (rtac (hoare_lemma25 RS sym) 1), - (atac 1), - (atac 1), - (rtac trans 1), - (rtac q_def3 1), - (res_inst_tac [("s","TT"),("t","b1`(iterate xa g x)")] ssubst 1), - (rtac (hoare_lemma8 RS spec RS mp) 1), - (atac 1), - (atac 1), - (Simp_tac 1), - (asm_simp_tac HOLCF_ss 1), - (rtac trans 1), - (rtac q_def3 1), - (asm_simp_tac HOLCF_ss 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("n","k")] natE 1), + (hyp_subst_tac 1), + (Simp_tac 1), + (strip_tac 1), + (etac conjE 1), + (rtac (q_def3 RS ssubst) 1), + (asm_simp_tac HOLCF_ss 1), + (hyp_subst_tac 1), + (Simp_tac 1), + (strip_tac 1), + (etac conjE 1), + (rtac trans 1), + (rtac (hoare_lemma25 RS sym) 1), + (atac 1), + (atac 1), + (rtac trans 1), + (rtac q_def3 1), + (res_inst_tac [("s","TT"),("t","b1`(iterate xa g x)")] ssubst 1), + (rtac (hoare_lemma8 RS spec RS mp) 1), + (atac 1), + (atac 1), + (Simp_tac 1), + (asm_simp_tac HOLCF_ss 1), + (rtac trans 1), + (rtac q_def3 1), + (asm_simp_tac HOLCF_ss 1) + ]); (* ------- (! k. b1`(iterate k g x)=TT) ==> q o p = q ----- *) val hoare_lemma23 = prove_goal Hoare.thy "(! k. b1`(iterate k g x)=TT) ==> q`(p`x) = q`x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (hoare_lemma16 RS ssubst) 1), - (atac 1), - (rtac (hoare_lemma19 RS disjE) 1), - (atac 1), - (rtac (hoare_lemma18 RS ssubst) 1), - (atac 1), - (rtac (hoare_lemma22 RS ssubst) 1), - (atac 1), - (rtac refl 1), - (rtac (hoare_lemma21 RS ssubst) 1), - (atac 1), - (rtac (hoare_lemma21 RS ssubst) 1), - (atac 1), - (rtac refl 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (hoare_lemma16 RS ssubst) 1), + (atac 1), + (rtac (hoare_lemma19 RS disjE) 1), + (atac 1), + (rtac (hoare_lemma18 RS ssubst) 1), + (atac 1), + (rtac (hoare_lemma22 RS ssubst) 1), + (atac 1), + (rtac refl 1), + (rtac (hoare_lemma21 RS ssubst) 1), + (atac 1), + (rtac (hoare_lemma21 RS ssubst) 1), + (atac 1), + (rtac refl 1) + ]); (* ------------ ? k. b1~(iterate k g x) ~= TT ==> q o p = q ----- *) val hoare_lemma29 = prove_goal Hoare.thy "? k. b1`(iterate k g x) ~= TT ==> q`(p`x) = q`x" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (hoare_lemma5 RS disjE) 1), - (atac 1), - (rtac refl 1), - (rtac (hoare_lemma11 RS mp RS ssubst) 1), - (atac 1), - (rtac conjI 1), - (rtac refl 1), - (atac 1), - (rtac (hoare_lemma26 RS mp RS subst) 1), - (atac 1), - (rtac conjI 1), - (rtac refl 1), - (atac 1), - (rtac refl 1), - (rtac (hoare_lemma12 RS mp RS ssubst) 1), - (atac 1), - (rtac conjI 1), - (rtac refl 1), - (atac 1), - (rtac (hoare_lemma22 RS ssubst) 1), - (rtac (hoare_lemma28 RS ssubst) 1), - (atac 1), - (rtac refl 1), - (rtac sym 1), - (rtac (hoare_lemma27 RS mp RS ssubst) 1), - (atac 1), - (rtac conjI 1), - (rtac refl 1), - (atac 1), - (rtac refl 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (hoare_lemma5 RS disjE) 1), + (atac 1), + (rtac refl 1), + (rtac (hoare_lemma11 RS mp RS ssubst) 1), + (atac 1), + (rtac conjI 1), + (rtac refl 1), + (atac 1), + (rtac (hoare_lemma26 RS mp RS subst) 1), + (atac 1), + (rtac conjI 1), + (rtac refl 1), + (atac 1), + (rtac refl 1), + (rtac (hoare_lemma12 RS mp RS ssubst) 1), + (atac 1), + (rtac conjI 1), + (rtac refl 1), + (atac 1), + (rtac (hoare_lemma22 RS ssubst) 1), + (rtac (hoare_lemma28 RS ssubst) 1), + (atac 1), + (rtac refl 1), + (rtac sym 1), + (rtac (hoare_lemma27 RS mp RS ssubst) 1), + (atac 1), + (rtac conjI 1), + (rtac refl 1), + (atac 1), + (rtac refl 1) + ]); (* ------ the main prove q o p = q ------ *) val hoare_main = prove_goal Hoare.thy "q oo p = q" (fn prems => - [ - (rtac ext_cfun 1), - (rtac (cfcomp2 RS ssubst) 1), - (rtac (hoare_lemma3 RS disjE) 1), - (etac hoare_lemma23 1), - (etac hoare_lemma29 1) - ]); + [ + (rtac ext_cfun 1), + (rtac (cfcomp2 RS ssubst) 1), + (rtac (hoare_lemma3 RS disjE) 1), + (etac hoare_lemma23 1), + (etac hoare_lemma29 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/ex/Loop.ML --- a/src/HOLCF/ex/Loop.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/ex/Loop.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,7 +1,7 @@ -(* Title: HOLCF/ex/Loop.ML +(* Title: HOLCF/ex/Loop.ML ID: $Id$ - Author: Franz Regensburger - Copyright 1993 Technische Universitaet Muenchen + Author: Franz Regensburger + Copyright 1993 Technische Universitaet Muenchen Lemmas for theory loop.thy *) @@ -15,16 +15,16 @@ val step_def2 = prove_goalw Loop.thy [step_def] "step`b`g`x = If b`x then g`x else x fi" (fn prems => - [ - (Simp_tac 1) - ]); + [ + (Simp_tac 1) + ]); val while_def2 = prove_goalw Loop.thy [while_def] "while`b`g = fix`(LAM f x. If b`x then f`(g`x) else x fi)" (fn prems => - [ - (Simp_tac 1) - ]); + [ + (Simp_tac 1) + ]); (* ------------------------------------------------------------------------- *) @@ -34,48 +34,48 @@ val while_unfold = prove_goal Loop.thy "while`b`g`x = If b`x then while`b`g`(g`x) else x fi" (fn prems => - [ - (fix_tac5 while_def2 1), - (Simp_tac 1) - ]); + [ + (fix_tac5 while_def2 1), + (Simp_tac 1) + ]); val while_unfold2 = prove_goal Loop.thy - "!x.while`b`g`x = while`b`g`(iterate k (step`b`g) x)" + "!x.while`b`g`x = while`b`g`(iterate k (step`b`g) x)" (fn prems => - [ - (nat_ind_tac "k" 1), - (simp_tac HOLCF_ss 1), - (rtac allI 1), - (rtac trans 1), - (rtac (while_unfold RS ssubst) 1), - (rtac refl 2), - (rtac (iterate_Suc2 RS ssubst) 1), - (rtac trans 1), - (etac spec 2), - (rtac (step_def2 RS ssubst) 1), - (res_inst_tac [("p","b`x")] trE 1), - (asm_simp_tac HOLCF_ss 1), - (rtac (while_unfold RS ssubst) 1), - (res_inst_tac [("s","UU"),("t","b`UU")]ssubst 1), - (etac (flat_tr RS flat_codom RS disjE) 1), - (atac 1), - (etac spec 1), - (simp_tac HOLCF_ss 1), - (asm_simp_tac HOLCF_ss 1), - (asm_simp_tac HOLCF_ss 1), - (rtac (while_unfold RS ssubst) 1), - (asm_simp_tac HOLCF_ss 1) - ]); + [ + (nat_ind_tac "k" 1), + (simp_tac HOLCF_ss 1), + (rtac allI 1), + (rtac trans 1), + (rtac (while_unfold RS ssubst) 1), + (rtac refl 2), + (rtac (iterate_Suc2 RS ssubst) 1), + (rtac trans 1), + (etac spec 2), + (rtac (step_def2 RS ssubst) 1), + (res_inst_tac [("p","b`x")] trE 1), + (asm_simp_tac HOLCF_ss 1), + (rtac (while_unfold RS ssubst) 1), + (res_inst_tac [("s","UU"),("t","b`UU")]ssubst 1), + (etac (flat_tr RS flat_codom RS disjE) 1), + (atac 1), + (etac spec 1), + (simp_tac HOLCF_ss 1), + (asm_simp_tac HOLCF_ss 1), + (asm_simp_tac HOLCF_ss 1), + (rtac (while_unfold RS ssubst) 1), + (asm_simp_tac HOLCF_ss 1) + ]); val while_unfold3 = prove_goal Loop.thy - "while`b`g`x = while`b`g`(step`b`g`x)" + "while`b`g`x = while`b`g`(step`b`g`x)" (fn prems => - [ - (res_inst_tac [("s", - "while`b`g`(iterate (Suc 0) (step`b`g) x)")] trans 1), - (rtac (while_unfold2 RS spec) 1), - (Simp_tac 1) - ]); + [ + (res_inst_tac [("s", + "while`b`g`(iterate (Suc 0) (step`b`g) x)")] trans 1), + (rtac (while_unfold2 RS spec) 1), + (Simp_tac 1) + ]); (* --------------------------------------------------------------------------- *) @@ -85,130 +85,130 @@ val loop_lemma1 = prove_goal Loop.thy "[|? y.b`y=FF; iterate k (step`b`g) x = UU|]==>iterate(Suc k) (step`b`g) x=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (Simp_tac 1), - (rtac trans 1), - (rtac step_def2 1), - (asm_simp_tac HOLCF_ss 1), - (etac exE 1), - (etac (flat_tr RS flat_codom RS disjE) 1), - (asm_simp_tac HOLCF_ss 1), - (asm_simp_tac HOLCF_ss 1) - ]); + [ + (cut_facts_tac prems 1), + (Simp_tac 1), + (rtac trans 1), + (rtac step_def2 1), + (asm_simp_tac HOLCF_ss 1), + (etac exE 1), + (etac (flat_tr RS flat_codom RS disjE) 1), + (asm_simp_tac HOLCF_ss 1), + (asm_simp_tac HOLCF_ss 1) + ]); val loop_lemma2 = prove_goal Loop.thy "[|? y.b`y=FF;iterate (Suc k) (step`b`g) x ~=UU |]==>\ \iterate k (step`b`g) x ~=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac contrapos 1), - (etac loop_lemma1 2), - (atac 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac contrapos 1), + (etac loop_lemma1 2), + (atac 1), + (atac 1) + ]); val loop_lemma3 = prove_goal Loop.thy "[|!x. INV x & b`x=TT & g`x~=UU --> INV (g`x);\ \? y.b`y=FF; INV x|] ==>\ \iterate k (step`b`g) x ~=UU --> INV (iterate k (step`b`g) x)" (fn prems => - [ - (cut_facts_tac prems 1), - (nat_ind_tac "k" 1), - (Asm_simp_tac 1), - (strip_tac 1), - (simp_tac (!simpset addsimps [step_def2]) 1), - (res_inst_tac [("p","b`(iterate k1 (step`b`g) x)")] trE 1), - (etac notE 1), - (asm_simp_tac (HOLCF_ss addsimps [step_def2] ) 1), - (asm_simp_tac HOLCF_ss 1), - (rtac mp 1), - (etac spec 1), - (asm_simp_tac (HOLCF_ss delsimps [iterate_Suc] + [ + (cut_facts_tac prems 1), + (nat_ind_tac "k" 1), + (Asm_simp_tac 1), + (strip_tac 1), + (simp_tac (!simpset addsimps [step_def2]) 1), + (res_inst_tac [("p","b`(iterate k1 (step`b`g) x)")] trE 1), + (etac notE 1), + (asm_simp_tac (HOLCF_ss addsimps [step_def2] ) 1), + (asm_simp_tac HOLCF_ss 1), + (rtac mp 1), + (etac spec 1), + (asm_simp_tac (HOLCF_ss delsimps [iterate_Suc] addsimps [loop_lemma2] ) 1), - (res_inst_tac [("s","iterate (Suc k1) (step`b`g) x"), - ("t","g`(iterate k1 (step`b`g) x)")] ssubst 1), - (atac 2), - (asm_simp_tac (HOLCF_ss addsimps [step_def2] ) 1), - (asm_simp_tac (HOLCF_ss delsimps [iterate_Suc] + (res_inst_tac [("s","iterate (Suc k1) (step`b`g) x"), + ("t","g`(iterate k1 (step`b`g) x)")] ssubst 1), + (atac 2), + (asm_simp_tac (HOLCF_ss addsimps [step_def2] ) 1), + (asm_simp_tac (HOLCF_ss delsimps [iterate_Suc] addsimps [loop_lemma2] ) 1) - ]); + ]); val loop_lemma4 = prove_goal Loop.thy "!x. b`(iterate k (step`b`g) x)=FF --> while`b`g`x= iterate k (step`b`g) x" (fn prems => - [ - (nat_ind_tac "k" 1), - (Simp_tac 1), - (strip_tac 1), - (rtac (while_unfold RS ssubst) 1), - (asm_simp_tac HOLCF_ss 1), - (rtac allI 1), - (rtac (iterate_Suc2 RS ssubst) 1), - (strip_tac 1), - (rtac trans 1), - (rtac while_unfold3 1), - (asm_simp_tac HOLCF_ss 1) - ]); + [ + (nat_ind_tac "k" 1), + (Simp_tac 1), + (strip_tac 1), + (rtac (while_unfold RS ssubst) 1), + (asm_simp_tac HOLCF_ss 1), + (rtac allI 1), + (rtac (iterate_Suc2 RS ssubst) 1), + (strip_tac 1), + (rtac trans 1), + (rtac while_unfold3 1), + (asm_simp_tac HOLCF_ss 1) + ]); val loop_lemma5 = prove_goal Loop.thy "!k. b`(iterate k (step`b`g) x) ~= FF ==>\ \ !m. while`b`g`(iterate m (step`b`g) x)=UU" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (while_def2 RS ssubst) 1), - (rtac fix_ind 1), - (rtac (allI RS adm_all) 1), - (rtac adm_eq 1), - (cont_tacR 1), - (Simp_tac 1), - (rtac allI 1), - (Simp_tac 1), - (res_inst_tac [("p","b`(iterate m (step`b`g) x)")] trE 1), - (asm_simp_tac HOLCF_ss 1), - (asm_simp_tac HOLCF_ss 1), - (res_inst_tac [("s","xa`(iterate (Suc m) (step`b`g) x)")] trans 1), - (etac spec 2), - (rtac cfun_arg_cong 1), - (rtac trans 1), - (rtac (iterate_Suc RS sym) 2), - (asm_simp_tac (HOLCF_ss addsimps [step_def2]) 1), - (dtac spec 1), - (contr_tac 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (while_def2 RS ssubst) 1), + (rtac fix_ind 1), + (rtac (allI RS adm_all) 1), + (rtac adm_eq 1), + (cont_tacR 1), + (Simp_tac 1), + (rtac allI 1), + (Simp_tac 1), + (res_inst_tac [("p","b`(iterate m (step`b`g) x)")] trE 1), + (asm_simp_tac HOLCF_ss 1), + (asm_simp_tac HOLCF_ss 1), + (res_inst_tac [("s","xa`(iterate (Suc m) (step`b`g) x)")] trans 1), + (etac spec 2), + (rtac cfun_arg_cong 1), + (rtac trans 1), + (rtac (iterate_Suc RS sym) 2), + (asm_simp_tac (HOLCF_ss addsimps [step_def2]) 1), + (dtac spec 1), + (contr_tac 1) + ]); val loop_lemma6 = prove_goal Loop.thy "!k. b`(iterate k (step`b`g) x) ~= FF ==> while`b`g`x=UU" (fn prems => - [ - (res_inst_tac [("t","x")] (iterate_0 RS subst) 1), - (rtac (loop_lemma5 RS spec) 1), - (resolve_tac prems 1) - ]); + [ + (res_inst_tac [("t","x")] (iterate_0 RS subst) 1), + (rtac (loop_lemma5 RS spec) 1), + (resolve_tac prems 1) + ]); val loop_lemma7 = prove_goal Loop.thy "while`b`g`x ~= UU ==> ? k. b`(iterate k (step`b`g) x) = FF" (fn prems => - [ - (cut_facts_tac prems 1), - (etac swap 1), - (rtac loop_lemma6 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (etac swap 1), + (rtac loop_lemma6 1), + (fast_tac HOL_cs 1) + ]); val loop_lemma8 = prove_goal Loop.thy "while`b`g`x ~= UU ==> ? y. b`y=FF" (fn prems => - [ - (cut_facts_tac prems 1), - (dtac loop_lemma7 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (dtac loop_lemma7 1), + (fast_tac HOL_cs 1) + ]); (* ------------------------------------------------------------------------- *) @@ -220,49 +220,49 @@ \ (!y. INV y & b`y=FF --> Q y);\ \ INV x; while`b`g`x~=UU |] ==> Q (while`b`g`x)" (fn prems => - [ - (res_inst_tac [("P","%k. b`(iterate k (step`b`g) x)=FF")] exE 1), - (rtac loop_lemma7 1), - (resolve_tac prems 1), - (rtac ((loop_lemma4 RS spec RS mp) RS ssubst) 1), - (atac 1), - (rtac (nth_elem (1,prems) RS spec RS mp) 1), - (rtac conjI 1), - (atac 2), - (rtac (loop_lemma3 RS mp) 1), - (resolve_tac prems 1), - (rtac loop_lemma8 1), - (resolve_tac prems 1), - (resolve_tac prems 1), - (rtac classical3 1), - (resolve_tac prems 1), - (etac box_equals 1), - (rtac (loop_lemma4 RS spec RS mp RS sym) 1), - (atac 1), - (rtac refl 1) - ]); + [ + (res_inst_tac [("P","%k. b`(iterate k (step`b`g) x)=FF")] exE 1), + (rtac loop_lemma7 1), + (resolve_tac prems 1), + (rtac ((loop_lemma4 RS spec RS mp) RS ssubst) 1), + (atac 1), + (rtac (nth_elem (1,prems) RS spec RS mp) 1), + (rtac conjI 1), + (atac 2), + (rtac (loop_lemma3 RS mp) 1), + (resolve_tac prems 1), + (rtac loop_lemma8 1), + (resolve_tac prems 1), + (resolve_tac prems 1), + (rtac classical3 1), + (resolve_tac prems 1), + (etac box_equals 1), + (rtac (loop_lemma4 RS spec RS mp RS sym) 1), + (atac 1), + (rtac refl 1) + ]); val loop_inv3 = prove_goal Loop.thy "[| !!y.[| INV y; b`y=TT; g`y~=UU|] ==> INV (g`y);\ \ !!y.[| INV y; b`y=FF|]==> Q y;\ \ INV x; while`b`g`x~=UU |] ==> Q (while`b`g`x)" (fn prems => - [ - (rtac loop_inv2 1), - (rtac allI 1), - (rtac impI 1), - (resolve_tac prems 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1), - (rtac allI 1), - (rtac impI 1), - (resolve_tac prems 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); + [ + (rtac loop_inv2 1), + (rtac allI 1), + (rtac impI 1), + (resolve_tac prems 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1), + (rtac allI 1), + (rtac impI 1), + (resolve_tac prems 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); val loop_inv = prove_goal Loop.thy "[| P(x);\ @@ -271,15 +271,15 @@ \ !!y.[| INV y; b`y=FF|]==> Q y;\ \ while`b`g`x ~= UU |] ==> Q (while`b`g`x)" (fn prems => - [ - (rtac loop_inv3 1), - (eresolve_tac prems 1), - (atac 1), - (atac 1), - (resolve_tac prems 1), - (atac 1), - (atac 1), - (resolve_tac prems 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); + [ + (rtac loop_inv3 1), + (eresolve_tac prems 1), + (atac 1), + (atac 1), + (resolve_tac prems 1), + (atac 1), + (atac 1), + (resolve_tac prems 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/ex/ROOT.ML --- a/src/HOLCF/ex/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/ex/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/ex/ROOT +(* Title: HOLCF/ex/ROOT ID: $Id$ - Author: Tobias Nipkow + Author: Tobias Nipkow Copyright 1994 TU Muenchen Executes all examples for HOLCF. diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/ex/loeckx.ML --- a/src/HOLCF/ex/loeckx.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/ex/loeckx.ML Tue Jan 30 13:42:57 1996 +0100 @@ -27,7 +27,7 @@ Since we already proved the theorem val cont_lubcfun = prove_goal Cfun2.thy - "is_chain ?F ==> cont (%x. lub (range (%j. ?F j`x)))" + "is_chain ?F ==> cont (%x. lub (range (%j. ?F j`x)))" it suffices to prove: @@ -51,12 +51,12 @@ [("t","Ifix"),("s","(%f.lub(range(%j.(LAM f. iterate j f UU)`f)))")] ssubst 1); by (rtac ext 1); -by (rewrite_goals_tac [Ifix_def] ); +by (rewtac Ifix_def ); by (subgoal_tac "range(% i.iterate i f UU) = range(%j.(LAM f. iterate j f UU)`f)" 1); by (etac arg_cong 1); by (subgoal_tac - "(% i.iterate i f UU) = (%j.(LAM f. iterate j f UU)`f)" 1); + "(% i.iterate i f UU) = (%j.(LAM f. iterate j f UU)`f)" 1); by (etac arg_cong 1); by (rtac ext 1); by (rtac (beta_cfun RS ssubst) 1); @@ -90,7 +90,7 @@ val prems = goal Fix.thy " Ifix = (%f.lub(range(%j.(LAM f.iterate j f UU)`f)))"; by (rtac ext 1); -by (rewrite_goals_tac [Ifix_def] ); +by (rewtac Ifix_def ); by (rtac (fix_lemma1 RS ssubst) 1); by (rtac refl 1); val fix_lemma2 = result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/explicit_domains/Coind.ML --- a/src/HOLCF/explicit_domains/Coind.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/explicit_domains/Coind.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/Coind.ML +(* Title: HOLCF/Coind.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen *) @@ -12,10 +12,10 @@ val nats_def2 = fix_prover2 Coind.thy nats_def - "nats = scons`dzero`(smap`dsucc`nats)"; + "nats = scons`dzero`(smap`dsucc`nats)"; val from_def2 = fix_prover2 Coind.thy from_def - "from = (LAM n.scons`n`(from`(dsucc`n)))"; + "from = (LAM n.scons`n`(from`(dsucc`n)))"; @@ -26,25 +26,25 @@ val from = prove_goal Coind.thy "from`n = scons`n`(from`(dsucc`n))" (fn prems => - [ - (rtac trans 1), - (rtac (from_def2 RS ssubst) 1), - (Simp_tac 1), - (rtac refl 1) - ]); + [ + (rtac trans 1), + (rtac (from_def2 RS ssubst) 1), + (Simp_tac 1), + (rtac refl 1) + ]); val from1 = prove_goal Coind.thy "from`UU = UU" (fn prems => - [ - (rtac trans 1), - (rtac (from RS ssubst) 1), - (resolve_tac stream_constrdef 1), - (rtac refl 1) - ]); + [ + (rtac trans 1), + (rtac (from RS ssubst) 1), + (resolve_tac stream_constrdef 1), + (rtac refl 1) + ]); val coind_rews = - [iterator1, iterator2, iterator3, smap1, smap2,from1]; + [iterator1, iterator2, iterator3, smap1, smap2,from1]; (* ------------------------------------------------------------------------- *) @@ -54,85 +54,85 @@ val coind_lemma1 = prove_goal Coind.thy "iterator`n`(smap`dsucc)`nats =\ -\ scons`n`(iterator`(dsucc`n)`(smap`dsucc)`nats)" +\ scons`n`(iterator`(dsucc`n)`(smap`dsucc)`nats)" (fn prems => - [ - (res_inst_tac [("s","n")] dnat_ind 1), - (simp_tac (!simpset addsimps (coind_rews @ stream_rews)) 1), - (simp_tac (!simpset addsimps (coind_rews @ stream_rews)) 1), - (rtac trans 1), - (rtac nats_def2 1), - (simp_tac (!simpset addsimps (coind_rews @ dnat_rews)) 1), - (rtac trans 1), - (etac iterator3 1), - (rtac trans 1), - (Asm_simp_tac 1), - (rtac trans 1), - (etac smap2 1), - (rtac cfun_arg_cong 1), - (asm_simp_tac (!simpset addsimps ([iterator3 RS sym] @ dnat_rews)) 1) - ]); + [ + (res_inst_tac [("s","n")] dnat_ind 1), + (simp_tac (!simpset addsimps (coind_rews @ stream_rews)) 1), + (simp_tac (!simpset addsimps (coind_rews @ stream_rews)) 1), + (rtac trans 1), + (rtac nats_def2 1), + (simp_tac (!simpset addsimps (coind_rews @ dnat_rews)) 1), + (rtac trans 1), + (etac iterator3 1), + (rtac trans 1), + (Asm_simp_tac 1), + (rtac trans 1), + (etac smap2 1), + (rtac cfun_arg_cong 1), + (asm_simp_tac (!simpset addsimps ([iterator3 RS sym] @ dnat_rews)) 1) + ]); val nats_eq_from = prove_goal Coind.thy "nats = from`dzero" (fn prems => - [ - (res_inst_tac [("R", + [ + (res_inst_tac [("R", "% p q.? n. p = iterator`n`(smap`dsucc)`nats & q = from`n")] stream_coind 1), - (res_inst_tac [("x","dzero")] exI 2), - (asm_simp_tac (!simpset addsimps coind_rews) 2), - (rewrite_goals_tac [stream_bisim_def]), - (strip_tac 1), - (etac exE 1), - (res_inst_tac [("Q","n=UU")] classical2 1), - (rtac disjI1 1), - (asm_simp_tac (!simpset addsimps coind_rews) 1), - (rtac disjI2 1), - (etac conjE 1), - (hyp_subst_tac 1), - (res_inst_tac [("x","n")] exI 1), - (res_inst_tac [("x","iterator`(dsucc`n)`(smap`dsucc)`nats")] exI 1), - (res_inst_tac [("x","from`(dsucc`n)")] exI 1), - (etac conjI 1), - (rtac conjI 1), - (rtac coind_lemma1 1), - (rtac conjI 1), - (rtac from 1), - (res_inst_tac [("x","dsucc`n")] exI 1), - (fast_tac HOL_cs 1) - ]); + (res_inst_tac [("x","dzero")] exI 2), + (asm_simp_tac (!simpset addsimps coind_rews) 2), + (rewtac stream_bisim_def), + (strip_tac 1), + (etac exE 1), + (res_inst_tac [("Q","n=UU")] classical2 1), + (rtac disjI1 1), + (asm_simp_tac (!simpset addsimps coind_rews) 1), + (rtac disjI2 1), + (etac conjE 1), + (hyp_subst_tac 1), + (res_inst_tac [("x","n")] exI 1), + (res_inst_tac [("x","iterator`(dsucc`n)`(smap`dsucc)`nats")] exI 1), + (res_inst_tac [("x","from`(dsucc`n)")] exI 1), + (etac conjI 1), + (rtac conjI 1), + (rtac coind_lemma1 1), + (rtac conjI 1), + (rtac from 1), + (res_inst_tac [("x","dsucc`n")] exI 1), + (fast_tac HOL_cs 1) + ]); (* another proof using stream_coind_lemma2 *) val nats_eq_from = prove_goal Coind.thy "nats = from`dzero" (fn prems => - [ - (res_inst_tac [("R","% p q.? n. p = \ -\ iterator`n`(smap`dsucc)`nats & q = from`n")] stream_coind 1), - (rtac stream_coind_lemma2 1), - (strip_tac 1), - (etac exE 1), - (res_inst_tac [("Q","n=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps coind_rews) 1), - (res_inst_tac [("x","UU::dnat")] exI 1), - (simp_tac (!simpset addsimps coind_rews addsimps stream_rews) 1), - (etac conjE 1), - (hyp_subst_tac 1), - (rtac conjI 1), - (rtac (coind_lemma1 RS ssubst) 1), - (rtac (from RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (res_inst_tac [("x","dsucc`n")] exI 1), - (rtac conjI 1), - (rtac trans 1), - (rtac (coind_lemma1 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (rtac refl 1), - (rtac trans 1), - (rtac (from RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (rtac refl 1), - (res_inst_tac [("x","dzero")] exI 1), - (asm_simp_tac (!simpset addsimps coind_rews) 1) - ]); + [ + (res_inst_tac [("R","% p q.? n. p = \ +\ iterator`n`(smap`dsucc)`nats & q = from`n")] stream_coind 1), + (rtac stream_coind_lemma2 1), + (strip_tac 1), + (etac exE 1), + (res_inst_tac [("Q","n=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps coind_rews) 1), + (res_inst_tac [("x","UU::dnat")] exI 1), + (simp_tac (!simpset addsimps coind_rews addsimps stream_rews) 1), + (etac conjE 1), + (hyp_subst_tac 1), + (rtac conjI 1), + (rtac (coind_lemma1 RS ssubst) 1), + (rtac (from RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (res_inst_tac [("x","dsucc`n")] exI 1), + (rtac conjI 1), + (rtac trans 1), + (rtac (coind_lemma1 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (rtac refl 1), + (rtac trans 1), + (rtac (from RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (rtac refl 1), + (res_inst_tac [("x","dzero")] exI 1), + (asm_simp_tac (!simpset addsimps coind_rews) 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/explicit_domains/Dagstuhl.ML --- a/src/HOLCF/explicit_domains/Dagstuhl.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/explicit_domains/Dagstuhl.ML Tue Jan 30 13:42:57 1996 +0100 @@ -7,7 +7,7 @@ val prems = goal Dagstuhl.thy "YYS << scons`y`YYS"; -by (rewrite_goals_tac [YYS_def]); +by (rewtac YYS_def); by (rtac fix_ind 1); by (resolve_tac adm_thms 1); by (cont_tacR 1); @@ -52,7 +52,7 @@ (* ------------------------------------------------------------------------ *) val prems = goal Dagstuhl.thy "YYS << YS"; -by (rewrite_goals_tac [YYS_def]); +by (rewtac YYS_def); by (rtac fix_least 1); by (rtac (beta_cfun RS ssubst) 1); by (cont_tacR 1); @@ -60,7 +60,7 @@ val lemma6=result(); val prems = goal Dagstuhl.thy "YS << YYS"; -by (rewrite_goals_tac [YS_def]); +by (rewtac YS_def); by (rtac fix_ind 1); by (resolve_tac adm_thms 1); by (cont_tacR 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/explicit_domains/Dlist.ML --- a/src/HOLCF/explicit_domains/Dlist.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/explicit_domains/Dlist.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,5 +1,5 @@ -(* Title: HOLCF/Dlist.ML - Author: Franz Regensburger +(* Title: HOLCF/Dlist.ML + Author: Franz Regensburger ID: $ $ Copyright 1994 Technische Universitaet Muenchen @@ -13,10 +13,10 @@ (* ------------------------------------------------------------------------*) val dlist_iso_strict= dlist_rep_iso RS (dlist_abs_iso RS - (allI RSN (2,allI RS iso_strict))); + (allI RSN (2,allI RS iso_strict))); val dlist_rews = [dlist_iso_strict RS conjunct1, - dlist_iso_strict RS conjunct2]; + dlist_iso_strict RS conjunct2]; (* ------------------------------------------------------------------------*) (* Properties of dlist_copy *) @@ -24,10 +24,10 @@ val temp = prove_goalw Dlist.thy [dlist_copy_def] "dlist_copy`f`UU=UU" (fn prems => - [ - (asm_simp_tac (!simpset addsimps - (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1) - ]); + [ + (asm_simp_tac (!simpset addsimps + (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1) + ]); val dlist_copy = [temp]; @@ -35,25 +35,25 @@ val temp = prove_goalw Dlist.thy [dlist_copy_def,dnil_def] "dlist_copy`f`dnil=dnil" (fn prems => - [ - (asm_simp_tac (!simpset addsimps - (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1) - ]); + [ + (asm_simp_tac (!simpset addsimps + (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1) + ]); val dlist_copy = temp :: dlist_copy; val temp = prove_goalw Dlist.thy [dlist_copy_def,dcons_def] - "xl~=UU ==> dlist_copy`f`(dcons`x`xl)= dcons`x`(f`xl)" + "xl~=UU ==> dlist_copy`f`(dcons`x`xl)= dcons`x`(f`xl)" (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps - (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1), - (res_inst_tac [("Q","x=UU")] classical2 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps [defined_spair]) 1) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps + (dlist_rews @ [dlist_abs_iso,dlist_rep_iso])) 1), + (res_inst_tac [("Q","x=UU")] classical2 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps [defined_spair]) 1) + ]); val dlist_copy = temp :: dlist_copy; @@ -64,45 +64,45 @@ (* ------------------------------------------------------------------------*) qed_goalw "Exh_dlist" Dlist.thy [dcons_def,dnil_def] - "l = UU | l = dnil | (? x xl. x~=UU &xl~=UU & l = dcons`x`xl)" + "l = UU | l = dnil | (? x xl. x~=UU &xl~=UU & l = dcons`x`xl)" (fn prems => - [ - (Simp_tac 1), - (rtac (dlist_rep_iso RS subst) 1), - (res_inst_tac [("p","dlist_rep`l")] ssumE 1), - (rtac disjI1 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (rtac disjI2 1), - (rtac disjI1 1), - (res_inst_tac [("p","x")] oneE 1), - (contr_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (rtac disjI2 1), - (rtac disjI2 1), - (res_inst_tac [("p","y")] sprodE 1), - (contr_tac 1), - (rtac exI 1), - (rtac exI 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (fast_tac HOL_cs 1) - ]); + [ + (Simp_tac 1), + (rtac (dlist_rep_iso RS subst) 1), + (res_inst_tac [("p","dlist_rep`l")] ssumE 1), + (rtac disjI1 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (rtac disjI2 1), + (rtac disjI1 1), + (res_inst_tac [("p","x")] oneE 1), + (contr_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (rtac disjI2 1), + (rtac disjI2 1), + (res_inst_tac [("p","y")] sprodE 1), + (contr_tac 1), + (rtac exI 1), + (rtac exI 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (fast_tac HOL_cs 1) + ]); qed_goal "dlistE" Dlist.thy "[| l=UU ==> Q; l=dnil ==> Q;!!x xl.[|l=dcons`x`xl;x~=UU;xl~=UU|]==>Q|]==>Q" (fn prems => - [ - (rtac (Exh_dlist RS disjE) 1), - (eresolve_tac prems 1), - (etac disjE 1), - (eresolve_tac prems 1), - (etac exE 1), - (etac exE 1), - (resolve_tac prems 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); + [ + (rtac (Exh_dlist RS disjE) 1), + (eresolve_tac prems 1), + (etac disjE 1), + (eresolve_tac prems 1), + (etac exE 1), + (etac exE 1), + (resolve_tac prems 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); (* ------------------------------------------------------------------------*) (* Properties of dlist_when *) @@ -110,28 +110,28 @@ val temp = prove_goalw Dlist.thy [dlist_when_def] "dlist_when`f1`f2`UU=UU" (fn prems => - [ - (asm_simp_tac (!simpset addsimps [dlist_iso_strict]) 1) - ]); + [ + (asm_simp_tac (!simpset addsimps [dlist_iso_strict]) 1) + ]); val dlist_when = [temp]; val temp = prove_goalw Dlist.thy [dlist_when_def,dnil_def] "dlist_when`f1`f2`dnil= f1" (fn prems => - [ - (asm_simp_tac (!simpset addsimps [dlist_abs_iso]) 1) - ]); + [ + (asm_simp_tac (!simpset addsimps [dlist_abs_iso]) 1) + ]); val dlist_when = temp::dlist_when; val temp = prove_goalw Dlist.thy [dlist_when_def,dcons_def] "[|x~=UU;xl~=UU|] ==> dlist_when`f1`f2`(dcons`x`xl)= f2`x`xl" (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps [dlist_abs_iso,defined_spair]) 1) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps [dlist_abs_iso,defined_spair]) 1) + ]); val dlist_when = temp::dlist_when; @@ -143,23 +143,23 @@ fun prover defs thm = prove_goalw Dlist.thy defs thm (fn prems => - [ - (simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (simp_tac (!simpset addsimps dlist_rews) 1) + ]); val dlist_discsel = [ - prover [is_dnil_def] "is_dnil`UU=UU", - prover [is_dcons_def] "is_dcons`UU=UU", - prover [dhd_def] "dhd`UU=UU", - prover [dtl_def] "dtl`UU=UU" - ]; + prover [is_dnil_def] "is_dnil`UU=UU", + prover [is_dcons_def] "is_dcons`UU=UU", + prover [dhd_def] "dhd`UU=UU", + prover [dtl_def] "dtl`UU=UU" + ]; fun prover defs thm = prove_goalw Dlist.thy defs thm (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); val dlist_discsel = [ prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] @@ -187,32 +187,32 @@ fun prover contr thm = prove_goal Dlist.thy thm (fn prems => - [ - (res_inst_tac [("P1",contr)] classical3 1), - (simp_tac (!simpset addsimps dlist_rews) 1), - (dtac sym 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps (prems @ dlist_rews)) 1) - ]); + [ + (res_inst_tac [("P1",contr)] classical3 1), + (simp_tac (!simpset addsimps dlist_rews) 1), + (dtac sym 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps (prems @ dlist_rews)) 1) + ]); val dlist_constrdef = [ prover "is_dnil`(UU::'a dlist) ~= UU" "dnil~=(UU::'a dlist)", prover "is_dcons`(UU::'a dlist) ~= UU" - "[|x~=UU;xl~=UU|]==>dcons`(x::'a)`xl ~= UU" + "[|x~=UU;xl~=UU|]==>dcons`(x::'a)`xl ~= UU" ]; fun prover defs thm = prove_goalw Dlist.thy defs thm (fn prems => - [ - (simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (simp_tac (!simpset addsimps dlist_rews) 1) + ]); val dlist_constrdef = [ - prover [dcons_def] "dcons`UU`xl=UU", - prover [dcons_def] "dcons`x`UU=UU" - ] @ dlist_constrdef; + prover [dcons_def] "dcons`UU`xl=UU", + prover [dcons_def] "dcons`x`UU=UU" + ] @ dlist_constrdef; val dlist_rews = dlist_constrdef @ dlist_rews; @@ -223,49 +223,49 @@ val temp = prove_goal Dlist.thy "~dnil << dcons`(x::'a)`xl" (fn prems => - [ - (res_inst_tac [("P1","TT << FF")] classical3 1), - (resolve_tac dist_less_tr 1), - (dres_inst_tac [("fo5","is_dnil")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("Q","(x::'a)=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("Q","(xl ::'a dlist)=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (res_inst_tac [("P1","TT << FF")] classical3 1), + (resolve_tac dist_less_tr 1), + (dres_inst_tac [("fo5","is_dnil")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("Q","(x::'a)=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("Q","(xl ::'a dlist)=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); val dlist_dist_less = [temp]; val temp = prove_goal Dlist.thy "[|x~=UU;xl~=UU|]==>~ dcons`x`xl << dnil" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("P1","TT << FF")] classical3 1), - (resolve_tac dist_less_tr 1), - (dres_inst_tac [("fo5","is_dcons")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("P1","TT << FF")] classical3 1), + (resolve_tac dist_less_tr 1), + (dres_inst_tac [("fo5","is_dcons")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); val dlist_dist_less = temp::dlist_dist_less; val temp = prove_goal Dlist.thy "dnil ~= dcons`x`xl" (fn prems => - [ - (res_inst_tac [("Q","x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("Q","xl=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("P1","TT = FF")] classical3 1), - (resolve_tac dist_eq_tr 1), - (dres_inst_tac [("f","is_dnil")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (res_inst_tac [("Q","x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("Q","xl=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("P1","TT = FF")] classical3 1), + (resolve_tac dist_eq_tr 1), + (dres_inst_tac [("f","is_dnil")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); val dlist_dist_eq = [temp,temp RS not_sym]; @@ -278,18 +278,18 @@ val temp = prove_goal Dlist.thy "[|x1~=UU; y1~=UU;x2~=UU; y2~=UU;\ \ dcons`x1`x2 << dcons`y1`y2 |] ==> x1<< y1 & x2 << y2" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - (dres_inst_tac [("fo5","dlist_when`UU`(LAM x l.x)")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (dres_inst_tac [("fo5","dlist_when`UU`(LAM x l.l)")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac conjI 1), + (dres_inst_tac [("fo5","dlist_when`UU`(LAM x l.x)")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (dres_inst_tac [("fo5","dlist_when`UU`(LAM x l.l)")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1) + ]); val dlist_invert =[temp]; @@ -300,18 +300,18 @@ val temp = prove_goal Dlist.thy "[|x1~=UU; y1~=UU;x2~=UU; y2~=UU;\ \ dcons`x1`x2 = dcons`y1`y2 |] ==> x1= y1 & x2 = y2" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - (dres_inst_tac [("f","dlist_when`UU`(LAM x l.x)")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (dres_inst_tac [("f","dlist_when`UU`(LAM x l.l)")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1), - (asm_simp_tac (!simpset addsimps dlist_when) 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac conjI 1), + (dres_inst_tac [("f","dlist_when`UU`(LAM x l.x)")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (dres_inst_tac [("f","dlist_when`UU`(LAM x l.l)")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1), + (asm_simp_tac (!simpset addsimps dlist_when) 1) + ]); val dlist_inject = [temp]; @@ -322,18 +322,18 @@ fun prover thm = prove_goal Dlist.thy thm (fn prems => - [ - (cut_facts_tac prems 1), - (rtac dlistE 1), - (contr_tac 1), - (REPEAT (asm_simp_tac (!simpset addsimps dlist_discsel) 1)) - ]); + [ + (cut_facts_tac prems 1), + (rtac dlistE 1), + (contr_tac 1), + (REPEAT (asm_simp_tac (!simpset addsimps dlist_discsel) 1)) + ]); val dlist_discsel_def = - [ - prover "l~=UU ==> is_dnil`l~=UU", - prover "l~=UU ==> is_dcons`l~=UU" - ]; + [ + prover "l~=UU ==> is_dnil`l~=UU", + prover "l~=UU ==> is_dcons`l~=UU" + ]; val dlist_rews = dlist_discsel_def @ dlist_rews; @@ -343,21 +343,21 @@ qed_goal "dhd2" Dlist.thy "xl~=UU ==> dhd`(dcons`x`xl)=x" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("Q","x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("Q","x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); qed_goal "dtl2" Dlist.thy "x~=UU ==> dtl`(dcons`x`xl)=xl" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("Q","xl=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("Q","xl=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); val dlist_rews = dhd2 :: dtl2 :: dlist_rews; @@ -367,52 +367,52 @@ val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take n`UU=UU" (fn prems => - [ - (res_inst_tac [("n","n")] natE 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (res_inst_tac [("n","n")] natE 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps dlist_rews) 1) + ]); val dlist_take = [temp]; val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take 0`xs=UU" (fn prems => - [ - (Asm_simp_tac 1) - ]); + [ + (Asm_simp_tac 1) + ]); val dlist_take = temp::dlist_take; val temp = prove_goalw Dlist.thy [dlist_take_def] - "dlist_take (Suc n)`dnil=dnil" + "dlist_take (Suc n)`dnil=dnil" (fn prems => - [ - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps dlist_rews) 1) + ]); val dlist_take = temp::dlist_take; val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take (Suc n)`(dcons`x`xl)= dcons`x`(dlist_take n`xl)" (fn prems => - [ - (res_inst_tac [("Q","x=UU")] classical2 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("Q","xl=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("n","n")] natE 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (res_inst_tac [("Q","x=UU")] classical2 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("Q","xl=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("n","n")] natE 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); val dlist_take = temp::dlist_take; @@ -424,23 +424,23 @@ fun prover reach defs thm = prove_goalw Dlist.thy defs thm (fn prems => - [ - (res_inst_tac [("t","l1")] (reach RS subst) 1), - (res_inst_tac [("t","l2")] (reach RS subst) 1), - (rtac (fix_def2 RS ssubst) 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac lub_equal 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac allI 1), - (resolve_tac prems 1) - ]); + [ + (res_inst_tac [("t","l1")] (reach RS subst) 1), + (res_inst_tac [("t","l2")] (reach RS subst) 1), + (rtac (fix_def2 RS ssubst) 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac lub_equal 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac allI 1), + (resolve_tac prems 1) + ]); val dlist_take_lemma = prover dlist_reach [dlist_take_def] - "(!!n.dlist_take n`l1 = dlist_take n`l2) ==> l1=l2"; + "(!!n.dlist_take n`l1 = dlist_take n`l2) ==> l1=l2"; (* ------------------------------------------------------------------------*) @@ -450,33 +450,33 @@ qed_goalw "dlist_coind_lemma" Dlist.thy [dlist_bisim_def] "dlist_bisim R ==> ! p q. R p q --> dlist_take n`p = dlist_take n`q" (fn prems => - [ - (cut_facts_tac prems 1), - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dlist_rews) 1), - (strip_tac 1), - ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), - (atac 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (etac disjE 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (etac exE 1), - (etac exE 1), - (etac exE 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (REPEAT (etac conjE 1)), - (rtac cfun_arg_cong 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dlist_rews) 1), + (strip_tac 1), + ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), + (atac 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (etac disjE 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (etac exE 1), + (etac exE 1), + (etac exE 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (REPEAT (etac conjE 1)), + (rtac cfun_arg_cong 1), + (fast_tac HOL_cs 1) + ]); qed_goal "dlist_coind" Dlist.thy "[|dlist_bisim R ; R p q |] ==> p = q" (fn prems => - [ - (rtac dlist_take_lemma 1), - (rtac (dlist_coind_lemma RS spec RS spec RS mp) 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); + [ + (rtac dlist_take_lemma 1), + (rtac (dlist_coind_lemma RS spec RS spec RS mp) 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); (* ------------------------------------------------------------------------*) (* structural induction *) @@ -487,77 +487,77 @@ \ !! x l1.[|x~=UU;l1~=UU;P(l1)|] ==> P(dcons`x`l1)\ \ |] ==> !l.P(dlist_take n`l)" (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dlist_rews) 1), - (resolve_tac prems 1), - (rtac allI 1), - (res_inst_tac [("l","l")] dlistE 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (res_inst_tac [("Q","dlist_take n1`xl=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (resolve_tac prems 1), - (resolve_tac prems 1), - (atac 1), - (atac 1), - (etac spec 1) - ]); + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dlist_rews) 1), + (resolve_tac prems 1), + (rtac allI 1), + (res_inst_tac [("l","l")] dlistE 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (res_inst_tac [("Q","dlist_take n1`xl=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (resolve_tac prems 1), + (resolve_tac prems 1), + (atac 1), + (atac 1), + (etac spec 1) + ]); qed_goal "dlist_all_finite_lemma1" Dlist.thy "!l.dlist_take n`l=UU |dlist_take n`l=l" (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dlist_rews) 1), - (rtac allI 1), - (res_inst_tac [("l","l")] dlistE 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (eres_inst_tac [("x","xl")] allE 1), - (etac disjE 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1) - ]); + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dlist_rews) 1), + (rtac allI 1), + (res_inst_tac [("l","l")] dlistE 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (eres_inst_tac [("x","xl")] allE 1), + (etac disjE 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1) + ]); qed_goal "dlist_all_finite_lemma2" Dlist.thy "? n.dlist_take n`l=l" (fn prems => - [ - (res_inst_tac [("Q","l=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (subgoal_tac "(!n.dlist_take n`l=UU) |(? n.dlist_take n`l = l)" 1), - (etac disjE 1), - (eres_inst_tac [("P","l=UU")] notE 1), - (rtac dlist_take_lemma 1), - (asm_simp_tac (!simpset addsimps dlist_rews) 1), - (atac 1), - (subgoal_tac "!n.!l.dlist_take n`l=UU |dlist_take n`l=l" 1), - (fast_tac HOL_cs 1), - (rtac allI 1), - (rtac dlist_all_finite_lemma1 1) - ]); + [ + (res_inst_tac [("Q","l=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (subgoal_tac "(!n.dlist_take n`l=UU) |(? n.dlist_take n`l = l)" 1), + (etac disjE 1), + (eres_inst_tac [("P","l=UU")] notE 1), + (rtac dlist_take_lemma 1), + (asm_simp_tac (!simpset addsimps dlist_rews) 1), + (atac 1), + (subgoal_tac "!n.!l.dlist_take n`l=UU |dlist_take n`l=l" 1), + (fast_tac HOL_cs 1), + (rtac allI 1), + (rtac dlist_all_finite_lemma1 1) + ]); qed_goalw "dlist_all_finite" Dlist.thy [dlist_finite_def] "dlist_finite(l)" (fn prems => - [ - (rtac dlist_all_finite_lemma2 1) - ]); + [ + (rtac dlist_all_finite_lemma2 1) + ]); qed_goal "dlist_ind" Dlist.thy "[|P(UU);P(dnil);\ \ !! x l1.[|x~=UU;l1~=UU;P(l1)|] ==> P(dcons`x`l1)|] ==> P(l)" (fn prems => - [ - (rtac (dlist_all_finite_lemma2 RS exE) 1), - (etac subst 1), - (rtac (dlist_finite_ind RS spec) 1), - (REPEAT (resolve_tac prems 1)), - (REPEAT (atac 1)) - ]); + [ + (rtac (dlist_all_finite_lemma2 RS exE) 1), + (etac subst 1), + (rtac (dlist_finite_ind RS spec) 1), + (REPEAT (resolve_tac prems 1)), + (REPEAT (atac 1)) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/explicit_domains/Dnat.ML --- a/src/HOLCF/explicit_domains/Dnat.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/explicit_domains/Dnat.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: HOLCF/Dnat.ML +(* Title: HOLCF/Dnat.ML ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for dnat.thy @@ -13,10 +13,10 @@ (* ------------------------------------------------------------------------*) val dnat_iso_strict = dnat_rep_iso RS (dnat_abs_iso RS - (allI RSN (2,allI RS iso_strict))); + (allI RSN (2,allI RS iso_strict))); val dnat_rews = [dnat_iso_strict RS conjunct1, - dnat_iso_strict RS conjunct2]; + dnat_iso_strict RS conjunct2]; (* ------------------------------------------------------------------------*) (* Properties of dnat_copy *) @@ -24,19 +24,19 @@ fun prover defs thm = prove_goalw Dnat.thy defs thm (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps - (dnat_rews @ [dnat_abs_iso,dnat_rep_iso])) 1) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps + (dnat_rews @ [dnat_abs_iso,dnat_rep_iso])) 1) + ]); val dnat_copy = - [ - prover [dnat_copy_def] "dnat_copy`f`UU=UU", - prover [dnat_copy_def,dzero_def] "dnat_copy`f`dzero= dzero", - prover [dnat_copy_def,dsucc_def] - "n~=UU ==> dnat_copy`f`(dsucc`n) = dsucc`(f`n)" - ]; + [ + prover [dnat_copy_def] "dnat_copy`f`UU=UU", + prover [dnat_copy_def,dzero_def] "dnat_copy`f`dzero= dzero", + prover [dnat_copy_def,dsucc_def] + "n~=UU ==> dnat_copy`f`(dsucc`n) = dsucc`(f`n)" + ]; val dnat_rews = dnat_copy @ dnat_rews; @@ -45,37 +45,37 @@ (* ------------------------------------------------------------------------*) qed_goalw "Exh_dnat" Dnat.thy [dsucc_def,dzero_def] - "n = UU | n = dzero | (? x . x~=UU & n = dsucc`x)" + "n = UU | n = dzero | (? x . x~=UU & n = dsucc`x)" (fn prems => - [ - (Simp_tac 1), - (rtac (dnat_rep_iso RS subst) 1), - (res_inst_tac [("p","dnat_rep`n")] ssumE 1), - (rtac disjI1 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (rtac (disjI1 RS disjI2) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (res_inst_tac [("p","x")] oneE 1), - (contr_tac 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (rtac (disjI2 RS disjI2) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (fast_tac HOL_cs 1) - ]); + [ + (Simp_tac 1), + (rtac (dnat_rep_iso RS subst) 1), + (res_inst_tac [("p","dnat_rep`n")] ssumE 1), + (rtac disjI1 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (rtac (disjI1 RS disjI2) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (res_inst_tac [("p","x")] oneE 1), + (contr_tac 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (rtac (disjI2 RS disjI2) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (fast_tac HOL_cs 1) + ]); qed_goal "dnatE" Dnat.thy "[| n=UU ==> Q; n=dzero ==> Q; !!x.[|n=dsucc`x;x~=UU|]==>Q|]==>Q" (fn prems => - [ - (rtac (Exh_dnat RS disjE) 1), - (eresolve_tac prems 1), - (etac disjE 1), - (eresolve_tac prems 1), - (REPEAT (etac exE 1)), - (resolve_tac prems 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); + [ + (rtac (Exh_dnat RS disjE) 1), + (eresolve_tac prems 1), + (etac disjE 1), + (eresolve_tac prems 1), + (REPEAT (etac exE 1)), + (resolve_tac prems 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); (* ------------------------------------------------------------------------*) (* Properties of dnat_when *) @@ -83,19 +83,19 @@ fun prover defs thm = prove_goalw Dnat.thy defs thm (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps - (dnat_rews @ [dnat_abs_iso,dnat_rep_iso])) 1) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps + (dnat_rews @ [dnat_abs_iso,dnat_rep_iso])) 1) + ]); val dnat_when = [ - prover [dnat_when_def] "dnat_when`c`f`UU=UU", - prover [dnat_when_def,dzero_def] "dnat_when`c`f`dzero=c", - prover [dnat_when_def,dsucc_def] - "n~=UU ==> dnat_when`c`f`(dsucc`n)=f`n" - ]; + prover [dnat_when_def] "dnat_when`c`f`UU=UU", + prover [dnat_when_def,dzero_def] "dnat_when`c`f`dzero=c", + prover [dnat_when_def,dsucc_def] + "n~=UU ==> dnat_when`c`f`(dsucc`n)=f`n" + ]; val dnat_rews = dnat_when @ dnat_rews; @@ -105,32 +105,32 @@ fun prover defs thm = prove_goalw Dnat.thy defs thm (fn prems => - [ - (simp_tac (!simpset addsimps dnat_rews) 1) - ]); + [ + (simp_tac (!simpset addsimps dnat_rews) 1) + ]); val dnat_discsel = [ - prover [is_dzero_def] "is_dzero`UU=UU", - prover [is_dsucc_def] "is_dsucc`UU=UU", - prover [dpred_def] "dpred`UU=UU" - ]; + prover [is_dzero_def] "is_dzero`UU=UU", + prover [is_dsucc_def] "is_dsucc`UU=UU", + prover [dpred_def] "dpred`UU=UU" + ]; fun prover defs thm = prove_goalw Dnat.thy defs thm (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); val dnat_discsel = [ - prover [is_dzero_def] "is_dzero`dzero=TT", - prover [is_dzero_def] "n~=UU ==>is_dzero`(dsucc`n)=FF", - prover [is_dsucc_def] "is_dsucc`dzero=FF", - prover [is_dsucc_def] "n~=UU ==> is_dsucc`(dsucc`n)=TT", - prover [dpred_def] "dpred`dzero=UU", - prover [dpred_def] "n~=UU ==> dpred`(dsucc`n)=n" - ] @ dnat_discsel; + prover [is_dzero_def] "is_dzero`dzero=TT", + prover [is_dzero_def] "n~=UU ==>is_dzero`(dsucc`n)=FF", + prover [is_dsucc_def] "is_dsucc`dzero=FF", + prover [is_dsucc_def] "n~=UU ==> is_dsucc`(dsucc`n)=TT", + prover [dpred_def] "dpred`dzero=UU", + prover [dpred_def] "n~=UU ==> dpred`(dsucc`n)=n" + ] @ dnat_discsel; val dnat_rews = dnat_discsel @ dnat_rews; @@ -140,29 +140,29 @@ fun prover contr thm = prove_goal Dnat.thy thm (fn prems => - [ - (res_inst_tac [("P1",contr)] classical3 1), - (simp_tac (!simpset addsimps dnat_rews) 1), - (dtac sym 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps (prems @ dnat_rews)) 1) - ]); + [ + (res_inst_tac [("P1",contr)] classical3 1), + (simp_tac (!simpset addsimps dnat_rews) 1), + (dtac sym 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps (prems @ dnat_rews)) 1) + ]); val dnat_constrdef = [ - prover "is_dzero`UU ~= UU" "dzero~=UU", - prover "is_dsucc`UU ~= UU" "n~=UU ==> dsucc`n~=UU" - ]; + prover "is_dzero`UU ~= UU" "dzero~=UU", + prover "is_dsucc`UU ~= UU" "n~=UU ==> dsucc`n~=UU" + ]; fun prover defs thm = prove_goalw Dnat.thy defs thm (fn prems => - [ - (simp_tac (!simpset addsimps dnat_rews) 1) - ]); + [ + (simp_tac (!simpset addsimps dnat_rews) 1) + ]); val dnat_constrdef = [ - prover [dsucc_def] "dsucc`UU=UU" - ] @ dnat_constrdef; + prover [dsucc_def] "dsucc`UU=UU" + ] @ dnat_constrdef; val dnat_rews = dnat_constrdef @ dnat_rews; @@ -173,45 +173,45 @@ val temp = prove_goal Dnat.thy "~dzero << dsucc`n" (fn prems => - [ - (res_inst_tac [("P1","TT << FF")] classical3 1), - (resolve_tac dist_less_tr 1), - (dres_inst_tac [("fo5","is_dzero")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (res_inst_tac [("Q","n=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); + [ + (res_inst_tac [("P1","TT << FF")] classical3 1), + (resolve_tac dist_less_tr 1), + (dres_inst_tac [("fo5","is_dzero")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (res_inst_tac [("Q","n=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); val dnat_dist_less = [temp]; val temp = prove_goal Dnat.thy "n~=UU ==> ~dsucc`n << dzero" (fn prems => - [ - (cut_facts_tac prems 1), - (res_inst_tac [("P1","TT << FF")] classical3 1), - (resolve_tac dist_less_tr 1), - (dres_inst_tac [("fo5","is_dsucc")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); + [ + (cut_facts_tac prems 1), + (res_inst_tac [("P1","TT << FF")] classical3 1), + (resolve_tac dist_less_tr 1), + (dres_inst_tac [("fo5","is_dsucc")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); val dnat_dist_less = temp::dnat_dist_less; val temp = prove_goal Dnat.thy "dzero ~= dsucc`n" (fn prems => - [ - (res_inst_tac [("Q","n=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (res_inst_tac [("P1","TT = FF")] classical3 1), - (resolve_tac dist_eq_tr 1), - (dres_inst_tac [("f","is_dzero")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); + [ + (res_inst_tac [("Q","n=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (res_inst_tac [("P1","TT = FF")] classical3 1), + (resolve_tac dist_eq_tr 1), + (dres_inst_tac [("f","is_dzero")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); val dnat_dist_eq = [temp, temp RS not_sym]; @@ -222,36 +222,36 @@ (* ------------------------------------------------------------------------*) val dnat_invert = - [ + [ prove_goal Dnat.thy "[|x1~=UU; y1~=UU; dsucc`x1 << dsucc`y1 |] ==> x1<< y1" (fn prems => - [ - (cut_facts_tac prems 1), - (dres_inst_tac [("fo5","dnat_when`c`(LAM x.x)")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]) - ]; + [ + (cut_facts_tac prems 1), + (dres_inst_tac [("fo5","dnat_when`c`(LAM x.x)")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]) + ]; (* ------------------------------------------------------------------------*) (* Injectivity *) (* ------------------------------------------------------------------------*) val dnat_inject = - [ + [ prove_goal Dnat.thy "[|x1~=UU; y1~=UU; dsucc`x1 = dsucc`y1 |] ==> x1= y1" (fn prems => - [ - (cut_facts_tac prems 1), - (dres_inst_tac [("f","dnat_when`c`(LAM x.x)")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]) - ]; + [ + (cut_facts_tac prems 1), + (dres_inst_tac [("f","dnat_when`c`(LAM x.x)")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]) + ]; (* ------------------------------------------------------------------------*) (* definedness for discriminators and selectors *) @@ -260,18 +260,18 @@ fun prover thm = prove_goal Dnat.thy thm (fn prems => - [ - (cut_facts_tac prems 1), - (rtac dnatE 1), - (contr_tac 1), - (REPEAT (asm_simp_tac (!simpset addsimps dnat_rews) 1)) - ]); + [ + (cut_facts_tac prems 1), + (rtac dnatE 1), + (contr_tac 1), + (REPEAT (asm_simp_tac (!simpset addsimps dnat_rews) 1)) + ]); val dnat_discsel_def = - [ - prover "n~=UU ==> is_dzero`n ~= UU", - prover "n~=UU ==> is_dsucc`n ~= UU" - ]; + [ + prover "n~=UU ==> is_dzero`n ~= UU", + prover "n~=UU ==> is_dsucc`n ~= UU" + ]; val dnat_rews = dnat_discsel_def @ dnat_rews; @@ -281,49 +281,49 @@ (* ------------------------------------------------------------------------*) val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take n`UU = UU" (fn prems => - [ - (res_inst_tac [("n","n")] natE 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps dnat_rews) 1) - ]); + [ + (res_inst_tac [("n","n")] natE 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps dnat_rews) 1) + ]); val dnat_take = [temp]; val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take 0`xs = UU" (fn prems => - [ - (Asm_simp_tac 1) - ]); + [ + (Asm_simp_tac 1) + ]); val dnat_take = temp::dnat_take; val temp = prove_goalw Dnat.thy [dnat_take_def] - "dnat_take (Suc n)`dzero=dzero" + "dnat_take (Suc n)`dzero=dzero" (fn prems => - [ - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps dnat_rews) 1) - ]); + [ + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps dnat_rews) 1) + ]); val dnat_take = temp::dnat_take; val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take (Suc n)`(dsucc`xs)=dsucc`(dnat_take n ` xs)" (fn prems => - [ - (res_inst_tac [("Q","xs=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (res_inst_tac [("n","n")] natE 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); + [ + (res_inst_tac [("Q","xs=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (res_inst_tac [("n","n")] natE 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (Asm_simp_tac 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); val dnat_take = temp::dnat_take; @@ -336,23 +336,23 @@ fun prover reach defs thm = prove_goalw Dnat.thy defs thm (fn prems => - [ - (res_inst_tac [("t","s1")] (reach RS subst) 1), - (res_inst_tac [("t","s2")] (reach RS subst) 1), - (rtac (fix_def2 RS ssubst) 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac lub_equal 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac allI 1), - (resolve_tac prems 1) - ]); + [ + (res_inst_tac [("t","s1")] (reach RS subst) 1), + (res_inst_tac [("t","s2")] (reach RS subst) 1), + (rtac (fix_def2 RS ssubst) 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac lub_equal 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac allI 1), + (resolve_tac prems 1) + ]); val dnat_take_lemma = prover dnat_reach [dnat_take_def] - "(!!n.dnat_take n`s1 = dnat_take n`s2) ==> s1=s2"; + "(!!n.dnat_take n`s1 = dnat_take n`s2) ==> s1=s2"; (* ------------------------------------------------------------------------*) @@ -362,32 +362,32 @@ qed_goalw "dnat_coind_lemma" Dnat.thy [dnat_bisim_def] "dnat_bisim R ==> ! p q. R p q --> dnat_take n`p = dnat_take n`q" (fn prems => - [ - (cut_facts_tac prems 1), - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dnat_take) 1), - (strip_tac 1), - ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), - (atac 1), - (asm_simp_tac (!simpset addsimps dnat_take) 1), - (etac disjE 1), - (asm_simp_tac (!simpset addsimps dnat_take) 1), - (etac exE 1), - (etac exE 1), - (asm_simp_tac (!simpset addsimps dnat_take) 1), - (REPEAT (etac conjE 1)), - (rtac cfun_arg_cong 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dnat_take) 1), + (strip_tac 1), + ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), + (atac 1), + (asm_simp_tac (!simpset addsimps dnat_take) 1), + (etac disjE 1), + (asm_simp_tac (!simpset addsimps dnat_take) 1), + (etac exE 1), + (etac exE 1), + (asm_simp_tac (!simpset addsimps dnat_take) 1), + (REPEAT (etac conjE 1)), + (rtac cfun_arg_cong 1), + (fast_tac HOL_cs 1) + ]); qed_goal "dnat_coind" Dnat.thy "[|dnat_bisim R;R p q|] ==> p = q" (fn prems => - [ - (rtac dnat_take_lemma 1), - (rtac (dnat_coind_lemma RS spec RS spec RS mp) 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); + [ + (rtac dnat_take_lemma 1), + (rtac (dnat_coind_lemma RS spec RS spec RS mp) 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); (* ------------------------------------------------------------------------*) @@ -401,29 +401,29 @@ \ P(dzero);\ \ !! s1.[|s1~=UU ; P(s1)|] ==> P(dsucc`s1)|] ==> P(s)" (fn prems => - [ - (rtac (dnat_reach RS subst) 1), - (res_inst_tac [("x","s")] spec 1), - (rtac fix_ind 1), - (rtac adm_all2 1), - (rtac adm_subst 1), - (cont_tacR 1), - (resolve_tac prems 1), - (Simp_tac 1), - (resolve_tac prems 1), - (strip_tac 1), - (res_inst_tac [("n","xa")] dnatE 1), - (asm_simp_tac (!simpset addsimps dnat_copy) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dnat_copy) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dnat_copy) 1), - (res_inst_tac [("Q","x`xb=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (resolve_tac prems 1), - (eresolve_tac prems 1), - (etac spec 1) - ]); + [ + (rtac (dnat_reach RS subst) 1), + (res_inst_tac [("x","s")] spec 1), + (rtac fix_ind 1), + (rtac adm_all2 1), + (rtac adm_subst 1), + (cont_tacR 1), + (resolve_tac prems 1), + (Simp_tac 1), + (resolve_tac prems 1), + (strip_tac 1), + (res_inst_tac [("n","xa")] dnatE 1), + (asm_simp_tac (!simpset addsimps dnat_copy) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dnat_copy) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dnat_copy) 1), + (res_inst_tac [("Q","x`xb=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (resolve_tac prems 1), + (eresolve_tac prems 1), + (etac spec 1) + ]); *) qed_goal "dnat_finite_ind" Dnat.thy @@ -431,58 +431,58 @@ \ !! s1.[|s1~=UU;P(s1)|] ==> P(dsucc`s1)\ \ |] ==> !s.P(dnat_take n`s)" (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dnat_rews) 1), - (resolve_tac prems 1), - (rtac allI 1), - (res_inst_tac [("n","s")] dnatE 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (res_inst_tac [("Q","dnat_take n1`x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (resolve_tac prems 1), - (resolve_tac prems 1), - (atac 1), - (etac spec 1) - ]); + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dnat_rews) 1), + (resolve_tac prems 1), + (rtac allI 1), + (res_inst_tac [("n","s")] dnatE 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (res_inst_tac [("Q","dnat_take n1`x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (resolve_tac prems 1), + (resolve_tac prems 1), + (atac 1), + (etac spec 1) + ]); qed_goal "dnat_all_finite_lemma1" Dnat.thy "!s.dnat_take n`s=UU |dnat_take n`s=s" (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps dnat_rews) 1), - (rtac allI 1), - (res_inst_tac [("n","s")] dnatE 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (eres_inst_tac [("x","x")] allE 1), - (etac disjE 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1) - ]); + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps dnat_rews) 1), + (rtac allI 1), + (res_inst_tac [("n","s")] dnatE 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (eres_inst_tac [("x","x")] allE 1), + (etac disjE 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1) + ]); qed_goal "dnat_all_finite_lemma2" Dnat.thy "? n.dnat_take n`s=s" (fn prems => - [ - (res_inst_tac [("Q","s=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (subgoal_tac "(!n.dnat_take(n)`s=UU) |(? n.dnat_take(n)`s=s)" 1), - (etac disjE 1), - (eres_inst_tac [("P","s=UU")] notE 1), - (rtac dnat_take_lemma 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (atac 1), - (subgoal_tac "!n.!s.dnat_take(n)`s=UU |dnat_take(n)`s=s" 1), - (fast_tac HOL_cs 1), - (rtac allI 1), - (rtac dnat_all_finite_lemma1 1) - ]); + [ + (res_inst_tac [("Q","s=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (subgoal_tac "(!n.dnat_take(n)`s=UU) |(? n.dnat_take(n)`s=s)" 1), + (etac disjE 1), + (eres_inst_tac [("P","s=UU")] notE 1), + (rtac dnat_take_lemma 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (atac 1), + (subgoal_tac "!n.!s.dnat_take(n)`s=UU |dnat_take(n)`s=s" 1), + (fast_tac HOL_cs 1), + (rtac allI 1), + (rtac dnat_all_finite_lemma1 1) + ]); qed_goal "dnat_ind" Dnat.thy @@ -490,42 +490,42 @@ \ !! s1.[|s1~=UU;P(s1)|] ==> P(dsucc`s1)\ \ |] ==> P(s)" (fn prems => - [ - (rtac (dnat_all_finite_lemma2 RS exE) 1), - (etac subst 1), - (rtac (dnat_finite_ind RS spec) 1), - (REPEAT (resolve_tac prems 1)), - (REPEAT (atac 1)) - ]); + [ + (rtac (dnat_all_finite_lemma2 RS exE) 1), + (etac subst 1), + (rtac (dnat_finite_ind RS spec) 1), + (REPEAT (resolve_tac prems 1)), + (REPEAT (atac 1)) + ]); qed_goalw "dnat_flat" Dnat.thy [is_flat_def] "is_flat(dzero)" (fn prems => - [ - (rtac allI 1), - (res_inst_tac [("s","x")] dnat_ind 1), - (fast_tac HOL_cs 1), - (rtac allI 1), - (res_inst_tac [("n","y")] dnatE 1), - (fast_tac (HOL_cs addSIs [UU_I]) 1), - (Asm_simp_tac 1), - (asm_simp_tac (!simpset addsimps dnat_dist_less) 1), - (rtac allI 1), - (res_inst_tac [("n","y")] dnatE 1), - (fast_tac (HOL_cs addSIs [UU_I]) 1), - (asm_simp_tac (!simpset addsimps dnat_dist_less) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (strip_tac 1), - (subgoal_tac "s1< - [ - (rtac (iterator_def2 RS ssubst) 1), - (simp_tac (!simpset addsimps dnat_when) 1) - ]); + [ + (rtac (iterator_def2 RS ssubst) 1), + (simp_tac (!simpset addsimps dnat_when) 1) + ]); qed_goal "iterator2" Dnat2.thy "iterator`dzero`f`x = x" (fn prems => - [ - (rtac (iterator_def2 RS ssubst) 1), - (simp_tac (!simpset addsimps dnat_when) 1) - ]); + [ + (rtac (iterator_def2 RS ssubst) 1), + (simp_tac (!simpset addsimps dnat_when) 1) + ]); qed_goal "iterator3" Dnat2.thy "n~=UU ==> iterator`(dsucc`n)`f`x = f`(iterator`n`f`x)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans 1), - (rtac (iterator_def2 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps dnat_rews) 1), - (rtac refl 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac trans 1), + (rtac (iterator_def2 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps dnat_rews) 1), + (rtac refl 1) + ]); val dnat2_rews = - [iterator1, iterator2, iterator3]; + [iterator1, iterator2, iterator3]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/explicit_domains/Focus_ex.ML --- a/src/HOLCF/explicit_domains/Focus_ex.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/explicit_domains/Focus_ex.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ (* ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1995 Technische Universitaet Muenchen *) @@ -12,19 +12,19 @@ val prems = goal Focus_ex.thy "is_g(g) = \ \ (? f. is_f(f) & (!x.(? z. = f` & \ -\ (! w y. = f` --> z << w))))"; +\ (! w y. = f` --> z << w))))"; by (simp_tac (!simpset addsimps [is_g,is_net_g]) 1); by (fast_tac HOL_cs 1); val lemma1 = result(); val prems = goal Focus_ex.thy "(? f. is_f(f) & (!x. (? z. = f` & \ -\ (! w y. = f` --> z << w)))) \ +\ (! w y. = f` --> z << w)))) \ \ = \ \ (? f. is_f(f) & (!x. ? z.\ \ g`x = cfst`(f`) & \ \ z = csnd`(f`) & \ -\ (! w y. = f` --> z << w)))"; +\ (! w y. = f` --> z << w)))"; by (rtac iffI 1); by (etac exE 1); by (res_inst_tac [("x","f")] exI 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/explicit_domains/ROOT.ML --- a/src/HOLCF/explicit_domains/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/explicit_domains/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ (* ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1995 Technische Universitaet Muenchen *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/explicit_domains/Stream.ML --- a/src/HOLCF/explicit_domains/Stream.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/explicit_domains/Stream.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ (* ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for stream.thy @@ -13,10 +13,10 @@ (* ------------------------------------------------------------------------*) val stream_iso_strict= stream_rep_iso RS (stream_abs_iso RS - (allI RSN (2,allI RS iso_strict))); + (allI RSN (2,allI RS iso_strict))); val stream_rews = [stream_iso_strict RS conjunct1, - stream_iso_strict RS conjunct2]; + stream_iso_strict RS conjunct2]; (* ------------------------------------------------------------------------*) (* Properties of stream_copy *) @@ -24,18 +24,18 @@ fun prover defs thm = prove_goalw Stream.thy defs thm (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps - (stream_rews @ [stream_abs_iso,stream_rep_iso])) 1) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps + (stream_rews @ [stream_abs_iso,stream_rep_iso])) 1) + ]); val stream_copy = - [ - prover [stream_copy_def] "stream_copy`f`UU=UU", - prover [stream_copy_def,scons_def] - "x~=UU ==> stream_copy`f`(scons`x`xs)= scons`x`(f`xs)" - ]; + [ + prover [stream_copy_def] "stream_copy`f`UU=UU", + prover [stream_copy_def,scons_def] + "x~=UU ==> stream_copy`f`(scons`x`xs)= scons`x`(f`xs)" + ]; val stream_rews = stream_copy @ stream_rews; @@ -44,35 +44,35 @@ (* ------------------------------------------------------------------------*) qed_goalw "Exh_stream" Stream.thy [scons_def] - "s = UU | (? x xs. x~=UU & s = scons`x`xs)" + "s = UU | (? x xs. x~=UU & s = scons`x`xs)" (fn prems => - [ - (Simp_tac 1), - (rtac (stream_rep_iso RS subst) 1), - (res_inst_tac [("p","stream_rep`s")] sprodE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (Asm_simp_tac 1), - (res_inst_tac [("p","y")] liftE1 1), - (contr_tac 1), - (rtac disjI2 1), - (rtac exI 1), - (rtac exI 1), - (etac conjI 1), - (Asm_simp_tac 1) - ]); + [ + (Simp_tac 1), + (rtac (stream_rep_iso RS subst) 1), + (res_inst_tac [("p","stream_rep`s")] sprodE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (Asm_simp_tac 1), + (res_inst_tac [("p","y")] liftE1 1), + (contr_tac 1), + (rtac disjI2 1), + (rtac exI 1), + (rtac exI 1), + (etac conjI 1), + (Asm_simp_tac 1) + ]); qed_goal "streamE" Stream.thy - "[| s=UU ==> Q; !!x xs.[|s=scons`x`xs;x~=UU|]==>Q|]==>Q" + "[| s=UU ==> Q; !!x xs.[|s=scons`x`xs;x~=UU|]==>Q|]==>Q" (fn prems => - [ - (rtac (Exh_stream RS disjE) 1), - (eresolve_tac prems 1), - (etac exE 1), - (etac exE 1), - (resolve_tac prems 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); + [ + (rtac (Exh_stream RS disjE) 1), + (eresolve_tac prems 1), + (etac exE 1), + (etac exE 1), + (resolve_tac prems 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); (* ------------------------------------------------------------------------*) (* Properties of stream_when *) @@ -80,18 +80,18 @@ fun prover defs thm = prove_goalw Stream.thy defs thm (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps - (stream_rews @ [stream_abs_iso,stream_rep_iso])) 1) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps + (stream_rews @ [stream_abs_iso,stream_rep_iso])) 1) + ]); val stream_when = [ - prover [stream_when_def] "stream_when`f`UU=UU", - prover [stream_when_def,scons_def] - "x~=UU ==> stream_when`f`(scons`x`xs)= f`x`xs" - ]; + prover [stream_when_def] "stream_when`f`UU=UU", + prover [stream_when_def,scons_def] + "x~=UU ==> stream_when`f`(scons`x`xs)= f`x`xs" + ]; val stream_rews = stream_when @ stream_rews; @@ -101,28 +101,28 @@ fun prover defs thm = prove_goalw Stream.thy defs thm (fn prems => - [ - (simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (simp_tac (!simpset addsimps stream_rews) 1) + ]); val stream_discsel = [ - prover [is_scons_def] "is_scons`UU=UU", - prover [shd_def] "shd`UU=UU", - prover [stl_def] "stl`UU=UU" - ]; + prover [is_scons_def] "is_scons`UU=UU", + prover [shd_def] "shd`UU=UU", + prover [stl_def] "stl`UU=UU" + ]; fun prover defs thm = prove_goalw Stream.thy defs thm (fn prems => - [ - (cut_facts_tac prems 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (cut_facts_tac prems 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); val stream_discsel = [ prover [is_scons_def,shd_def,stl_def] "x~=UU ==> is_scons`(scons`x`xs)=TT", prover [is_scons_def,shd_def,stl_def] "x~=UU ==> shd`(scons`x`xs)=x", prover [is_scons_def,shd_def,stl_def] "x~=UU ==> stl`(scons`x`xs)=xs" - ] @ stream_discsel; + ] @ stream_discsel; val stream_rews = stream_discsel @ stream_rews; @@ -132,27 +132,27 @@ fun prover contr thm = prove_goal Stream.thy thm (fn prems => - [ - (res_inst_tac [("P1",contr)] classical3 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (dtac sym 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps (prems @ stream_rews)) 1) - ]); + [ + (res_inst_tac [("P1",contr)] classical3 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (dtac sym 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps (prems @ stream_rews)) 1) + ]); val stream_constrdef = [ - prover "is_scons`(UU::'a stream)~=UU" "x~=UU ==> scons`(x::'a)`xs~=UU" - ]; + prover "is_scons`(UU::'a stream)~=UU" "x~=UU ==> scons`(x::'a)`xs~=UU" + ]; fun prover defs thm = prove_goalw Stream.thy defs thm (fn prems => - [ - (simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (simp_tac (!simpset addsimps stream_rews) 1) + ]); val stream_constrdef = [ - prover [scons_def] "scons`UU`xs=UU" - ] @ stream_constrdef; + prover [scons_def] "scons`UU`xs=UU" + ] @ stream_constrdef; val stream_rews = stream_constrdef @ stream_rews; @@ -167,46 +167,46 @@ (* ------------------------------------------------------------------------*) val stream_invert = - [ + [ prove_goal Stream.thy "[|x1~=UU; y1~=UU;\ \ scons`x1`x2 << scons`y1`y2|] ==> x1<< y1 & x2 << y2" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - (dres_inst_tac [("fo5","stream_when`(LAM x l.x)")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (dres_inst_tac [("fo5","stream_when`(LAM x l.l)")] monofun_cfun_arg 1), - (etac box_less 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (asm_simp_tac (!simpset addsimps stream_when) 1) - ]) - ]; + [ + (cut_facts_tac prems 1), + (rtac conjI 1), + (dres_inst_tac [("fo5","stream_when`(LAM x l.x)")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (dres_inst_tac [("fo5","stream_when`(LAM x l.l)")] monofun_cfun_arg 1), + (etac box_less 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (asm_simp_tac (!simpset addsimps stream_when) 1) + ]) + ]; (* ------------------------------------------------------------------------*) (* Injectivity *) (* ------------------------------------------------------------------------*) val stream_inject = - [ + [ prove_goal Stream.thy "[|x1~=UU; y1~=UU;\ \ scons`x1`x2 = scons`y1`y2 |] ==> x1= y1 & x2 = y2" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac conjI 1), - (dres_inst_tac [("f","stream_when`(LAM x l.x)")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (dres_inst_tac [("f","stream_when`(LAM x l.l)")] cfun_arg_cong 1), - (etac box_equals 1), - (asm_simp_tac (!simpset addsimps stream_when) 1), - (asm_simp_tac (!simpset addsimps stream_when) 1) - ]) - ]; + [ + (cut_facts_tac prems 1), + (rtac conjI 1), + (dres_inst_tac [("f","stream_when`(LAM x l.x)")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (dres_inst_tac [("f","stream_when`(LAM x l.l)")] cfun_arg_cong 1), + (etac box_equals 1), + (asm_simp_tac (!simpset addsimps stream_when) 1), + (asm_simp_tac (!simpset addsimps stream_when) 1) + ]) + ]; (* ------------------------------------------------------------------------*) (* definedness for discriminators and selectors *) @@ -214,18 +214,18 @@ fun prover thm = prove_goal Stream.thy thm (fn prems => - [ - (cut_facts_tac prems 1), - (rtac streamE 1), - (contr_tac 1), - (REPEAT (asm_simp_tac (!simpset addsimps stream_discsel) 1)) - ]); + [ + (cut_facts_tac prems 1), + (rtac streamE 1), + (contr_tac 1), + (REPEAT (asm_simp_tac (!simpset addsimps stream_discsel) 1)) + ]); val stream_discsel_def = - [ - prover "s~=UU ==> is_scons`s ~= UU", - prover "s~=UU ==> shd`s ~=UU" - ]; + [ + prover "s~=UU ==> is_scons`s ~= UU", + prover "s~=UU ==> shd`s ~=UU" + ]; val stream_rews = stream_discsel_def @ stream_rews; @@ -235,33 +235,33 @@ (* ------------------------------------------------------------------------*) val stream_take = - [ + [ prove_goalw Stream.thy [stream_take_def] "stream_take n`UU = UU" (fn prems => - [ - (res_inst_tac [("n","n")] natE 1), - (Asm_simp_tac 1), - (Asm_simp_tac 1), - (simp_tac (!simpset addsimps stream_rews) 1) - ]), + [ + (res_inst_tac [("n","n")] natE 1), + (Asm_simp_tac 1), + (Asm_simp_tac 1), + (simp_tac (!simpset addsimps stream_rews) 1) + ]), prove_goalw Stream.thy [stream_take_def] "stream_take 0`xs=UU" (fn prems => - [ - (Asm_simp_tac 1) - ])]; + [ + (Asm_simp_tac 1) + ])]; fun prover thm = prove_goalw Stream.thy [stream_take_def] thm (fn prems => - [ - (cut_facts_tac prems 1), - (Simp_tac 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (cut_facts_tac prems 1), + (Simp_tac 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); val stream_take = [ prover "x~=UU ==> stream_take (Suc n)`(scons`x`xs) = scons`x`(stream_take n`xs)" - ] @ stream_take; + ] @ stream_take; val stream_rews = stream_take @ stream_rews; @@ -272,31 +272,31 @@ qed_goal "stream_copy2" Stream.thy "stream_copy`f`(scons`x`xs) = scons`x`(f`xs)" (fn prems => - [ - (res_inst_tac [("Q","x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (res_inst_tac [("Q","x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); qed_goal "shd2" Stream.thy "shd`(scons`x`xs) = x" (fn prems => - [ - (res_inst_tac [("Q","x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (res_inst_tac [("Q","x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); qed_goal "stream_take2" Stream.thy "stream_take (Suc n)`(scons`x`xs) = scons`x`(stream_take n`xs)" (fn prems => - [ - (res_inst_tac [("Q","x=UU")] classical2 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (res_inst_tac [("Q","x=UU")] classical2 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); val stream_rews = [stream_iso_strict RS conjunct1, - stream_iso_strict RS conjunct2, + stream_iso_strict RS conjunct2, hd stream_copy, stream_copy2] @ stream_when @ [hd stream_discsel,shd2] @ (tl (tl stream_discsel)) @@ -311,35 +311,35 @@ fun prover reach defs thm = prove_goalw Stream.thy defs thm (fn prems => - [ - (res_inst_tac [("t","s1")] (reach RS subst) 1), - (res_inst_tac [("t","s2")] (reach RS subst) 1), - (rtac (fix_def2 RS ssubst) 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac lub_equal 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac (is_chain_iterate RS ch2ch_fappL) 1), - (rtac allI 1), - (resolve_tac prems 1) - ]); + [ + (res_inst_tac [("t","s1")] (reach RS subst) 1), + (res_inst_tac [("t","s2")] (reach RS subst) 1), + (rtac (fix_def2 RS ssubst) 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac lub_equal 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac (is_chain_iterate RS ch2ch_fappL) 1), + (rtac allI 1), + (resolve_tac prems 1) + ]); val stream_take_lemma = prover stream_reach [stream_take_def] - "(!!n.stream_take n`s1 = stream_take n`s2) ==> s1=s2"; + "(!!n.stream_take n`s1 = stream_take n`s2) ==> s1=s2"; qed_goal "stream_reach2" Stream.thy "lub(range(%i.stream_take i`s))=s" (fn prems => - [ - (res_inst_tac [("t","s")] (stream_reach RS subst) 1), - (rtac (fix_def2 RS ssubst) 1), - (rewrite_goals_tac [stream_take_def]), - (rtac (contlub_cfun_fun RS ssubst) 1), - (rtac is_chain_iterate 1), - (rtac refl 1) - ]); + [ + (res_inst_tac [("t","s")] (stream_reach RS subst) 1), + (rtac (fix_def2 RS ssubst) 1), + (rewtac stream_take_def), + (rtac (contlub_cfun_fun RS ssubst) 1), + (rtac is_chain_iterate 1), + (rtac refl 1) + ]); (* ------------------------------------------------------------------------*) (* Co -induction for streams *) @@ -348,31 +348,31 @@ qed_goalw "stream_coind_lemma" Stream.thy [stream_bisim_def] "stream_bisim R ==> ! p q. R p q --> stream_take n`p = stream_take n`q" (fn prems => - [ - (cut_facts_tac prems 1), - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1), - ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), - (atac 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (etac exE 1), - (etac exE 1), - (etac exE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (REPEAT (etac conjE 1)), - (rtac cfun_arg_cong 1), - (fast_tac HOL_cs 1) - ]); + [ + (cut_facts_tac prems 1), + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1), + ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), + (atac 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (etac exE 1), + (etac exE 1), + (etac exE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (REPEAT (etac conjE 1)), + (rtac cfun_arg_cong 1), + (fast_tac HOL_cs 1) + ]); qed_goal "stream_coind" Stream.thy "[|stream_bisim R ;R p q|] ==> p = q" (fn prems => - [ - (rtac stream_take_lemma 1), - (rtac (stream_coind_lemma RS spec RS spec RS mp) 1), - (resolve_tac prems 1), - (resolve_tac prems 1) - ]); + [ + (rtac stream_take_lemma 1), + (rtac (stream_coind_lemma RS spec RS spec RS mp) 1), + (resolve_tac prems 1), + (resolve_tac prems 1) + ]); (* ------------------------------------------------------------------------*) (* structural induction for admissible predicates *) @@ -383,41 +383,41 @@ \ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ \ |] ==> !s.P(stream_take n`s)" (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (resolve_tac prems 1), - (rtac allI 1), - (res_inst_tac [("s","s")] streamE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (resolve_tac prems 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (resolve_tac prems 1), - (atac 1), - (etac spec 1) - ]); + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (resolve_tac prems 1), + (rtac allI 1), + (res_inst_tac [("s","s")] streamE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (resolve_tac prems 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (resolve_tac prems 1), + (atac 1), + (etac spec 1) + ]); qed_goalw "stream_finite_ind2" Stream.thy [stream_finite_def] "(!!n.P(stream_take n`s)) ==> stream_finite(s) -->P(s)" (fn prems => - [ - (strip_tac 1), - (etac exE 1), - (etac subst 1), - (resolve_tac prems 1) - ]); + [ + (strip_tac 1), + (etac exE 1), + (etac subst 1), + (resolve_tac prems 1) + ]); qed_goal "stream_finite_ind3" Stream.thy "[|P(UU);\ \ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ \ |] ==> stream_finite(s) --> P(s)" (fn prems => - [ - (rtac stream_finite_ind2 1), - (rtac (stream_finite_ind RS spec) 1), - (REPEAT (resolve_tac prems 1)), - (REPEAT (atac 1)) - ]); + [ + (rtac stream_finite_ind2 1), + (rtac (stream_finite_ind RS spec) 1), + (REPEAT (resolve_tac prems 1)), + (REPEAT (atac 1)) + ]); (* prove induction using definition of admissibility stream_reach rsp. stream_reach2 @@ -429,18 +429,18 @@ \ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ \ |] ==> P(s)" (fn prems => - [ - (rtac (stream_reach2 RS subst) 1), - (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), - (resolve_tac prems 1), - (SELECT_GOAL (rewrite_goals_tac [stream_take_def]) 1), - (rtac ch2ch_fappL 1), - (rtac is_chain_iterate 1), - (rtac allI 1), - (rtac (stream_finite_ind RS spec) 1), - (REPEAT (resolve_tac prems 1)), - (REPEAT (atac 1)) - ]); + [ + (rtac (stream_reach2 RS subst) 1), + (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), + (resolve_tac prems 1), + (SELECT_GOAL (rewtac stream_take_def) 1), + (rtac ch2ch_fappL 1), + (rtac is_chain_iterate 1), + (rtac allI 1), + (rtac (stream_finite_ind RS spec) 1), + (REPEAT (resolve_tac prems 1)), + (REPEAT (atac 1)) + ]); (* prove induction with usual LCF-Method using fixed point induction *) qed_goal "stream_ind" Stream.thy @@ -449,20 +449,20 @@ \ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ \ |] ==> P(s)" (fn prems => - [ - (rtac (stream_reach RS subst) 1), - (res_inst_tac [("x","s")] spec 1), - (rtac wfix_ind 1), - (rtac adm_impl_admw 1), - (REPEAT (resolve_tac adm_thms 1)), - (rtac adm_subst 1), - (cont_tacR 1), - (resolve_tac prems 1), - (rtac allI 1), - (rtac (rewrite_rule [stream_take_def] stream_finite_ind) 1), - (REPEAT (resolve_tac prems 1)), - (REPEAT (atac 1)) - ]); + [ + (rtac (stream_reach RS subst) 1), + (res_inst_tac [("x","s")] spec 1), + (rtac wfix_ind 1), + (rtac adm_impl_admw 1), + (REPEAT (resolve_tac adm_thms 1)), + (rtac adm_subst 1), + (cont_tacR 1), + (resolve_tac prems 1), + (rtac allI 1), + (rtac (rewrite_rule [stream_take_def] stream_finite_ind) 1), + (REPEAT (resolve_tac prems 1)), + (REPEAT (atac 1)) + ]); (* ------------------------------------------------------------------------*) @@ -471,48 +471,48 @@ qed_goal "surjectiv_scons" Stream.thy "scons`(shd`s)`(stl`s)=s" (fn prems => - [ - (res_inst_tac [("s","s")] streamE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (res_inst_tac [("s","s")] streamE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); qed_goalw "stream_coind_lemma2" Stream.thy [stream_bisim_def] "!s1 s2. R s1 s2 --> shd`s1 = shd`s2 & R (stl`s1) (stl`s2) ==> stream_bisim R" (fn prems => - [ - (cut_facts_tac prems 1), - (strip_tac 1), - (etac allE 1), - (etac allE 1), - (dtac mp 1), - (atac 1), - (etac conjE 1), - (res_inst_tac [("Q","s1 = UU & s2 = UU")] classical2 1), - (rtac disjI1 1), - (fast_tac HOL_cs 1), - (rtac disjI2 1), - (rtac disjE 1), - (etac (de_morgan2 RS ssubst) 1), - (res_inst_tac [("x","shd`s1")] exI 1), - (res_inst_tac [("x","stl`s1")] exI 1), - (res_inst_tac [("x","stl`s2")] exI 1), - (rtac conjI 1), - (eresolve_tac stream_discsel_def 1), - (asm_simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), - (eres_inst_tac [("s","shd`s1"),("t","shd`s2")] subst 1), - (simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), - (res_inst_tac [("x","shd`s2")] exI 1), - (res_inst_tac [("x","stl`s1")] exI 1), - (res_inst_tac [("x","stl`s2")] exI 1), - (rtac conjI 1), - (eresolve_tac stream_discsel_def 1), - (asm_simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), - (res_inst_tac [("s","shd`s1"),("t","shd`s2")] ssubst 1), - (etac sym 1), - (simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1) - ]); + [ + (cut_facts_tac prems 1), + (strip_tac 1), + (etac allE 1), + (etac allE 1), + (dtac mp 1), + (atac 1), + (etac conjE 1), + (res_inst_tac [("Q","s1 = UU & s2 = UU")] classical2 1), + (rtac disjI1 1), + (fast_tac HOL_cs 1), + (rtac disjI2 1), + (rtac disjE 1), + (etac (de_morgan2 RS ssubst) 1), + (res_inst_tac [("x","shd`s1")] exI 1), + (res_inst_tac [("x","stl`s1")] exI 1), + (res_inst_tac [("x","stl`s2")] exI 1), + (rtac conjI 1), + (eresolve_tac stream_discsel_def 1), + (asm_simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), + (eres_inst_tac [("s","shd`s1"),("t","shd`s2")] subst 1), + (simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), + (res_inst_tac [("x","shd`s2")] exI 1), + (res_inst_tac [("x","stl`s1")] exI 1), + (res_inst_tac [("x","stl`s2")] exI 1), + (rtac conjI 1), + (eresolve_tac stream_discsel_def 1), + (asm_simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), + (res_inst_tac [("s","shd`s1"),("t","shd`s2")] ssubst 1), + (etac sym 1), + (simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1) + ]); (* ------------------------------------------------------------------------*) @@ -524,22 +524,22 @@ (* ----------------------------------------------------------------------- *) qed_goalw "stream_finite_UU" Stream.thy [stream_finite_def] - "stream_finite(UU)" + "stream_finite(UU)" (fn prems => - [ - (rtac exI 1), - (simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (rtac exI 1), + (simp_tac (!simpset addsimps stream_rews) 1) + ]); qed_goal "inf_stream_not_UU" Stream.thy "~stream_finite(s) ==> s ~= UU" (fn prems => - [ - (cut_facts_tac prems 1), - (etac swap 1), - (dtac notnotD 1), - (hyp_subst_tac 1), - (rtac stream_finite_UU 1) - ]); + [ + (cut_facts_tac prems 1), + (etac swap 1), + (dtac notnotD 1), + (hyp_subst_tac 1), + (rtac stream_finite_UU 1) + ]); (* ----------------------------------------------------------------------- *) (* a lemma about shd *) @@ -547,12 +547,12 @@ qed_goal "stream_shd_lemma1" Stream.thy "shd`s=UU --> s=UU" (fn prems => - [ - (res_inst_tac [("s","s")] streamE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (hyp_subst_tac 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (res_inst_tac [("s","s")] streamE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (hyp_subst_tac 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); (* ----------------------------------------------------------------------- *) @@ -563,134 +563,134 @@ "!x xs.x~=UU --> \ \ stream_take (Suc n)`(scons`x`xs) = scons`x`xs --> stream_take n`xs=xs" (fn prems => - [ - (rtac allI 1), - (rtac allI 1), - (rtac impI 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1), - (rtac ((hd stream_inject) RS conjunct2) 1), - (atac 1), - (atac 1), - (atac 1) - ]); + [ + (rtac allI 1), + (rtac allI 1), + (rtac impI 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1), + (rtac ((hd stream_inject) RS conjunct2) 1), + (atac 1), + (atac 1), + (atac 1) + ]); qed_goal "stream_take_lemma2" Stream.thy "! s2. stream_take n`s2 = s2 --> stream_take (Suc n)`s2=s2" (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1 ), - (hyp_subst_tac 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (rtac allI 1), - (res_inst_tac [("s","s2")] streamE 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1 ), - (subgoal_tac "stream_take n1`xs = xs" 1), - (rtac ((hd stream_inject) RS conjunct2) 2), - (atac 4), - (atac 2), - (atac 2), - (rtac cfun_arg_cong 1), - (fast_tac HOL_cs 1) - ]); + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1 ), + (hyp_subst_tac 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (rtac allI 1), + (res_inst_tac [("s","s2")] streamE 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1 ), + (subgoal_tac "stream_take n1`xs = xs" 1), + (rtac ((hd stream_inject) RS conjunct2) 2), + (atac 4), + (atac 2), + (atac 2), + (rtac cfun_arg_cong 1), + (fast_tac HOL_cs 1) + ]); qed_goal "stream_take_lemma3" Stream.thy "!x xs.x~=UU --> \ \ stream_take n`(scons`x`xs) = scons`x`xs --> stream_take n`xs=xs" (fn prems => - [ - (nat_ind_tac "n" 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1 ), - (res_inst_tac [("P","scons`x`xs=UU")] notE 1), - (eresolve_tac stream_constrdef 1), - (etac sym 1), - (strip_tac 1 ), - (rtac (stream_take_lemma2 RS spec RS mp) 1), - (res_inst_tac [("x1.1","x")] ((hd stream_inject) RS conjunct2) 1), - (atac 1), - (atac 1), - (etac (stream_take2 RS subst) 1) - ]); + [ + (nat_ind_tac "n" 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1 ), + (res_inst_tac [("P","scons`x`xs=UU")] notE 1), + (eresolve_tac stream_constrdef 1), + (etac sym 1), + (strip_tac 1 ), + (rtac (stream_take_lemma2 RS spec RS mp) 1), + (res_inst_tac [("x1.1","x")] ((hd stream_inject) RS conjunct2) 1), + (atac 1), + (atac 1), + (etac (stream_take2 RS subst) 1) + ]); qed_goal "stream_take_lemma4" Stream.thy "!x xs.\ \stream_take n`xs=xs --> stream_take (Suc n)`(scons`x`xs) = scons`x`xs" (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (simp_tac (!simpset addsimps stream_rews) 1) + ]); (* ---- *) qed_goal "stream_take_lemma5" Stream.thy "!s. stream_take n`s=s --> iterate n stl s=UU" (fn prems => - [ - (nat_ind_tac "n" 1), - (Simp_tac 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1), - (res_inst_tac [("s","s")] streamE 1), - (hyp_subst_tac 1), - (rtac (iterate_Suc2 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (rtac (iterate_Suc2 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (etac allE 1), - (etac mp 1), - (hyp_subst_tac 1), - (etac (stream_take_lemma1 RS spec RS spec RS mp RS mp) 1), - (atac 1) - ]); + [ + (nat_ind_tac "n" 1), + (Simp_tac 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1), + (res_inst_tac [("s","s")] streamE 1), + (hyp_subst_tac 1), + (rtac (iterate_Suc2 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (rtac (iterate_Suc2 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (etac allE 1), + (etac mp 1), + (hyp_subst_tac 1), + (etac (stream_take_lemma1 RS spec RS spec RS mp RS mp) 1), + (atac 1) + ]); qed_goal "stream_take_lemma6" Stream.thy "!s.iterate n stl s =UU --> stream_take n`s=s" (fn prems => - [ - (nat_ind_tac "n" 1), - (Simp_tac 1), - (strip_tac 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (rtac allI 1), - (res_inst_tac [("s","s")] streamE 1), - (hyp_subst_tac 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (hyp_subst_tac 1), - (rtac (iterate_Suc2 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1) - ]); + [ + (nat_ind_tac "n" 1), + (Simp_tac 1), + (strip_tac 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (rtac allI 1), + (res_inst_tac [("s","s")] streamE 1), + (hyp_subst_tac 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (hyp_subst_tac 1), + (rtac (iterate_Suc2 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1) + ]); qed_goal "stream_take_lemma7" Stream.thy "(iterate n stl s=UU) = (stream_take n`s=s)" (fn prems => - [ - (rtac iffI 1), - (etac (stream_take_lemma6 RS spec RS mp) 1), - (etac (stream_take_lemma5 RS spec RS mp) 1) - ]); + [ + (rtac iffI 1), + (etac (stream_take_lemma6 RS spec RS mp) 1), + (etac (stream_take_lemma5 RS spec RS mp) 1) + ]); qed_goal "stream_take_lemma8" Stream.thy "[|adm(P); !n. ? m. n < m & P (stream_take m`s)|] ==> P(s)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac (stream_reach2 RS subst) 1), - (rtac adm_disj_lemma11 1), - (atac 1), - (atac 2), - (rewrite_goals_tac [stream_take_def]), - (rtac ch2ch_fappL 1), - (rtac is_chain_iterate 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac (stream_reach2 RS subst) 1), + (rtac adm_disj_lemma11 1), + (atac 1), + (atac 2), + (rewtac stream_take_def), + (rtac ch2ch_fappL 1), + (rtac is_chain_iterate 1) + ]); (* ----------------------------------------------------------------------- *) (* lemmas stream_finite *) @@ -699,103 +699,103 @@ qed_goalw "stream_finite_lemma1" Stream.thy [stream_finite_def] "stream_finite(xs) ==> stream_finite(scons`x`xs)" (fn prems => - [ - (cut_facts_tac prems 1), - (etac exE 1), - (rtac exI 1), - (etac (stream_take_lemma4 RS spec RS spec RS mp) 1) - ]); + [ + (cut_facts_tac prems 1), + (etac exE 1), + (rtac exI 1), + (etac (stream_take_lemma4 RS spec RS spec RS mp) 1) + ]); qed_goalw "stream_finite_lemma2" Stream.thy [stream_finite_def] "[|x~=UU; stream_finite(scons`x`xs)|] ==> stream_finite(xs)" (fn prems => - [ - (cut_facts_tac prems 1), - (etac exE 1), - (rtac exI 1), - (etac (stream_take_lemma3 RS spec RS spec RS mp RS mp) 1), - (atac 1) - ]); + [ + (cut_facts_tac prems 1), + (etac exE 1), + (rtac exI 1), + (etac (stream_take_lemma3 RS spec RS spec RS mp RS mp) 1), + (atac 1) + ]); qed_goal "stream_finite_lemma3" Stream.thy "x~=UU ==> stream_finite(scons`x`xs) = stream_finite(xs)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac iffI 1), - (etac stream_finite_lemma2 1), - (atac 1), - (etac stream_finite_lemma1 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac iffI 1), + (etac stream_finite_lemma2 1), + (atac 1), + (etac stream_finite_lemma1 1) + ]); qed_goalw "stream_finite_lemma5" Stream.thy [stream_finite_def] "(!n. s1 << s2 --> stream_take n`s2 = s2 --> stream_finite(s1))\ \=(s1 << s2 --> stream_finite(s2) --> stream_finite(s1))" (fn prems => - [ - (rtac iffI 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) - ]); + [ + (rtac iffI 1), + (fast_tac HOL_cs 1), + (fast_tac HOL_cs 1) + ]); qed_goal "stream_finite_lemma6" Stream.thy "!s1 s2. s1 << s2 --> stream_take n`s2 = s2 --> stream_finite(s1)" (fn prems => - [ - (nat_ind_tac "n" 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1 ), - (hyp_subst_tac 1), - (dtac UU_I 1), - (hyp_subst_tac 1), - (rtac stream_finite_UU 1), - (rtac allI 1), - (rtac allI 1), - (res_inst_tac [("s","s1")] streamE 1), - (hyp_subst_tac 1), - (strip_tac 1 ), - (rtac stream_finite_UU 1), - (hyp_subst_tac 1), - (res_inst_tac [("s","s2")] streamE 1), - (hyp_subst_tac 1), - (strip_tac 1 ), - (dtac UU_I 1), - (asm_simp_tac(!simpset addsimps (stream_rews @ [stream_finite_UU])) 1), - (hyp_subst_tac 1), - (simp_tac (!simpset addsimps stream_rews) 1), - (strip_tac 1 ), - (rtac stream_finite_lemma1 1), - (subgoal_tac "xs << xsa" 1), - (subgoal_tac "stream_take n1`xsa = xsa" 1), - (fast_tac HOL_cs 1), - (res_inst_tac [("x1.1","xa"),("y1.1","xa")] + [ + (nat_ind_tac "n" 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1 ), + (hyp_subst_tac 1), + (dtac UU_I 1), + (hyp_subst_tac 1), + (rtac stream_finite_UU 1), + (rtac allI 1), + (rtac allI 1), + (res_inst_tac [("s","s1")] streamE 1), + (hyp_subst_tac 1), + (strip_tac 1 ), + (rtac stream_finite_UU 1), + (hyp_subst_tac 1), + (res_inst_tac [("s","s2")] streamE 1), + (hyp_subst_tac 1), + (strip_tac 1 ), + (dtac UU_I 1), + (asm_simp_tac(!simpset addsimps (stream_rews @ [stream_finite_UU])) 1), + (hyp_subst_tac 1), + (simp_tac (!simpset addsimps stream_rews) 1), + (strip_tac 1 ), + (rtac stream_finite_lemma1 1), + (subgoal_tac "xs << xsa" 1), + (subgoal_tac "stream_take n1`xsa = xsa" 1), + (fast_tac HOL_cs 1), + (res_inst_tac [("x1.1","xa"),("y1.1","xa")] ((hd stream_inject) RS conjunct2) 1), - (atac 1), - (atac 1), - (atac 1), - (res_inst_tac [("x1.1","x"),("y1.1","xa")] - ((hd stream_invert) RS conjunct2) 1), - (atac 1), - (atac 1), - (atac 1) - ]); + (atac 1), + (atac 1), + (atac 1), + (res_inst_tac [("x1.1","x"),("y1.1","xa")] + ((hd stream_invert) RS conjunct2) 1), + (atac 1), + (atac 1), + (atac 1) + ]); qed_goal "stream_finite_lemma7" Stream.thy "s1 << s2 --> stream_finite(s2) --> stream_finite(s1)" (fn prems => - [ - (rtac (stream_finite_lemma5 RS iffD1) 1), - (rtac allI 1), - (rtac (stream_finite_lemma6 RS spec RS spec) 1) - ]); + [ + (rtac (stream_finite_lemma5 RS iffD1) 1), + (rtac allI 1), + (rtac (stream_finite_lemma6 RS spec RS spec) 1) + ]); qed_goalw "stream_finite_lemma8" Stream.thy [stream_finite_def] "stream_finite(s) = (? n. iterate n stl s = UU)" (fn prems => - [ - (simp_tac (!simpset addsimps [stream_take_lemma7]) 1) - ]); + [ + (simp_tac (!simpset addsimps [stream_take_lemma7]) 1) + ]); (* ----------------------------------------------------------------------- *) @@ -805,17 +805,17 @@ qed_goalw "adm_not_stream_finite" Stream.thy [adm_def] "adm(%s. ~ stream_finite(s))" (fn prems => - [ - (strip_tac 1 ), - (res_inst_tac [("P1","!i. ~ stream_finite(Y(i))")] classical3 1), - (atac 2), - (subgoal_tac "!i.stream_finite(Y(i))" 1), - (fast_tac HOL_cs 1), - (rtac allI 1), - (rtac (stream_finite_lemma7 RS mp RS mp) 1), - (etac is_ub_thelub 1), - (atac 1) - ]); + [ + (strip_tac 1 ), + (res_inst_tac [("P1","!i. ~ stream_finite(Y(i))")] classical3 1), + (atac 2), + (subgoal_tac "!i.stream_finite(Y(i))" 1), + (fast_tac HOL_cs 1), + (rtac allI 1), + (rtac (stream_finite_lemma7 RS mp RS mp) 1), + (etac is_ub_thelub 1), + (atac 1) + ]); (* ----------------------------------------------------------------------- *) (* alternative prove for admissibility of ~stream_finite *) @@ -827,14 +827,14 @@ qed_goal "adm_not_stream_finite" Stream.thy "adm(%s. ~ stream_finite(s))" (fn prems => - [ - (subgoal_tac "(!s.(~stream_finite(s))=(!n.iterate n stl s ~=UU))" 1), - (etac (adm_cong RS iffD2)1), - (REPEAT(resolve_tac adm_thms 1)), - (rtac cont_iterate2 1), - (rtac allI 1), - (rtac (stream_finite_lemma8 RS ssubst) 1), - (fast_tac HOL_cs 1) - ]); + [ + (subgoal_tac "(!s.(~stream_finite(s))=(!n.iterate n stl s ~=UU))" 1), + (etac (adm_cong RS iffD2)1), + (REPEAT(resolve_tac adm_thms 1)), + (rtac cont_iterate2 1), + (rtac allI 1), + (rtac (stream_finite_lemma8 RS ssubst) 1), + (fast_tac HOL_cs 1) + ]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/HOLCF/explicit_domains/Stream2.ML --- a/src/HOLCF/explicit_domains/Stream2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/HOLCF/explicit_domains/Stream2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ (* ID: $Id$ - Author: Franz Regensburger + Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Lemmas for theory Stream2.thy @@ -13,7 +13,7 @@ (* ------------------------------------------------------------------------- *) val smap_def2 = fix_prover2 Stream2.thy smap_def - "smap = (LAM f s. stream_when`(LAM x l.scons`(f`x) `(smap`f`l)) `s)"; + "smap = (LAM f s. stream_when`(LAM x l.scons`(f`x) `(smap`f`l)) `s)"; (* ------------------------------------------------------------------------- *) @@ -23,21 +23,21 @@ qed_goal "smap1" Stream2.thy "smap`f`UU = UU" (fn prems => - [ - (rtac (smap_def2 RS ssubst) 1), - (simp_tac (!simpset addsimps stream_when) 1) - ]); + [ + (rtac (smap_def2 RS ssubst) 1), + (simp_tac (!simpset addsimps stream_when) 1) + ]); qed_goal "smap2" Stream2.thy - "x~=UU ==> smap`f`(scons`x`xs) = scons `(f`x) `(smap`f`xs)" + "x~=UU ==> smap`f`(scons`x`xs) = scons `(f`x) `(smap`f`xs)" (fn prems => - [ - (cut_facts_tac prems 1), - (rtac trans 1), - (rtac (smap_def2 RS ssubst) 1), - (asm_simp_tac (!simpset addsimps stream_rews) 1), - (rtac refl 1) - ]); + [ + (cut_facts_tac prems 1), + (rtac trans 1), + (rtac (smap_def2 RS ssubst) 1), + (asm_simp_tac (!simpset addsimps stream_rews) 1), + (rtac refl 1) + ]); val stream2_rews = [smap1, smap2]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LCF/LCF.ML --- a/src/LCF/LCF.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LCF/LCF.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: LCF/lcf.ML +(* Title: LCF/lcf.ML ID: $Id$ - Author: Tobias Nipkow + Author: Tobias Nipkow Copyright 1992 University of Cambridge For lcf.thy. Basic lemmas about LCF @@ -53,26 +53,26 @@ fun strip_tac i = REPEAT(rstac [impI,allI] i); val eq_imp_less1 = prove_goal thy "x=y ==> x << y" - (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct1) 1]); + (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct1) 1]); val eq_imp_less2 = prove_goal thy "x=y ==> y << x" - (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct2) 1]); + (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct2) 1]); val less_refl = refl RS eq_imp_less1; val less_anti_sym = prove_goal thy "[| x << y; y << x |] ==> x=y" - (fn prems => [rewrite_goals_tac[eq_def], - REPEAT(rstac(conjI::prems)1)]); + (fn prems => [rewtac eq_def, + REPEAT(rstac(conjI::prems)1)]); val ext = prove_goal thy - "(!!x::'a::cpo. f(x)=(g(x)::'b::cpo)) ==> (%x.f(x))=(%x.g(x))" - (fn [prem] => [REPEAT(rstac[less_anti_sym, less_ext, allI, - prem RS eq_imp_less1, - prem RS eq_imp_less2]1)]); + "(!!x::'a::cpo. f(x)=(g(x)::'b::cpo)) ==> (%x.f(x))=(%x.g(x))" + (fn [prem] => [REPEAT(rstac[less_anti_sym, less_ext, allI, + prem RS eq_imp_less1, + prem RS eq_imp_less2]1)]); val cong = prove_goal thy "[| f=g; x=y |] ==> f(x)=g(y)" - (fn prems => [cut_facts_tac prems 1, etac subst 1, etac subst 1, - rtac refl 1]); + (fn prems => [cut_facts_tac prems 1, etac subst 1, etac subst 1, + rtac refl 1]); val less_ap_term = less_refl RS mono; val less_ap_thm = less_refl RSN (2,mono); @@ -80,24 +80,24 @@ val ap_thm = refl RSN (2,cong); val UU_abs = prove_goal thy "(%x::'a::cpo.UU) = UU" - (fn _ => [rtac less_anti_sym 1, rtac minimal 2, - rtac less_ext 1, rtac allI 1, rtac minimal 1]); + (fn _ => [rtac less_anti_sym 1, rtac minimal 2, + rtac less_ext 1, rtac allI 1, rtac minimal 1]); val UU_app = UU_abs RS sym RS ap_thm; val less_UU = prove_goal thy "x << UU ==> x=UU" - (fn prems=> [rtac less_anti_sym 1,rstac prems 1,rtac minimal 1]); + (fn prems=> [rtac less_anti_sym 1,rstac prems 1,rtac minimal 1]); val tr_induct = prove_goal thy "[| P(UU); P(TT); P(FF) |] ==> ALL b.P(b)" - (fn prems => [rtac allI 1, rtac mp 1, - res_inst_tac[("p","b")]tr_cases 2, - fast_tac (FOL_cs addIs prems) 1]); + (fn prems => [rtac allI 1, rtac mp 1, + res_inst_tac[("p","b")]tr_cases 2, + fast_tac (FOL_cs addIs prems) 1]); val Contrapos = prove_goal thy "(A ==> B) ==> (~B ==> ~A)" - (fn prems => [rtac notI 1, rtac notE 1, rstac prems 1, - rstac prems 1, atac 1]); + (fn prems => [rtac notI 1, rtac notE 1, rstac prems 1, + rstac prems 1, atac 1]); val not_less_imp_not_eq1 = eq_imp_less1 COMP Contrapos; val not_less_imp_not_eq2 = eq_imp_less2 COMP Contrapos; @@ -112,14 +112,14 @@ val COND_cases_iff = (prove_goal thy "ALL b. P(b=>x|y) <-> (b=UU-->P(UU)) & (b=TT-->P(x)) & (b=FF-->P(y))" - (fn _ => [cut_facts_tac [not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU, - not_TT_eq_FF,not_FF_eq_UU,not_FF_eq_TT]1, - rtac tr_induct 1, stac COND_UU 1, stac COND_TT 2, - stac COND_FF 3, REPEAT(fast_tac FOL_cs 1)])) RS spec; + (fn _ => [cut_facts_tac [not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU, + not_TT_eq_FF,not_FF_eq_UU,not_FF_eq_TT]1, + rtac tr_induct 1, stac COND_UU 1, stac COND_TT 2, + stac COND_FF 3, REPEAT(fast_tac FOL_cs 1)])) RS spec; val lemma = prove_goal thy "A<->B ==> B ==> A" - (fn prems => [cut_facts_tac prems 1, rewrite_goals_tac [iff_def], - fast_tac FOL_cs 1]); + (fn prems => [cut_facts_tac prems 1, rewtac iff_def, + fast_tac FOL_cs 1]); val COND_cases = conjI RSN (2,conjI RS (COND_cases_iff RS lemma)); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LCF/ROOT.ML --- a/src/LCF/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LCF/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: LCF/ROOT +(* Title: LCF/ROOT ID: $Id$ - Author: Tobias Nipkow + Author: Tobias Nipkow Copyright 1992 University of Cambridge Adds LCF to a database containing First-Order Logic. @@ -17,4 +17,4 @@ use"pair.ML"; use"fix.ML"; -val LCF_build_completed = (); (*indicate successful build*) +val LCF_build_completed = (); (*indicate successful build*) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LCF/fix.ML --- a/src/LCF/fix.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LCF/fix.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: LCF/fix +(* Title: LCF/fix ID: $Id$ - Author: Tobias Nipkow + Author: Tobias Nipkow Copyright 1992 University of Cambridge Fixedpoint theory @@ -23,11 +23,11 @@ struct val adm_eq = prove_goal LCF.thy "adm(%x.t(x)=(u(x)::'a::cpo))" - (fn _ => [rewrite_goals_tac [eq_def], - REPEAT(rstac[adm_conj,adm_less]1)]); + (fn _ => [rewtac eq_def, + REPEAT(rstac[adm_conj,adm_less]1)]); val adm_not_not = prove_goal LCF.thy "adm(P) ==> adm(%x.~~P(x))" - (fn prems => [simp_tac (LCF_ss addsimps prems) 1]); + (fn prems => [simp_tac (LCF_ss addsimps prems) 1]); val tac = rtac tr_induct 1 THEN REPEAT(simp_tac LCF_ss 1); @@ -44,24 +44,24 @@ val adm_not_eq_tr = prove_goal LCF.thy "ALL p::tr.adm(%x. ~t(x)=p)" (fn _ => [rtac tr_induct 1, REPEAT(simp_tac (LCF_ss addsimps [not_eq_TT,not_eq_FF,not_eq_UU]) 1 THEN - REPEAT(rstac [adm_disj,adm_eq] 1))]) RS spec; + REPEAT(rstac [adm_disj,adm_eq] 1))]) RS spec; val adm_lemmas = [adm_not_free,adm_eq,adm_less,adm_not_less,adm_not_eq_tr, - adm_conj,adm_disj,adm_imp,adm_all]; + adm_conj,adm_disj,adm_imp,adm_all]; fun induct_tac v i = res_inst_tac[("f",v)] induct i THEN - REPEAT(rstac adm_lemmas i); + REPEAT(rstac adm_lemmas i); val least_FIX = prove_goal LCF.thy "f(p) = p ==> FIX(f) << p" - (fn [prem] => [induct_tac "f" 1, rtac minimal 1, strip_tac 1, - stac (prem RS sym) 1, etac less_ap_term 1]); + (fn [prem] => [induct_tac "f" 1, rtac minimal 1, strip_tac 1, + stac (prem RS sym) 1, etac less_ap_term 1]); val lfp_is_FIX = prove_goal LCF.thy - "[| f(p) = p; ALL q. f(q)=q --> p << q |] ==> p = FIX(f)" - (fn [prem1,prem2] => [rtac less_anti_sym 1, - rtac (prem2 RS spec RS mp) 1, rtac FIX_eq 1, - rtac least_FIX 1, rtac prem1 1]); + "[| f(p) = p; ALL q. f(q)=q --> p << q |] ==> p = FIX(f)" + (fn [prem1,prem2] => [rtac less_anti_sym 1, + rtac (prem2 RS spec RS mp) 1, rtac FIX_eq 1, + rtac least_FIX 1, rtac prem1 1]); val ffix = read_instantiate [("f","f::?'a=>?'a")] FIX_eq; val gfix = read_instantiate [("f","g::?'a=>?'a")] FIX_eq; @@ -70,9 +70,9 @@ val FIX_pair = prove_goal LCF.thy " = FIX(%p.)" (fn _ => [rtac lfp_is_FIX 1, simp_tac ss 1, - strip_tac 1, simp_tac (LCF_ss addsimps [PROD_less]) 1, - rtac conjI 1, rtac least_FIX 1, etac subst 1, rtac (FST RS sym) 1, - rtac least_FIX 1, etac subst 1, rtac (SND RS sym) 1]); + strip_tac 1, simp_tac (LCF_ss addsimps [PROD_less]) 1, + rtac conjI 1, rtac least_FIX 1, etac subst 1, rtac (FST RS sym) 1, + rtac least_FIX 1, etac subst 1, rtac (SND RS sym) 1]); val FIX_pair_conj = rewrite_rule (map mk_meta_eq [PROD_eq,FST,SND]) FIX_pair; @@ -80,17 +80,17 @@ val FIX2 = FIX_pair_conj RS conjunct2; val induct2 = prove_goal LCF.thy - "[| adm(%p.P(FST(p),SND(p))); P(UU::'a,UU::'b);\ -\ ALL x y. P(x,y) --> P(f(x),g(y)) |] ==> P(FIX(f),FIX(g))" - (fn prems => [EVERY1 - [res_inst_tac [("f","f"),("g","g")] (standard(FIX1 RS ssubst)), - res_inst_tac [("f","f"),("g","g")] (standard(FIX2 RS ssubst)), - res_inst_tac [("f","%x. ")] induct, - rstac prems, simp_tac ss, rstac prems, - simp_tac (LCF_ss addsimps [expand_all_PROD]), rstac prems]]); + "[| adm(%p.P(FST(p),SND(p))); P(UU::'a,UU::'b);\ +\ ALL x y. P(x,y) --> P(f(x),g(y)) |] ==> P(FIX(f),FIX(g))" + (fn prems => [EVERY1 + [res_inst_tac [("f","f"),("g","g")] (standard(FIX1 RS ssubst)), + res_inst_tac [("f","f"),("g","g")] (standard(FIX2 RS ssubst)), + res_inst_tac [("f","%x. ")] induct, + rstac prems, simp_tac ss, rstac prems, + simp_tac (LCF_ss addsimps [expand_all_PROD]), rstac prems]]); fun induct2_tac (f,g) i = res_inst_tac[("f",f),("g",g)] induct2 i THEN - REPEAT(rstac adm_lemmas i); + REPEAT(rstac adm_lemmas i); end; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LCF/pair.ML --- a/src/LCF/pair.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LCF/pair.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,41 +1,41 @@ -(* Title: LCF/pair +(* Title: LCF/pair ID: $Id$ - Author: Tobias Nipkow + Author: Tobias Nipkow Copyright 1992 University of Cambridge Theory of ordered pairs and products *) val expand_all_PROD = prove_goal LCF.thy - "(ALL p. P(p)) <-> (ALL x y. P())" - (fn _ => [rtac iffI 1, fast_tac FOL_cs 1, rtac allI 1, - rtac (surj_pairing RS subst) 1, fast_tac FOL_cs 1]); + "(ALL p. P(p)) <-> (ALL x y. P())" + (fn _ => [rtac iffI 1, fast_tac FOL_cs 1, rtac allI 1, + rtac (surj_pairing RS subst) 1, fast_tac FOL_cs 1]); local val ppair = read_instantiate [("z","p::'a*'b")] surj_pairing; val qpair = read_instantiate [("z","q::'a*'b")] surj_pairing; in val PROD_less = prove_goal LCF.thy - "(p::'a*'b) << q <-> FST(p) << FST(q) & SND(p) << SND(q)" - (fn _ => [EVERY1[rtac iffI, - rtac conjI, etac less_ap_term, etac less_ap_term, - rtac (ppair RS subst), rtac (qpair RS subst), - etac conjE, rtac mono, etac less_ap_term, atac]]); + "(p::'a*'b) << q <-> FST(p) << FST(q) & SND(p) << SND(q)" + (fn _ => [EVERY1[rtac iffI, + rtac conjI, etac less_ap_term, etac less_ap_term, + rtac (ppair RS subst), rtac (qpair RS subst), + etac conjE, rtac mono, etac less_ap_term, atac]]); end; val PROD_eq = prove_goal LCF.thy "p=q <-> FST(p)=FST(q) & SND(p)=SND(q)" - (fn _ => [rtac iffI 1, asm_simp_tac LCF_ss 1, - rewrite_goals_tac [eq_def], - asm_simp_tac (LCF_ss addsimps [PROD_less]) 1]); + (fn _ => [rtac iffI 1, asm_simp_tac LCF_ss 1, + rewtac eq_def, + asm_simp_tac (LCF_ss addsimps [PROD_less]) 1]); val PAIR_less = prove_goal LCF.thy " << <-> a< [simp_tac (LCF_ss addsimps [PROD_less])1]); + (fn _ => [simp_tac (LCF_ss addsimps [PROD_less])1]); val PAIR_eq = prove_goal LCF.thy " = <-> a=c & b=d" - (fn _ => [simp_tac (LCF_ss addsimps [PROD_eq])1]); + (fn _ => [simp_tac (LCF_ss addsimps [PROD_eq])1]); val UU_is_UU_UU = prove_goal LCF.thy " << UU" - (fn _ => [simp_tac (LCF_ss addsimps [PROD_less]) 1]) - RS less_UU RS sym; + (fn _ => [simp_tac (LCF_ss addsimps [PROD_less]) 1]) + RS less_UU RS sym; val LCF_ss = LCF_ss addsimps [PAIR_less,PAIR_eq,UU_is_UU_UU]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LCF/simpdata.ML --- a/src/LCF/simpdata.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LCF/simpdata.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: LCF/simpdata +(* Title: LCF/simpdata ID: $Id$ - Author: Tobias Nipkow, Cambridge University Computer Laboratory + Author: Tobias Nipkow, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Simplification data for LCF @@ -8,9 +8,9 @@ val LCF_ss = FOL_ss addsimps [minimal, - UU_app, UU_app RS ap_thm, UU_app RS ap_thm RS ap_thm, - not_TT_less_FF,not_FF_less_TT,not_TT_less_UU,not_FF_less_UU, - not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU,not_TT_eq_FF, - not_FF_eq_UU,not_FF_eq_TT, - COND_UU,COND_TT,COND_FF, - surj_pairing,FST,SND]; + UU_app, UU_app RS ap_thm, UU_app RS ap_thm RS ap_thm, + not_TT_less_FF,not_FF_less_TT,not_TT_less_UU,not_FF_less_UU, + not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU,not_TT_eq_FF, + not_FF_eq_UU,not_FF_eq_TT, + COND_UU,COND_TT,COND_FF, + surj_pairing,FST,SND]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LK/LK.ML --- a/src/LK/LK.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LK/LK.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: LK/lk.ML +(* Title: LK/lk.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge Tactics and lemmas for lk.thy (thanks also to Philippe de Groote) @@ -74,14 +74,14 @@ (*The rules of LK*) val prop_pack = empty_pack add_safes - [basic, refl, conjL, conjR, disjL, disjR, impL, impR, - notL, notR, iffL, iffR]; + [basic, refl, conjL, conjR, disjL, disjR, impL, impR, + notL, notR, iffL, iffR]; val LK_pack = prop_pack add_safes [allR, exL] - add_unsafes [allL_thin, exR_thin]; + add_unsafes [allL_thin, exR_thin]; val LK_dup_pack = prop_pack add_safes [allR, exL] - add_unsafes [allL, exR]; + add_unsafes [allL, exR]; @@ -105,7 +105,7 @@ case (prem,conc) of (_ $ Abs(_,_,leftp) $ Abs(_,_,rightp), _ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) => - could_res (leftp,leftc) andalso could_res (rightp,rightc) + could_res (leftp,leftc) andalso could_res (rightp,rightc) | _ => false; @@ -134,12 +134,12 @@ fun RESOLVE_THEN rules = let val [rls0,rls1,rls2] = partition_list has_prems 0 2 rules; fun tac nextac i = STATE (fn state => - filseq_resolve_tac rls0 9999 i - ORELSE - (DETERM(filseq_resolve_tac rls1 9999 i) THEN TRY(nextac i)) - ORELSE - (DETERM(filseq_resolve_tac rls2 9999 i) THEN TRY(nextac(i+1)) - THEN TRY(nextac i)) ) + filseq_resolve_tac rls0 9999 i + ORELSE + (DETERM(filseq_resolve_tac rls1 9999 i) THEN TRY(nextac i)) + ORELSE + (DETERM(filseq_resolve_tac rls2 9999 i) THEN TRY(nextac(i+1)) + THEN TRY(nextac i)) ) in tac end; @@ -179,7 +179,7 @@ fun best_tac thm_pack = SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, size_of_thm) - (step_tac thm_pack 1)); + (step_tac thm_pack 1)); (** Contraction. Useful since some rules are not complete. **) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LK/ROOT.ML --- a/src/LK/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LK/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: LK/ROOT +(* Title: LK/ROOT ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1991 University of Cambridge Adds Classical Sequent Calculus to a database containing pure Isabelle. @@ -19,4 +19,4 @@ use "../Pure/install_pp.ML"; print_depth 8; -val LK_build_completed = (); (*indicate successful build*) +val LK_build_completed = (); (*indicate successful build*) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LK/ex/ROOT.ML --- a/src/LK/ex/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LK/ex/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: LK/ex/ROOT +(* Title: LK/ex/ROOT ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge Executes all examples for Classical Logic. diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LK/ex/hardquant.ML --- a/src/LK/ex/hardquant.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LK/ex/hardquant.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: LK/ex/hard-quant +(* Title: LK/ex/hard-quant ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge Hard examples with quantifiers. Can be read to test the LK system. @@ -96,8 +96,8 @@ result(); writeln"Problem 26"; -goal LK.thy "|- ((EX x. p(x)) <-> (EX x. q(x))) & \ -\ (ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y))) \ +goal LK.thy "|- ((EX x. p(x)) <-> (EX x. q(x))) & \ +\ (ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y))) \ \ --> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))"; by (pc_tac LK_pack 1); result(); @@ -157,12 +157,12 @@ writeln"Problem 34 AMENDED (TWICE!!) NOT PROVED AUTOMATICALLY"; (*Andrews's challenge*) -goal LK.thy "|- ((EX x. ALL y. p(x) <-> p(y)) <-> \ -\ ((EX x. q(x)) <-> (ALL y. p(y)))) <-> \ -\ ((EX x. ALL y. q(x) <-> q(y)) <-> \ +goal LK.thy "|- ((EX x. ALL y. p(x) <-> p(y)) <-> \ +\ ((EX x. q(x)) <-> (ALL y. p(y)))) <-> \ +\ ((EX x. ALL y. q(x) <-> q(y)) <-> \ \ ((EX x. p(x)) <-> (ALL y. q(y))))"; by (safe_goal_tac LK_pack 1); (*53 secs*) (*13 secs*) -by (TRYALL (fast_tac LK_pack)); (*165 secs*) (*117 secs*) (*138 secs*) +by (TRYALL (fast_tac LK_pack)); (*165 secs*) (*117 secs*) (*138 secs*) (*for some reason, pc_tac leaves 14 subgoals instead of 6*) by (TRYALL (best_tac LK_dup_pack)); (*55 secs*) (*29 secs*) (*54 secs*) result(); @@ -176,7 +176,7 @@ writeln"Problem 36"; goal LK.thy "|- (ALL x. EX y. J(x,y)) & \ \ (ALL x. EX y. G(x,y)) & \ -\ (ALL x y. J(x,y) | G(x,y) --> \ +\ (ALL x y. J(x,y) | G(x,y) --> \ \ (ALL z. J(y,z) | G(y,z) --> H(x,z))) \ \ --> (ALL x. EX y. H(x,y))"; by (fast_tac LK_pack 1); @@ -193,10 +193,10 @@ writeln"Problem 38"; goal LK.thy - "|- (ALL x. p(a) & (p(x) --> (EX y. p(y) & r(x,y))) --> \ -\ (EX z. EX w. p(z) & r(x,w) & r(w,z))) <-> \ -\ (ALL x. (~p(a) | p(x) | (EX z. EX w. p(z) & r(x,w) & r(w,z))) & \ -\ (~p(a) | ~(EX y. p(y) & r(x,y)) | \ + "|- (ALL x. p(a) & (p(x) --> (EX y. p(y) & r(x,y))) --> \ +\ (EX z. EX w. p(z) & r(x,w) & r(w,z))) <-> \ +\ (ALL x. (~p(a) | p(x) | (EX z. EX w. p(z) & r(x,w) & r(w,z))) & \ +\ (~p(a) | ~(EX y. p(y) & r(x,y)) | \ \ (EX z. EX w. p(z) & r(x,w) & r(w,z))))"; by (pc_tac LK_pack 1); result(); @@ -213,7 +213,7 @@ result(); writeln"Problem 41"; -goal LK.thy "|- (ALL z. EX y. ALL x. f(x,y) <-> f(x,z) & ~ f(x,x)) \ +goal LK.thy "|- (ALL z. EX y. ALL x. f(x,y) <-> f(x,z) & ~ f(x,x)) \ \ --> ~ (EX z. ALL x. f(x,z))"; by (fast_tac LK_pack 1); result(); @@ -226,19 +226,19 @@ \ --> (ALL x. (ALL y. q(x,y) <-> q(y,x)))"; writeln"Problem 44"; -goal LK.thy "|- (ALL x. f(x) --> \ -\ (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) & \ -\ (EX x. j(x) & (ALL y. g(y) --> h(x,y))) \ +goal LK.thy "|- (ALL x. f(x) --> \ +\ (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) & \ +\ (EX x. j(x) & (ALL y. g(y) --> h(x,y))) \ \ --> (EX x. j(x) & ~f(x))"; by (fast_tac LK_pack 1); result(); writeln"Problem 45"; -goal LK.thy "|- (ALL x. f(x) & (ALL y. g(y) & h(x,y) --> j(x,y)) \ -\ --> (ALL y. g(y) & h(x,y) --> k(y))) & \ -\ ~ (EX y. l(y) & k(y)) & \ -\ (EX x. f(x) & (ALL y. h(x,y) --> l(y)) \ -\ & (ALL y. g(y) & h(x,y) --> j(x,y))) \ +goal LK.thy "|- (ALL x. f(x) & (ALL y. g(y) & h(x,y) --> j(x,y)) \ +\ --> (ALL y. g(y) & h(x,y) --> k(y))) & \ +\ ~ (EX y. l(y) & k(y)) & \ +\ (EX x. f(x) & (ALL y. h(x,y) --> l(y)) \ +\ & (ALL y. g(y) & h(x,y) --> j(x,y))) \ \ --> (EX x. f(x) & ~ (EX y. g(y) & h(x,y)))"; by (best_tac LK_pack 1); result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LK/ex/prop.ML --- a/src/LK/ex/prop.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LK/ex/prop.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: LK/ex/prop +(* Title: LK/ex/prop ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge Classical sequent calculus: examples with propositional connectives diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/LK/ex/quant.ML --- a/src/LK/ex/quant.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/LK/ex/quant.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: LK/ex/quant +(* Title: LK/ex/quant ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge Classical sequent calculus: examples with quantifiers. diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/Modal/ROOT.ML --- a/src/Modal/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/Modal/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: 91/Modal/ROOT +(* Title: 91/Modal/ROOT ID: $Id$ - Author: Martin Coen + Author: Martin Coen Copyright 1991 University of Cambridge *) @@ -18,13 +18,13 @@ val iffR = prove_goal thy "[| $H,P |- $E,Q,$F; $H,Q |- $E,P,$F |] ==> $H |- $E, P <-> Q, $F" (fn prems=> - [ (rewrite_goals_tac [iff_def]), + [ (rewtac iff_def), (REPEAT (resolve_tac (prems@[conjR,impR]) 1)) ]); val iffL = prove_goal thy "[| $H,$G |- $E,P,Q; $H,Q,P,$G |- $E |] ==> $H, P <-> Q, $G |- $E" (fn prems=> - [ rewrite_goals_tac [iff_def], + [ rewtac iff_def, (REPEAT (resolve_tac (prems@[conjL,impL,basic]) 1)) ]) in val safe_rls = [basic,conjL,conjR,disjL,disjR,impL,impR,notL,notR,iffL,iffR]; @@ -78,6 +78,6 @@ end; structure S43_Prover = Modal_ProverFun(MP_Rule); -val Modal_build_completed = (); (*indicate successful build*) +val Modal_build_completed = (); (*indicate successful build*) (*printing functions are inherited from LK*) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/Modal/ex/ROOT.ML --- a/src/Modal/ex/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/Modal/ex/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: Modal/ex/ROOT +(* Title: Modal/ex/ROOT ID: $Id$ - Author: Martin Coen + Author: Martin Coen Copyright 1991 University of Cambridge *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/Modal/ex/S43thms.ML --- a/src/Modal/ex/S43thms.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/Modal/ex/S43thms.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: 91/Modal/ex/S43thms +(* Title: 91/Modal/ex/S43thms ID: $Id$ - Author: Martin Coen + Author: Martin Coen Copyright 1991 University of Cambridge *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/Modal/ex/S4thms.ML --- a/src/Modal/ex/S4thms.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/Modal/ex/S4thms.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: 91/Modal/ex/S4thms +(* Title: 91/Modal/ex/S4thms ID: $Id$ - Author: Martin Coen + Author: Martin Coen Copyright 1991 University of Cambridge *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/Modal/ex/Tthms.ML --- a/src/Modal/ex/Tthms.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/Modal/ex/Tthms.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: 91/Modal/ex/Tthms +(* Title: 91/Modal/ex/Tthms ID: $Id$ - Author: Martin Coen + Author: Martin Coen Copyright 1991 University of Cambridge *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/Modal/prover.ML --- a/src/Modal/prover.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/Modal/prover.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: 91/Modal/prover +(* Title: 91/Modal/prover ID: $Id$ - Author: Martin Coen + Author: Martin Coen Copyright 1991 University of Cambridge *) @@ -45,7 +45,7 @@ case (prem,conc) of (_ $ Abs(_,_,leftp) $ Abs(_,_,rightp), _ $ Abs(_,_,leftc) $ Abs(_,_,rightc)) => - could_res (leftp,leftc) andalso could_res (rightp,rightc) + could_res (leftp,leftc) andalso could_res (rightp,rightc) | _ => false; (*Like filt_resolve_tac, using could_resolve_seq @@ -75,16 +75,16 @@ fun solven_tac d n = STATE (fn state => if d<0 then no_tac else if (nprems_of state = 0) then all_tac - else (DETERM(fres_safe_tac n) THEN UPTOGOAL n (solven_tac d)) ORELSE - ((fres_unsafe_tac n THEN UPTOGOAL n (solven_tac d)) APPEND - (fres_bound_tac n THEN UPTOGOAL n (solven_tac (d-1))))); + else (DETERM(fres_safe_tac n) THEN UPTOGOAL n (solven_tac d)) ORELSE + ((fres_unsafe_tac n THEN UPTOGOAL n (solven_tac d)) APPEND + (fres_bound_tac n THEN UPTOGOAL n (solven_tac (d-1))))); fun solve_tac d = rewrite_goals_tac rewrite_rls THEN solven_tac d 1; fun step_tac n = STATE (fn state => if (nprems_of state = 0) then all_tac else (DETERM(fres_safe_tac n)) ORELSE - (fres_unsafe_tac n APPEND fres_bound_tac n)); + (fres_unsafe_tac n APPEND fres_bound_tac n)); end; end; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC.ML --- a/src/ZF/AC.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC.ML +(* Title: ZF/AC.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge For AC.thy. The Axiom of Choice @@ -19,8 +19,8 @@ (*Using dtac, this has the advantage of DELETING the universal quantifier*) goal AC.thy "!!A B. ALL x:A. EX y. y:B(x) ==> EX y. y : Pi(A,B)"; -by (resolve_tac [AC_Pi] 1); -by (eresolve_tac [bspec] 1); +by (rtac AC_Pi 1); +by (etac bspec 1); by (assume_tac 1); qed "AC_ball_Pi"; @@ -31,7 +31,7 @@ qed "AC_Pi_Pow"; val [nonempty] = goal AC.thy - "[| !!x. x:A ==> (EX y. y:x) \ + "[| !!x. x:A ==> (EX y. y:x) \ \ |] ==> EX f: A->Union(A). ALL x:A. f`x : x"; by (res_inst_tac [("B1", "%x.x")] (AC_Pi RS exE) 1); by (etac nonempty 1); @@ -52,7 +52,7 @@ by (resolve_tac [AC_func0 RS bexE] 1); by (rtac bexI 2); by (assume_tac 2); -by (eresolve_tac [fun_weaken_type] 2); +by (etac fun_weaken_type 2); by (ALLGOALS (fast_tac ZF_cs)); qed "AC_func_Pow"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC0_AC1.ML --- a/src/ZF/AC/AC0_AC1.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC0_AC1.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC0_AC1.ML +(* Title: ZF/AC/AC0_AC1.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski AC0 is equivalent to AC1 AC0 comes from Suppes, AC1 from Rubin & Rubin diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC10_AC15.ML --- a/src/ZF/AC/AC10_AC15.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC10_AC15.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC10_AC15.ML +(* Title: ZF/AC/AC10_AC15.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proofs needed to state that AC10, ..., AC15 are equivalent to the rest. We need the following: @@ -21,10 +21,10 @@ *) (* ********************************************************************** *) -(* Lemmas used in the proofs in which the conclusion is AC13, AC14 *) -(* or AC15 *) -(* - cons_times_nat_not_Finite *) -(* - ex_fun_AC13_AC15 *) +(* Lemmas used in the proofs in which the conclusion is AC13, AC14 *) +(* or AC15 *) +(* - cons_times_nat_not_Finite *) +(* - ex_fun_AC13_AC15 *) (* ********************************************************************** *) goalw thy [lepoll_def] "!!A. A~=0 ==> B lepoll A*B"; @@ -40,7 +40,7 @@ by (rtac notI 1); by (dresolve_tac [subset_consI RS subset_imp_lepoll RS lepoll_Finite] 1); by (resolve_tac [lepoll_Sigma RS lepoll_Finite RS (nat_not_Finite RS notE)] 1 - THEN (assume_tac 2)); + THEN (assume_tac 2)); by (fast_tac AC_cs 1); val cons_times_nat_not_Finite = result(); @@ -49,17 +49,17 @@ val lemma1 = result(); goalw thy [pairwise_disjoint_def] - "!!A. [| pairwise_disjoint(A); B:A; C:A; a:B; a:C |] ==> B=C"; + "!!A. [| pairwise_disjoint(A); B:A; C:A; a:B; a:C |] ==> B=C"; by (dtac IntI 1 THEN (assume_tac 1)); by (dres_inst_tac [("A","B Int C")] not_emptyI 1); by (fast_tac ZF_cs 1); val lemma2 = result(); goalw thy [sets_of_size_between_def] - "!!A. ALL B:{cons(0, x*nat). x:A}. pairwise_disjoint(f`B) & \ -\ sets_of_size_between(f`B, 2, n) & Union(f`B)=B \ -\ ==> ALL B:A. EX! u. u:f`cons(0, B*nat) & u <= cons(0, B*nat) & \ -\ 0:u & 2 lepoll u & u lepoll n"; + "!!A. ALL B:{cons(0, x*nat). x:A}. pairwise_disjoint(f`B) & \ +\ sets_of_size_between(f`B, 2, n) & Union(f`B)=B \ +\ ==> ALL B:A. EX! u. u:f`cons(0, B*nat) & u <= cons(0, B*nat) & \ +\ 0:u & 2 lepoll u & u lepoll n"; by (rtac ballI 1); by (etac ballE 1); by (fast_tac ZF_cs 2); @@ -75,12 +75,12 @@ goalw thy [lepoll_def] "!!A. [| A lepoll i; Ord(i) |] ==> {P(a). a:A} lepoll i"; by (etac exE 1); by (res_inst_tac - [("x", "lam x:RepFun(A, P). LEAST j. EX a:A. x=P(a) & f`a=j")] exI 1); + [("x", "lam x:RepFun(A, P). LEAST j. EX a:A. x=P(a) & f`a=j")] exI 1); by (res_inst_tac [("d", "%y. P(converse(f)`y)")] lam_injective 1); by (etac RepFunE 1); by (forward_tac [inj_is_fun RS apply_type] 1 THEN (assume_tac 1)); by (fast_tac (AC_cs addIs [LeastI2] - addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1); + addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1); by (etac RepFunE 1); by (rtac LeastI2 1); by (fast_tac AC_cs 1); @@ -89,38 +89,38 @@ val lemma4 = result(); goal thy "!!A. [| n:nat; B:A; u(B) <= cons(0, B*nat); 0:u(B); 2 lepoll u(B); \ -\ u(B) lepoll succ(n) |] \ -\ ==> (lam x:A. {fst(x). x:u(x)-{0}})`B ~= 0 & \ -\ (lam x:A. {fst(x). x:u(x)-{0}})`B <= B & \ -\ (lam x:A. {fst(x). x:u(x)-{0}})`B lepoll n"; +\ u(B) lepoll succ(n) |] \ +\ ==> (lam x:A. {fst(x). x:u(x)-{0}})`B ~= 0 & \ +\ (lam x:A. {fst(x). x:u(x)-{0}})`B <= B & \ +\ (lam x:A. {fst(x). x:u(x)-{0}})`B lepoll n"; by (asm_simp_tac AC_ss 1); by (rtac conjI 1); by (fast_tac (empty_cs addSDs [RepFun_eq_0_iff RS iffD1] - addDs [lepoll_Diff_sing] - addEs [lepoll_trans RS succ_lepoll_natE, ssubst] - addSIs [notI, lepoll_refl, nat_0I]) 1); + addDs [lepoll_Diff_sing] + addEs [lepoll_trans RS succ_lepoll_natE, ssubst] + addSIs [notI, lepoll_refl, nat_0I]) 1); by (rtac conjI 1); by (fast_tac (ZF_cs addSIs [fst_type] addSEs [consE]) 1); by (fast_tac (ZF_cs addSEs [equalityE, - Diff_lepoll RS (nat_into_Ord RSN (2, lemma4))]) 1); + Diff_lepoll RS (nat_into_Ord RSN (2, lemma4))]) 1); val lemma5 = result(); goal thy "!!A. [| EX f. ALL B:{cons(0, x*nat). x:A}. \ -\ pairwise_disjoint(f`B) & \ -\ sets_of_size_between(f`B, 2, succ(n)) & \ -\ Union(f`B)=B; n:nat |] \ -\ ==> EX f. ALL B:A. f`B ~= 0 & f`B <= B & f`B lepoll n"; +\ pairwise_disjoint(f`B) & \ +\ sets_of_size_between(f`B, 2, succ(n)) & \ +\ Union(f`B)=B; n:nat |] \ +\ ==> EX f. ALL B:A. f`B ~= 0 & f`B <= B & f`B lepoll n"; by (fast_tac (empty_cs addSDs [lemma3, theI] addDs [bspec] - addSEs [exE, conjE] - addIs [exI, ballI, lemma5]) 1); + addSEs [exE, conjE] + addIs [exI, ballI, lemma5]) 1); val ex_fun_AC13_AC15 = result(); (* ********************************************************************** *) -(* The target proofs *) +(* The target proofs *) (* ********************************************************************** *) (* ********************************************************************** *) -(* AC10(n) ==> AC11 *) +(* AC10(n) ==> AC11 *) (* ********************************************************************** *) goalw thy AC_defs "!!Z. [| n:nat; 1 le n; AC10(n) |] ==> AC11"; @@ -129,7 +129,7 @@ qed "AC10_AC11"; (* ********************************************************************** *) -(* AC11 ==> AC12 *) +(* AC11 ==> AC12 *) (* ********************************************************************** *) goalw thy AC_defs "!! Z. AC11 ==> AC12"; @@ -137,7 +137,7 @@ qed "AC11_AC12"; (* ********************************************************************** *) -(* AC12 ==> AC15 *) +(* AC12 ==> AC15 *) (* ********************************************************************** *) goalw thy AC_defs "!!Z. AC12 ==> AC15"; @@ -149,35 +149,35 @@ qed "AC12_AC15"; (* ********************************************************************** *) -(* AC15 ==> WO6 *) +(* AC15 ==> WO6 *) (* ********************************************************************** *) (* in a separate file *) (* ********************************************************************** *) -(* The proof needed in the first case, not in the second *) +(* The proof needed in the first case, not in the second *) (* ********************************************************************** *) (* ********************************************************************** *) -(* AC10(n) ==> AC13(n-1) if 2 le n *) -(* *) -(* Because of the change to the formal definition of AC10(n) we prove *) -(* the following obviously equivalent theorem : *) -(* AC10(n) implies AC13(n) for (1 le n) *) +(* AC10(n) ==> AC13(n-1) if 2 le n *) +(* *) +(* Because of the change to the formal definition of AC10(n) we prove *) +(* the following obviously equivalent theorem : *) +(* AC10(n) implies AC13(n) for (1 le n) *) (* ********************************************************************** *) goalw thy AC_defs "!!n. [| n:nat; 1 le n; AC10(n) |] ==> AC13(n)"; by (safe_tac ZF_cs); by (fast_tac (empty_cs addSEs [allE, cons_times_nat_not_Finite RSN (2, impE), - ex_fun_AC13_AC15]) 1); + ex_fun_AC13_AC15]) 1); qed "AC10_AC13"; (* ********************************************************************** *) -(* The proofs needed in the second case, not in the first *) +(* The proofs needed in the second case, not in the first *) (* ********************************************************************** *) (* ********************************************************************** *) -(* AC1 ==> AC13(1) *) +(* AC1 ==> AC13(1) *) (* ********************************************************************** *) goalw thy AC_defs "!!Z. AC1 ==> AC13(1)"; @@ -188,13 +188,13 @@ by (etac exE 1); by (res_inst_tac [("x","lam x:A. {f`x}")] exI 1); by (asm_full_simp_tac (AC_ss addsimps - [singleton_eqpoll_1 RS eqpoll_imp_lepoll, - singletonI RS not_emptyI]) 1); + [singleton_eqpoll_1 RS eqpoll_imp_lepoll, + singletonI RS not_emptyI]) 1); by (fast_tac (AC_cs addSEs [singletonE, apply_type]) 1); qed "AC1_AC13"; (* ********************************************************************** *) -(* AC13(m) ==> AC13(n) for m <= n *) +(* AC13(m) ==> AC13(n) for m <= n *) (* ********************************************************************** *) goalw thy AC_defs "!!m n. [| m:nat; n:nat; m le n; AC13(m) |] ==> AC13(n)"; @@ -203,11 +203,11 @@ qed "AC13_mono"; (* ********************************************************************** *) -(* The proofs necessary for both cases *) +(* The proofs necessary for both cases *) (* ********************************************************************** *) (* ********************************************************************** *) -(* AC13(n) ==> AC14 if 1 <= n *) +(* AC13(n) ==> AC14 if 1 <= n *) (* ********************************************************************** *) goalw thy AC_defs "!!n. [| n:nat; 1 le n; AC13(n) |] ==> AC14"; @@ -215,7 +215,7 @@ qed "AC13_AC14"; (* ********************************************************************** *) -(* AC14 ==> AC15 *) +(* AC14 ==> AC15 *) (* ********************************************************************** *) goalw thy AC_defs "!!Z. AC14 ==> AC15"; @@ -223,11 +223,11 @@ qed "AC14_AC15"; (* ********************************************************************** *) -(* The redundant proofs; however cited by Rubin & Rubin *) +(* The redundant proofs; however cited by Rubin & Rubin *) (* ********************************************************************** *) (* ********************************************************************** *) -(* AC13(1) ==> AC1 *) +(* AC13(1) ==> AC1 *) (* ********************************************************************** *) goal thy "!!A. [| A~=0; A lepoll 1 |] ==> EX a. A={a}"; @@ -235,7 +235,7 @@ val lemma_aux = result(); goal thy "!!f. ALL B:A. f(B)~=0 & f(B)<=B & f(B) lepoll 1 \ -\ ==> (lam x:A. THE y. f(x)={y}) : (PROD X:A. X)"; +\ ==> (lam x:A. THE y. f(x)={y}) : (PROD X:A. X)"; by (rtac lam_type 1); by (dtac bspec 1 THEN (assume_tac 1)); by (REPEAT (etac conjE 1)); @@ -249,7 +249,7 @@ qed "AC13_AC1"; (* ********************************************************************** *) -(* AC11 ==> AC14 *) +(* AC11 ==> AC14 *) (* ********************************************************************** *) goalw thy [AC11_def, AC14_def] "!!Z. AC11 ==> AC14"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC15_WO6.ML --- a/src/ZF/AC/AC15_WO6.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC15_WO6.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC15_WO6.ML +(* Title: ZF/AC/AC15_WO6.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proof of AC1 ==> WO2 *) @@ -12,22 +12,22 @@ val OUN_eq_UN = result(); val [prem] = goal thy "ALL x:Pow(A)-{0}. f`x~=0 & f`x<=x & f`x lepoll m ==> \ -\ (UN i \ -\ ALL x WO6"; @@ -42,5 +42,5 @@ by (res_inst_tac [("x","lam j: (LEAST i. HH(f,A,i)={A}). HH(f,A,j)")] exI 1); by (asm_full_simp_tac AC_ss 1); by (fast_tac (AC_cs addSIs [Ord_Least, lam_type RS domain_of_fun] - addSEs [less_Least_subset_x, lemma1, lemma2]) 1); + addSEs [less_Least_subset_x, lemma1, lemma2]) 1); qed "AC15_WO6"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC16_WO4.ML --- a/src/ZF/AC/AC16_WO4.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC16_WO4.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC16_WO4.ML +(* Title: ZF/AC/AC16_WO4.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proof of AC16(n, k) ==> WO4(n-k) *) @@ -8,12 +8,12 @@ open AC16_WO4; (* ********************************************************************** *) -(* The case of finite set *) +(* The case of finite set *) (* ********************************************************************** *) goalw thy [Finite_def] "!!A. [| Finite(A); 0 \ -\ EX a f. Ord(a) & domain(f) = a & \ -\ (UN b well_ord({{y,z}. y:x},?s(x,z))"; @@ -40,68 +40,68 @@ val lepoll_trans1 = result(); goalw thy [lepoll_def] - "!!X.[| Y lepoll X; well_ord(X, R) |] ==> EX S. well_ord(Y, S)"; + "!!X.[| Y lepoll X; well_ord(X, R) |] ==> EX S. well_ord(Y, S)"; by (fast_tac (AC_cs addSEs [well_ord_rvimage]) 1); val well_ord_lepoll = result(); goal thy "!!X. [| well_ord(X,R); well_ord(Y,S) \ -\ |] ==> EX T. well_ord(X Un Y, T)"; +\ |] ==> EX T. well_ord(X Un Y, T)"; by (eresolve_tac [well_ord_radd RS (Un_lepoll_sum RS well_ord_lepoll)] 1); by (assume_tac 1); val well_ord_Un = result(); (* ********************************************************************** *) -(* There exists a well ordered set y such that ... *) +(* There exists a well ordered set y such that ... *) (* ********************************************************************** *) goal thy "EX y R. well_ord(y,R) & x Int y = 0 & ~y lepoll z & ~Finite(y)"; by (res_inst_tac [("x","{{a,x}. a:nat Un Hartog(z)}")] exI 1); by (resolve_tac [Ord_nat RS well_ord_Memrel RS (Ord_Hartog RS - well_ord_Memrel RSN (2, well_ord_Un)) RS exE] 1); + well_ord_Memrel RSN (2, well_ord_Un)) RS exE] 1); by (fast_tac (AC_cs addSIs [Ord_Hartog, well_ord_Memrel, well_ord_paired, - equals0I, HartogI RSN (2, lepoll_trans1), - subset_imp_lepoll RS (paired_eqpoll RS eqpoll_sym RS - eqpoll_imp_lepoll RSN (2, lepoll_trans))] - addSEs [RepFunE, nat_not_Finite RS notE] addEs [mem_asym] - addSDs [Un_upper1 RS subset_imp_lepoll RS lepoll_Finite, - paired_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll - RS lepoll_Finite]) 1); + equals0I, HartogI RSN (2, lepoll_trans1), + subset_imp_lepoll RS (paired_eqpoll RS eqpoll_sym RS + eqpoll_imp_lepoll RSN (2, lepoll_trans))] + addSEs [RepFunE, nat_not_Finite RS notE] addEs [mem_asym] + addSDs [Un_upper1 RS subset_imp_lepoll RS lepoll_Finite, + paired_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll + RS lepoll_Finite]) 1); val lemma2 = result(); val [prem] = goal thy "~Finite(B) ==> ~Finite(A Un B)"; by (fast_tac (AC_cs - addSIs [subset_imp_lepoll RS (prem RSN (2, lepoll_infinite))]) 1); + addSIs [subset_imp_lepoll RS (prem RSN (2, lepoll_infinite))]) 1); val infinite_Un = result(); (* ********************************************************************** *) -(* There is a v : s_u such that k lepoll x Int y (in our case succ(k)) *) -(* The idea of the proof is the following : *) -(* Suppose not, i.e. every element of s_u has exactly k-1 elements of y *) -(* Thence y is less than or equipollent to {v:Pow(x). v eqpoll n#-k} *) -(* We have obtained this result in two steps : *) -(* 1. y is less than or equipollent to {v:s_u. a <= v} *) -(* where a is certain k-2 element subset of y *) -(* 2. {v:s_u. a <= v} is less than or equipollent *) -(* to {v:Pow(x). v eqpoll n-k} *) +(* There is a v : s_u such that k lepoll x Int y (in our case succ(k)) *) +(* The idea of the proof is the following : *) +(* Suppose not, i.e. every element of s_u has exactly k-1 elements of y *) +(* Thence y is less than or equipollent to {v:Pow(x). v eqpoll n#-k} *) +(* We have obtained this result in two steps : *) +(* 1. y is less than or equipollent to {v:s_u. a <= v} *) +(* where a is certain k-2 element subset of y *) +(* 2. {v:s_u. a <= v} is less than or equipollent *) +(* to {v:Pow(x). v eqpoll n-k} *) (* ********************************************************************** *) goal thy "!!A. [| ~(EX x:A. f`x=y); f : inj(A, B); y:B |] \ -\ ==> (lam a:succ(A). if(a=A, y, f`a)) : inj(succ(A), B)"; +\ ==> (lam a:succ(A). if(a=A, y, f`a)) : inj(succ(A), B)"; by (res_inst_tac [("d","%z. if(z=y, A, converse(f)`z)")] lam_injective 1); by (fast_tac (AC_cs addSEs [inj_is_fun RS apply_type] - addIs [expand_if RS iffD2]) 1); + addIs [expand_if RS iffD2]) 1); by (REPEAT (split_tac [expand_if] 1)); by (fast_tac (AC_cs addSEs [left_inverse]) 1); val succ_not_lepoll_lemma = result(); goalw thy [lepoll_def, eqpoll_def, bij_def, surj_def] - "!!A. [| ~A eqpoll B; A lepoll B |] ==> succ(A) lepoll B"; + "!!A. [| ~A eqpoll B; A lepoll B |] ==> succ(A) lepoll B"; by (fast_tac (AC_cs addSEs [succ_not_lepoll_lemma, inj_is_fun]) 1); val succ_not_lepoll_imp_eqpoll = result(); val [prem] = goalw thy [s_u_def] - "(ALL v:s_u(u, t_n, k, y). k eqpoll v Int y ==> False) \ -\ ==> EX v : s_u(u, t_n, k, y). succ(k) lepoll v Int y"; + "(ALL v:s_u(u, t_n, k, y). k eqpoll v Int y ==> False) \ +\ ==> EX v : s_u(u, t_n, k, y). succ(k) lepoll v Int y"; by (excluded_middle_tac "?P" 1 THEN (assume_tac 2)); by (resolve_tac [prem RS FalseE] 1); by (rtac ballI 1); @@ -112,48 +112,48 @@ val suppose_not = result(); (* ********************************************************************** *) -(* There is a k-2 element subset of y *) +(* There is a k-2 element subset of y *) (* ********************************************************************** *) goalw thy [lepoll_def, eqpoll_def] - "!!X. [| n:nat; nat lepoll X |] ==> EX Y. Y<=X & n eqpoll Y"; + "!!X. [| n:nat; nat lepoll X |] ==> EX Y. Y<=X & n eqpoll Y"; by (fast_tac (FOL_cs addSDs [Ord_nat RSN (2, OrdmemD) RSN (2, restrict_inj)] - addSIs [subset_refl] - addSEs [restrict_bij, inj_is_fun RS fun_is_rel RS image_subset]) 1); + addSIs [subset_refl] + addSEs [restrict_bij, inj_is_fun RS fun_is_rel RS image_subset]) 1); val nat_lepoll_imp_ex_eqpoll_n = result(); val ordertype_eqpoll = - ordermap_bij RS (exI RS (eqpoll_def RS def_imp_iff RS iffD2)); + ordermap_bij RS (exI RS (eqpoll_def RS def_imp_iff RS iffD2)); goal thy "!!y. [| well_ord(y,R); ~Finite(y); n:nat \ -\ |] ==> EX z. z<=y & n eqpoll z"; +\ |] ==> EX z. z<=y & n eqpoll z"; by (etac nat_lepoll_imp_ex_eqpoll_n 1); by (resolve_tac [ordertype_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll - RSN (2, lepoll_trans)] 1 THEN (assume_tac 2)); + RSN (2, lepoll_trans)] 1 THEN (assume_tac 2)); by (fast_tac (AC_cs addSIs [nat_le_infinite_Ord RS le_imp_lepoll] - addSEs [Ord_ordertype, ordertype_eqpoll RS eqpoll_imp_lepoll - RS lepoll_infinite]) 1); + addSEs [Ord_ordertype, ordertype_eqpoll RS eqpoll_imp_lepoll + RS lepoll_infinite]) 1); val ex_subset_eqpoll_n = result(); goalw thy [lesspoll_def] "!!n. n: nat ==> n lesspoll nat"; by (fast_tac (AC_cs addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_lepoll, - eqpoll_sym RS eqpoll_imp_lepoll] - addIs [Ord_nat RSN (2, nat_succI RS ltI) RS leI - RS le_imp_lepoll RS lepoll_trans RS succ_lepoll_natE]) 1); + eqpoll_sym RS eqpoll_imp_lepoll] + addIs [Ord_nat RSN (2, nat_succI RS ltI) RS leI + RS le_imp_lepoll RS lepoll_trans RS succ_lepoll_natE]) 1); val n_lesspoll_nat = result(); goal thy "!!y. [| well_ord(y,R); ~Finite(y); k eqpoll a; a <= y; k: nat |] \ -\ ==> y - a eqpoll y"; +\ ==> y - a eqpoll y"; by (fast_tac (empty_cs addIs [lepoll_lesspoll_lesspoll] - addSIs [Card_cardinal, Diff_lesspoll_eqpoll_Card RS eqpoll_trans, - Card_cardinal RS Card_is_Ord RS nat_le_infinite_Ord - RS le_imp_lepoll] - addSEs [well_ord_cardinal_eqpoll, - well_ord_cardinal_eqpoll RS eqpoll_sym, - eqpoll_sym RS eqpoll_imp_lepoll, - n_lesspoll_nat RS lesspoll_lepoll_lesspoll, - well_ord_cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll - RS lepoll_infinite]) 1); + addSIs [Card_cardinal, Diff_lesspoll_eqpoll_Card RS eqpoll_trans, + Card_cardinal RS Card_is_Ord RS nat_le_infinite_Ord + RS le_imp_lepoll] + addSEs [well_ord_cardinal_eqpoll, + well_ord_cardinal_eqpoll RS eqpoll_sym, + eqpoll_sym RS eqpoll_imp_lepoll, + n_lesspoll_nat RS lesspoll_lepoll_lesspoll, + well_ord_cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll + RS lepoll_infinite]) 1); val Diff_Finite_eqpoll = result(); goal thy "!!x. [| a<=y; b:y-a; u:x |] ==> cons(b, cons(u, a)) : Pow(x Un y)"; @@ -161,7 +161,7 @@ val cons_cons_subset = result(); goal thy "!!x. [| a eqpoll k; a<=y; b:y-a; u:x; x Int y = 0 \ -\ |] ==> cons(b, cons(u, a)) eqpoll succ(succ(k))"; +\ |] ==> cons(b, cons(u, a)) eqpoll succ(succ(k))"; by (fast_tac (AC_cs addSIs [cons_eqpoll_succ] addEs [equals0D]) 1); val cons_cons_eqpoll = result(); @@ -170,10 +170,10 @@ val s_u_subset = result(); goalw thy [s_u_def, succ_def] - "!!w. [| w:t_n; cons(b,cons(u,a)) <= w; a <= y; b : y-a; k eqpoll a \ -\ |] ==> w: s_u(u, t_n, succ(k), y)"; + "!!w. [| w:t_n; cons(b,cons(u,a)) <= w; a <= y; b : y-a; k eqpoll a \ +\ |] ==> w: s_u(u, t_n, succ(k), y)"; by (fast_tac (AC_cs addDs [eqpoll_imp_lepoll RS cons_lepoll_cong] - addSEs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1); + addSEs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1); val s_uI = result(); goalw thy [s_u_def] "!!v. v : s_u(u, t_n, k, y) ==> u : v"; @@ -181,13 +181,13 @@ val in_s_u_imp_u_in = result(); goal thy - "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}. \ -\ EX! w. w:t_n & z <= w; \ -\ k eqpoll a; a <= y; b : y - a; u : x; x Int y = 0 |] \ -\ ==> EX! c. c:{v:s_u(u, t_n, succ(k), y). a <= v} & b:c"; + "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}. \ +\ EX! w. w:t_n & z <= w; \ +\ k eqpoll a; a <= y; b : y - a; u : x; x Int y = 0 |] \ +\ ==> EX! c. c:{v:s_u(u, t_n, succ(k), y). a <= v} & b:c"; by (etac ballE 1); by (fast_tac (FOL_cs addSIs [CollectI, cons_cons_subset, - eqpoll_sym RS cons_cons_eqpoll]) 2); + eqpoll_sym RS cons_cons_eqpoll]) 2); by (etac ex1E 1); by (res_inst_tac [("a","w")] ex1I 1); by (rtac conjI 1); @@ -202,8 +202,8 @@ val ex1_superset_a = result(); goal thy - "!!A. [| succ(k) eqpoll A; k eqpoll B; B <= A; a : A-B; k:nat \ -\ |] ==> A = cons(a, B)"; + "!!A. [| succ(k) eqpoll A; k eqpoll B; B <= A; a : A-B; k:nat \ +\ |] ==> A = cons(a, B)"; by (rtac equalityI 1); by (fast_tac AC_cs 2); by (resolve_tac [Diff_eq_0_iff RS iffD1] 1); @@ -214,17 +214,17 @@ by (dtac cons_eqpoll_succ 1); by (fast_tac AC_cs 1); by (fast_tac (AC_cs addSIs [nat_succI] - addSEs [[eqpoll_sym RS eqpoll_imp_lepoll, subset_imp_lepoll] MRS - (lepoll_trans RS lepoll_trans) RS succ_lepoll_natE]) 1); + addSEs [[eqpoll_sym RS eqpoll_imp_lepoll, subset_imp_lepoll] MRS + (lepoll_trans RS lepoll_trans) RS succ_lepoll_natE]) 1); val set_eq_cons = result(); goal thy - "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}. \ -\ EX! w. w:t_n & z <= w; \ -\ ALL v:s_u(u, t_n, succ(k), y). succ(k) eqpoll v Int y; \ -\ k eqpoll a; a <= y; b : y - a; u : x; x Int y = 0; k:nat \ -\ |] ==> (THE c. c:{v:s_u(u, t_n, succ(k), y). a <= v} & b:c) \ -\ Int y = cons(b, a)"; + "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}. \ +\ EX! w. w:t_n & z <= w; \ +\ ALL v:s_u(u, t_n, succ(k), y). succ(k) eqpoll v Int y; \ +\ k eqpoll a; a <= y; b : y - a; u : x; x Int y = 0; k:nat \ +\ |] ==> (THE c. c:{v:s_u(u, t_n, succ(k), y). a <= v} & b:c) \ +\ Int y = cons(b, a)"; by (dresolve_tac [ex1_superset_a RS theI] 1 THEN REPEAT (assume_tac 1)); by (rtac set_eq_cons 1); by (REPEAT (fast_tac AC_cs 1)); @@ -243,28 +243,28 @@ val msubst = result(); (* ********************************************************************** *) -(* 1. y is less than or equipollent to {v:s_u. a <= v} *) -(* where a is certain k-2 element subset of y *) +(* 1. y is less than or equipollent to {v:s_u. a <= v} *) +(* where a is certain k-2 element subset of y *) (* ********************************************************************** *) goal thy - "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}. \ -\ EX! w. w:t_n & z <= w; \ -\ ALL v:s_u(u, t_n, succ(k), y). succ(k) eqpoll v Int y; \ -\ well_ord(y,R); ~Finite(y); k eqpoll a; a <= y; \ -\ k:nat; u:x; x Int y = 0 |] \ -\ ==> y lepoll {v:s_u(u, t_n, succ(k), y). a <= v}"; + "!!y. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(k))}. \ +\ EX! w. w:t_n & z <= w; \ +\ ALL v:s_u(u, t_n, succ(k), y). succ(k) eqpoll v Int y; \ +\ well_ord(y,R); ~Finite(y); k eqpoll a; a <= y; \ +\ k:nat; u:x; x Int y = 0 |] \ +\ ==> y lepoll {v:s_u(u, t_n, succ(k), y). a <= v}"; by (resolve_tac [Diff_Finite_eqpoll RS eqpoll_sym RS - eqpoll_imp_lepoll RS lepoll_trans] 1 - THEN REPEAT (assume_tac 1)); + eqpoll_imp_lepoll RS lepoll_trans] 1 + THEN REPEAT (assume_tac 1)); by (res_inst_tac [("f3","lam b:y-a. \ -\ THE c. c:{v:s_u(u, t_n, succ(k), y). a <= v} & b:c")] - (exI RS (lepoll_def RS def_imp_iff RS iffD2)) 1); +\ THE c. c:{v:s_u(u, t_n, succ(k), y). a <= v} & b:c")] + (exI RS (lepoll_def RS def_imp_iff RS iffD2)) 1); by (resolve_tac [inj_def RS def_imp_eq RS ssubst] 1); by (rtac CollectI 1); by (rtac lam_type 1); by (resolve_tac [ex1_superset_a RS theI RS conjunct1] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (rtac ballI 1); by (rtac ballI 1); by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1)); @@ -274,20 +274,20 @@ by (fast_tac AC_cs 2); by (dres_inst_tac [("A","THE c. ?P(c)"), ("C","y")] eq_imp_Int_eq 1); by (eresolve_tac [[asm_rl, the_eq_cons, the_eq_cons] MRS msubst] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); val y_lepoll_subset_s_u = result(); (* ********************************************************************** *) -(* some arithmetic *) +(* some arithmetic *) (* ********************************************************************** *) goal thy - "!!k. [| k:nat; m:nat |] ==> \ -\ ALL A B. A eqpoll k #+ m & k lepoll B & B<=A --> A-B lepoll m"; + "!!k. [| k:nat; m:nat |] ==> \ +\ ALL A B. A eqpoll k #+ m & k lepoll B & B<=A --> A-B lepoll m"; by (eres_inst_tac [("n","k")] nat_induct 1); by (simp_tac (AC_ss addsimps [add_0]) 1); by (fast_tac (AC_cs addIs [eqpoll_imp_lepoll RS - (Diff_subset RS subset_imp_lepoll RS lepoll_trans)]) 1); + (Diff_subset RS subset_imp_lepoll RS lepoll_trans)]) 1); by (REPEAT (resolve_tac [allI,impI] 1)); by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1); by (fast_tac AC_cs 1); @@ -301,24 +301,24 @@ val eqpoll_sum_imp_Diff_lepoll_lemma = result(); goal thy "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) lepoll B; \ -\ k:nat; m:nat |] \ -\ ==> A-B lepoll m"; +\ k:nat; m:nat |] \ +\ ==> A-B lepoll m"; by (dresolve_tac [add_succ RS ssubst] 1); by (dresolve_tac [nat_succI RS eqpoll_sum_imp_Diff_lepoll_lemma] 1 - THEN (REPEAT (assume_tac 1))); + THEN (REPEAT (assume_tac 1))); by (fast_tac AC_cs 1); val eqpoll_sum_imp_Diff_lepoll = result(); (* ********************************************************************** *) -(* similar properties for eqpoll *) +(* similar properties for eqpoll *) (* ********************************************************************** *) goal thy - "!!k. [| k:nat; m:nat |] ==> \ -\ ALL A B. A eqpoll k #+ m & k eqpoll B & B<=A --> A-B eqpoll m"; + "!!k. [| k:nat; m:nat |] ==> \ +\ ALL A B. A eqpoll k #+ m & k eqpoll B & B<=A --> A-B eqpoll m"; by (eres_inst_tac [("n","k")] nat_induct 1); by (fast_tac (AC_cs addSDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_0_is_0] - addss (AC_ss addsimps [add_0])) 1); + addss (AC_ss addsimps [add_0])) 1); by (REPEAT (resolve_tac [allI,impI] 1)); by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1); by (fast_tac (AC_cs addSEs [eqpoll_imp_lepoll]) 1); @@ -326,72 +326,72 @@ by (eres_inst_tac [("x","B - {xa}")] allE 1); by (etac impE 1); by (fast_tac (AC_cs addSIs [Diff_sing_eqpoll, - eqpoll_sym RSN (2, Diff_sing_eqpoll) RS eqpoll_sym] - addss (AC_ss addsimps [add_succ])) 1); + eqpoll_sym RSN (2, Diff_sing_eqpoll) RS eqpoll_sym] + addss (AC_ss addsimps [add_succ])) 1); by (res_inst_tac [("P","%z. z eqpoll m")] subst 1 THEN (assume_tac 2)); by (fast_tac (AC_cs addSIs [equalityI]) 1); val eqpoll_sum_imp_Diff_eqpoll_lemma = result(); goal thy "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) eqpoll B; \ -\ k:nat; m:nat |] \ -\ ==> A-B eqpoll m"; +\ k:nat; m:nat |] \ +\ ==> A-B eqpoll m"; by (dresolve_tac [add_succ RS ssubst] 1); by (dresolve_tac [nat_succI RS eqpoll_sum_imp_Diff_eqpoll_lemma] 1 - THEN (REPEAT (assume_tac 1))); + THEN (REPEAT (assume_tac 1))); by (fast_tac AC_cs 1); val eqpoll_sum_imp_Diff_eqpoll = result(); (* ********************************************************************** *) -(* back to the second part *) +(* back to the second part *) (* ********************************************************************** *) goal thy "!!w. [| x Int y = 0; w <= x Un y |] \ -\ ==> w Int (x - {u}) = w - cons(u, w Int y)"; +\ ==> w Int (x - {u}) = w - cons(u, w Int y)"; by (fast_tac (AC_cs addSIs [equalityI] addEs [equals0D]) 1); val w_Int_eq_w_Diff = result(); goal thy "!!w. [| w:{v:s_u(u, t_n, succ(l), y). a <= v}; \ -\ l eqpoll a; t_n <= {v:Pow(x Un y). v eqpoll succ(succ(l) #+ m)}; \ -\ m:nat; l:nat; x Int y = 0; u : x; \ -\ ALL v:s_u(u, t_n, succ(l), y). succ(l) eqpoll v Int y \ -\ |] ==> w Int (x - {u}) eqpoll m"; +\ l eqpoll a; t_n <= {v:Pow(x Un y). v eqpoll succ(succ(l) #+ m)}; \ +\ m:nat; l:nat; x Int y = 0; u : x; \ +\ ALL v:s_u(u, t_n, succ(l), y). succ(l) eqpoll v Int y \ +\ |] ==> w Int (x - {u}) eqpoll m"; by (etac CollectE 1); by (resolve_tac [w_Int_eq_w_Diff RS ssubst] 1 THEN (assume_tac 1)); by (fast_tac (AC_cs addSDs [s_u_subset RS subsetD]) 1); by (fast_tac (AC_cs addEs [equals0D] addSDs [bspec] - addDs [s_u_subset RS subsetD] - addSEs [eqpoll_sym RS cons_eqpoll_succ RS eqpoll_sym, in_s_u_imp_u_in] - addSIs [nat_succI, eqpoll_sum_imp_Diff_eqpoll]) 1); + addDs [s_u_subset RS subsetD] + addSEs [eqpoll_sym RS cons_eqpoll_succ RS eqpoll_sym, in_s_u_imp_u_in] + addSIs [nat_succI, eqpoll_sum_imp_Diff_eqpoll]) 1); val w_Int_eqpoll_m = result(); goal thy "!!m. [| 0 x ~= 0"; by (fast_tac (empty_cs - addSEs [mem_irrefl, ltE, eqpoll_succ_imp_not_empty, natE]) 1); + addSEs [mem_irrefl, ltE, eqpoll_succ_imp_not_empty, natE]) 1); val eqpoll_m_not_empty = result(); goal thy - "!!z. [| z : xa Int (x - {u}); l eqpoll a; a <= y; x Int y = 0; u:x \ -\ |] ==> cons(z, cons(u, a)) : {v: Pow(x Un y). v eqpoll succ(succ(l))}"; + "!!z. [| z : xa Int (x - {u}); l eqpoll a; a <= y; x Int y = 0; u:x \ +\ |] ==> cons(z, cons(u, a)) : {v: Pow(x Un y). v eqpoll succ(succ(l))}"; by (fast_tac (AC_cs addSIs [cons_eqpoll_succ] addEs [equals0D, eqpoll_sym]) 1); val cons_cons_in = result(); (* ********************************************************************** *) -(* 2. {v:s_u. a <= v} is less than or equipollent *) -(* to {v:Pow(x). v eqpoll n-k} *) +(* 2. {v:s_u. a <= v} is less than or equipollent *) +(* to {v:Pow(x). v eqpoll n-k} *) (* ********************************************************************** *) goal thy - "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(l))}. \ -\ EX! w. w:t_n & z <= w; \ -\ t_n <= {v:Pow(x Un y). v eqpoll succ(succ(l) #+ m)}; \ -\ 0 {v:s_u(u, t_n, succ(l), y). a <= v} \ -\ lepoll {v:Pow(x). v eqpoll m}"; + "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(succ(l))}. \ +\ EX! w. w:t_n & z <= w; \ +\ t_n <= {v:Pow(x Un y). v eqpoll succ(succ(l) #+ m)}; \ +\ 0 {v:s_u(u, t_n, succ(l), y). a <= v} \ +\ lepoll {v:Pow(x). v eqpoll m}"; by (res_inst_tac [("f3","lam w:{v:s_u(u, t_n, succ(l), y). a <= v}. \ -\ w Int (x - {u})")] - (exI RS (lepoll_def RS def_imp_iff RS iffD2)) 1); +\ w Int (x - {u})")] + (exI RS (lepoll_def RS def_imp_iff RS iffD2)) 1); by (resolve_tac [inj_def RS def_imp_eq RS ssubst] 1); by (rtac CollectI 1); by (rtac lam_type 1); @@ -401,7 +401,7 @@ by (simp_tac AC_ss 1); by (REPEAT (resolve_tac [ballI, impI] 1)); by (eresolve_tac [w_Int_eqpoll_m RSN (2, eqpoll_m_not_empty) RS not_emptyE] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (dresolve_tac [equalityD1 RS subsetD] 1 THEN (assume_tac 1)); by (dresolve_tac [cons_cons_in RSN (2, bspec)] 1 THEN REPEAT (assume_tac 1)); by (etac ex1_two_eq 1); @@ -416,62 +416,62 @@ val ex_eq_succ = result(); goal thy - "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. \ -\ EX! w. w:t_n & z <= w; \ -\ well_ord(y,R); ~Finite(y); u:x; x Int y = 0; \ -\ t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)}; \ -\ ~ y lepoll {v:Pow(x). v eqpoll m}; 0 EX v : s_u(u, t_n, k, y). succ(k) lepoll v Int y"; + "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. \ +\ EX! w. w:t_n & z <= w; \ +\ well_ord(y,R); ~Finite(y); u:x; x Int y = 0; \ +\ t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)}; \ +\ ~ y lepoll {v:Pow(x). v eqpoll m}; 0 EX v : s_u(u, t_n, k, y). succ(k) lepoll v Int y"; by (rtac suppose_not 1); by (eresolve_tac [ex_eq_succ RS bexE] 1 THEN (assume_tac 1)); by (hyp_subst_tac 1); by (res_inst_tac [("n1","xa")] (ex_subset_eqpoll_n RS exE) 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (etac conjE 1); by (forward_tac [[y_lepoll_subset_s_u, subset_s_u_lepoll_w] MRS lepoll_trans] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (contr_tac 1); val exists_proper_in_s_u = result(); (* ********************************************************************** *) -(* LL can be well ordered *) +(* LL can be well ordered *) (* ********************************************************************** *) goal thy "{x:Pow(X). x lepoll 0} = {0}"; by (fast_tac (AC_cs addSDs [lepoll_0_is_0] - addSIs [singletonI, lepoll_refl, equalityI] - addSEs [singletonE]) 1); + addSIs [singletonI, lepoll_refl, equalityI] + addSEs [singletonE]) 1); val subsets_lepoll_0_eq_unit = result(); goal thy "!!X. [| well_ord(X, R); ~Finite(X); n:nat |] \ -\ ==> EX S. well_ord({Y: Pow(X) . Y eqpoll succ(n)}, S)"; +\ ==> EX S. well_ord({Y: Pow(X) . Y eqpoll succ(n)}, S)"; by (resolve_tac [well_ord_infinite_subsets_eqpoll_X - RS (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1 + RS (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1 THEN (REPEAT (assume_tac 1))); by (fast_tac (ZF_cs addSEs [bij_is_inj RS well_ord_rvimage]) 1); val well_ord_subsets_eqpoll_n = result(); goal thy "!!n. n:nat ==> {z:Pow(y). z lepoll succ(n)} = \ -\ {z:Pow(y). z lepoll n} Un {z:Pow(y). z eqpoll succ(n)}"; +\ {z:Pow(y). z lepoll n} Un {z:Pow(y). z eqpoll succ(n)}"; by (fast_tac (ZF_cs addIs [le_refl, leI, - le_imp_lepoll, equalityI] - addSDs [lepoll_succ_disj] - addSEs [nat_into_Ord, lepoll_trans, eqpoll_imp_lepoll]) 1); + le_imp_lepoll, equalityI] + addSDs [lepoll_succ_disj] + addSEs [nat_into_Ord, lepoll_trans, eqpoll_imp_lepoll]) 1); val subsets_lepoll_succ = result(); goal thy "!!n. n:nat ==> \ -\ {z:Pow(y). z lepoll n} Int {z:Pow(y). z eqpoll succ(n)} = 0"; +\ {z:Pow(y). z lepoll n} Int {z:Pow(y). z eqpoll succ(n)} = 0"; by (fast_tac (ZF_cs addSEs [eqpoll_sym RS eqpoll_imp_lepoll - RS lepoll_trans RS succ_lepoll_natE] - addSIs [equals0I]) 1); + RS lepoll_trans RS succ_lepoll_natE] + addSIs [equals0I]) 1); val Int_empty = result(); (* ********************************************************************** *) -(* unit set is well-ordered by the empty relation *) +(* unit set is well-ordered by the empty relation *) (* ********************************************************************** *) goalw thy [irrefl_def, trans_on_def, part_ord_def, linear_def, tot_ord_def] - "tot_ord({a},0)"; + "tot_ord({a},0)"; by (simp_tac ZF_ss 1); val tot_ord_unit = result(); @@ -484,49 +484,49 @@ val well_ord_unit = result(); (* ********************************************************************** *) -(* well_ord_subsets_lepoll_n *) +(* well_ord_subsets_lepoll_n *) (* ********************************************************************** *) goal thy "!!y r. [| well_ord(y,r); ~Finite(y); n:nat |] ==> \ -\ EX R. well_ord({z:Pow(y). z lepoll n}, R)"; +\ EX R. well_ord({z:Pow(y). z lepoll n}, R)"; by (etac nat_induct 1); by (fast_tac (ZF_cs addSIs [well_ord_unit] - addss (ZF_ss addsimps [subsets_lepoll_0_eq_unit])) 1); + addss (ZF_ss addsimps [subsets_lepoll_0_eq_unit])) 1); by (etac exE 1); by (eresolve_tac [well_ord_subsets_eqpoll_n RS exE] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (asm_simp_tac (ZF_ss addsimps [subsets_lepoll_succ]) 1); by (dtac well_ord_radd 1 THEN (assume_tac 1)); by (eresolve_tac [Int_empty RS disj_Un_eqpoll_sum RS - (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1); + (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1); by (fast_tac (ZF_cs addSEs [bij_is_inj RS well_ord_rvimage]) 1); val well_ord_subsets_lepoll_n = result(); goalw thy [LL_def, MM_def] - "!!x. t_n <= {v:Pow(x Un y). v eqpoll n} \ -\ ==> LL(t_n, k, y) <= {z:Pow(y). z lepoll n}"; + "!!x. t_n <= {v:Pow(x Un y). v eqpoll n} \ +\ ==> LL(t_n, k, y) <= {z:Pow(y). z lepoll n}"; by (fast_tac (AC_cs addSEs [RepFunE] - addIs [subset_imp_lepoll RS (eqpoll_imp_lepoll - RSN (2, lepoll_trans))]) 1); + addIs [subset_imp_lepoll RS (eqpoll_imp_lepoll + RSN (2, lepoll_trans))]) 1); val LL_subset = result(); goal thy "!!x. [| t_n <= {v:Pow(x Un y). v eqpoll n}; \ -\ well_ord(y, R); ~Finite(y); n:nat \ -\ |] ==> EX S. well_ord(LL(t_n, k, y), S)"; +\ well_ord(y, R); ~Finite(y); n:nat \ +\ |] ==> EX S. well_ord(LL(t_n, k, y), S)"; by (fast_tac (FOL_cs addIs [exI] - addSEs [LL_subset RSN (2, well_ord_subset)] - addEs [well_ord_subsets_lepoll_n RS exE]) 1); + addSEs [LL_subset RSN (2, well_ord_subset)] + addEs [well_ord_subsets_lepoll_n RS exE]) 1); val well_ord_LL = result(); (* ********************************************************************** *) -(* every element of LL is a contained in exactly one element of MM *) +(* every element of LL is a contained in exactly one element of MM *) (* ********************************************************************** *) goalw thy [MM_def, LL_def] - "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \ -\ t_n <= {v:Pow(x Un y). v eqpoll n}; \ -\ v:LL(t_n, k, y) \ -\ |] ==> EX! w. w:MM(t_n, k, y) & v<=w"; + "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \ +\ t_n <= {v:Pow(x Un y). v eqpoll n}; \ +\ v:LL(t_n, k, y) \ +\ |] ==> EX! w. w:MM(t_n, k, y) & v<=w"; by (step_tac (AC_cs addSEs [RepFunE]) 1); by (resolve_tac [lepoll_imp_eqpoll_subset RS exE] 1 THEN (assume_tac 1)); by (eres_inst_tac [("x","xa")] ballE 1); @@ -540,22 +540,22 @@ val unique_superset_in_MM = result(); (* ********************************************************************** *) -(* The function GG satisfies the conditions of WO4 *) +(* The function GG satisfies the conditions of WO4 *) (* ********************************************************************** *) (* ********************************************************************** *) -(* The union of appropriate values is the whole x *) +(* The union of appropriate values is the whole x *) (* ********************************************************************** *) goal thy - "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. \ -\ EX! w. w:t_n & z <= w; \ -\ well_ord(y,R); ~Finite(y); u:x; x Int y = 0; \ -\ t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)}; \ -\ ~ y lepoll {v:Pow(x). v eqpoll m}; 0 EX w:MM(t_n, succ(k), y). u:w"; + "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. \ +\ EX! w. w:t_n & z <= w; \ +\ well_ord(y,R); ~Finite(y); u:x; x Int y = 0; \ +\ t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)}; \ +\ ~ y lepoll {v:Pow(x). v eqpoll m}; 0 EX w:MM(t_n, succ(k), y). u:w"; by (eresolve_tac [exists_proper_in_s_u RS bexE] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (rewrite_goals_tac [MM_def, s_u_def]); by (fast_tac AC_cs 1); val exists_in_MM = result(); @@ -569,12 +569,12 @@ val MM_subset = result(); goal thy - "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. \ -\ EX! w. w:t_n & z <= w; \ -\ well_ord(y,R); ~Finite(y); u:x; x Int y = 0; \ -\ t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)}; \ -\ ~ y lepoll {v:Pow(x). v eqpoll m}; 0 EX w:LL(t_n, succ(k), y). u:GG(t_n, succ(k), y)`w"; + "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. \ +\ EX! w. w:t_n & z <= w; \ +\ well_ord(y,R); ~Finite(y); u:x; x Int y = 0; \ +\ t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)}; \ +\ ~ y lepoll {v:Pow(x). v eqpoll m}; 0 EX w:LL(t_n, succ(k), y). u:GG(t_n, succ(k), y)`w"; by (forward_tac [exists_in_MM] 1 THEN REPEAT (assume_tac 1)); by (etac bexE 1); by (res_inst_tac [("x","w Int y")] bexI 1); @@ -582,33 +582,33 @@ by (rewtac GG_def); by (asm_full_simp_tac (AC_ss addsimps [Int_in_LL]) 1); by (eresolve_tac [unique_superset_in_MM RS the_equality2 RS ssubst] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (REPEAT (fast_tac (AC_cs addEs [equals0D] addSEs [Int_in_LL]) 1)); val exists_in_LL = result(); goalw thy [LL_def] - "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \ -\ t_n <= {v:Pow(x Un y). v eqpoll n}; \ -\ v : LL(t_n, k, y) |] \ -\ ==> v = (THE x. x : MM(t_n, k, y) & v <= x) Int y"; + "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \ +\ t_n <= {v:Pow(x Un y). v eqpoll n}; \ +\ v : LL(t_n, k, y) |] \ +\ ==> v = (THE x. x : MM(t_n, k, y) & v <= x) Int y"; by (fast_tac (AC_cs addSEs [Int_in_LL, - unique_superset_in_MM RS the_equality2 RS ssubst]) 1); + unique_superset_in_MM RS the_equality2 RS ssubst]) 1); val in_LL_eq_Int = result(); goal thy - "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \ -\ t_n <= {v:Pow(x Un y). v eqpoll n}; \ -\ v : LL(t_n, k, y) |] \ -\ ==> (THE x. x : MM(t_n, k, y) & v <= x) <= x Un y"; + "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \ +\ t_n <= {v:Pow(x Un y). v eqpoll n}; \ +\ v : LL(t_n, k, y) |] \ +\ ==> (THE x. x : MM(t_n, k, y) & v <= x) <= x Un y"; by (fast_tac (AC_cs addSDs [unique_superset_in_MM RS theI RS conjunct1 RS - (MM_subset RS subsetD)]) 1); + (MM_subset RS subsetD)]) 1); val the_in_MM_subset = result(); goalw thy [GG_def] - "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \ -\ t_n <= {v:Pow(x Un y). v eqpoll n}; \ -\ v : LL(t_n, k, y) |] \ -\ ==> GG(t_n, k, y) ` v <= x"; + "!!x. [| ALL z:{z: Pow(x Un y) . z eqpoll k}. EX! w. w:t_n & z <= w; \ +\ t_n <= {v:Pow(x Un y). v eqpoll n}; \ +\ v : LL(t_n, k, y) |] \ +\ ==> GG(t_n, k, y) ` v <= x"; by (asm_full_simp_tac AC_ss 1); by (forward_tac [the_in_MM_subset] 1 THEN REPEAT (assume_tac 1)); by (dtac in_LL_eq_Int 1 THEN REPEAT (assume_tac 1)); @@ -619,20 +619,20 @@ val GG_subset = result(); goal thy - "!!x. [| well_ord(LL(t_n, succ(k), y), S); \ -\ ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. EX! w. w:t_n & z <= w; \ -\ well_ord(y,R); ~Finite(y); x Int y = 0; \ -\ t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)}; \ -\ ~ y lepoll {v:Pow(x). v eqpoll m}; 0 (UN b (UN b w eqpoll n"; + "!!w. [| w : MM(t_n, k, y); t_n <= {v:Pow(x Un y). v eqpoll n} \ +\ |] ==> w eqpoll n"; by (fast_tac AC_cs 1); val in_MM_eqpoll_n = result(); goalw thy [LL_def, MM_def] - "!!w. w : LL(t_n, k, y) ==> k lepoll w"; + "!!w. w : LL(t_n, k, y) ==> k lepoll w"; by (fast_tac AC_cs 1); val in_LL_eqpoll_n = result(); goalw thy [GG_def] - "!!x. [| \ -\ ALL z:{z: Pow(x Un y) . z eqpoll succ(k)}. EX! w. w:t_n & z <= w; \ -\ t_n <= {v:Pow(x Un y). v eqpoll succ(k #+ m)}; \ -\ well_ord(LL(t_n, succ(k), y), S); k:nat; m:nat |] \ -\ ==> ALL b ALL b WO4(n-k) *) +(* The main theorem AC16(n, k) ==> WO4(n-k) *) (* ********************************************************************** *) goalw thy [AC16_def,WO4_def] - "!!n k. [| AC16(k #+ m, k); 0 < k; 0 < m; k:nat; m:nat |] ==> WO4(m)"; + "!!n k. [| AC16(k #+ m, k); 0 < k; 0 < m; k:nat; m:nat |] ==> WO4(m)"; by (rtac allI 1); by (excluded_middle_tac "Finite(A)" 1); by (rtac lemma1 2 THEN REPEAT (assume_tac 2)); @@ -695,9 +695,9 @@ by (fast_tac (AC_cs addSIs [nat_succI, add_type]) 1); by (res_inst_tac [("x","ordertype(LL(T, succ(k), y), x)")] exI 1); by (res_inst_tac [("x","lam b:ordertype(LL(T, succ(k), y), x). \ -\ (GG(T, succ(k), y)) ` \ -\ (converse(ordermap(LL(T, succ(k), y), x)) ` b)")] exI 1); +\ (GG(T, succ(k), y)) ` \ +\ (converse(ordermap(LL(T, succ(k), y), x)) ` b)")] exI 1); by (simp_tac AC_ss 1); by (fast_tac (empty_cs addSIs [conjI, lam_funtype RS domain_of_fun] - addSEs [Ord_ordertype, all_in_lepoll_m, OUN_eq_x]) 1); + addSEs [Ord_ordertype, all_in_lepoll_m, OUN_eq_x]) 1); qed "AC16_WO4"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC16_lemmas.ML --- a/src/ZF/AC/AC16_lemmas.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC16_lemmas.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC16_lemmas.ML +(* Title: ZF/AC/AC16_lemmas.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski Lemmas used in the proofs concerning AC16 *) @@ -51,7 +51,7 @@ by (rewrite_goals_tac [inj_def, surj_def]); by (asm_full_simp_tac AC_ss 1); by (fast_tac (AC_cs addSIs [lam_type, RepFunI] - addIs [singleton_eq_iff RS iffD1]) 1); + addIs [singleton_eq_iff RS iffD1]) 1); by (asm_full_simp_tac AC_ss 1); by (fast_tac (AC_cs addSIs [lam_type]) 1); val eqpoll_RepFun_sing = result(); @@ -62,19 +62,19 @@ val subsets_eqpoll_1_eqpoll = result(); goal thy "!!x. [| InfCard(x); y:Pow(x); y eqpoll succ(z) |] \ -\ ==> (LEAST i. i:y) : y"; +\ ==> (LEAST i. i:y) : y"; by (eresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS - succ_lepoll_imp_not_empty RS not_emptyE] 1); + succ_lepoll_imp_not_empty RS not_emptyE] 1); by (fast_tac (AC_cs addIs [LeastI] - addSDs [InfCard_is_Card RS Card_is_Ord, PowD RS subsetD] - addEs [Ord_in_Ord]) 1); + addSDs [InfCard_is_Card RS Card_is_Ord, PowD RS subsetD] + addEs [Ord_in_Ord]) 1); val InfCard_Least_in = result(); goalw thy [lepoll_def] "!!i. [| InfCard(x); n:nat |] ==> \ -\ {y:Pow(x). y eqpoll succ(succ(n))} lepoll \ -\ x*{y:Pow(x). y eqpoll succ(n)}"; +\ {y:Pow(x). y eqpoll succ(succ(n))} lepoll \ +\ x*{y:Pow(x). y eqpoll succ(n)}"; by (res_inst_tac [("x","lam y:{y:Pow(x). y eqpoll succ(succ(n))}. \ -\ ")] exI 1); +\ ")] exI 1); by (res_inst_tac [("d","%z. cons(fst(z),snd(z))")] lam_injective 1); by (rtac SigmaI 1); by (etac CollectE 1); @@ -88,7 +88,7 @@ by (fast_tac AC_cs 2); by (resolve_tac [PowD RS subsetD] 1 THEN (assume_tac 1)); by (REPEAT (fast_tac (AC_cs addSIs [Diff_sing_eqpoll] - addIs [InfCard_Least_in]) 1)); + addIs [InfCard_Least_in]) 1)); val subsets_lepoll_lemma1 = result(); val prems = goal thy "(!!y. y:z ==> Ord(y)) ==> z <= succ(Union(z))"; @@ -119,7 +119,7 @@ goal thy "!!i. [| Ord(i); Ord(j) |] ==> i Un j = i | i Un j = j"; by (fast_tac (AC_cs addSDs [le_imp_subset] addIs [equalityI] - addEs [Ord_linear_le]) 1); + addEs [Ord_linear_le]) 1); val Un_Ord_disj = result(); goal thy "!!X. x:X ==> Union(X) = x Un Union(X-{x})"; @@ -127,19 +127,19 @@ val Union_eq_Un = result(); goal thy "!!n. n:nat ==> \ -\ ALL z. (ALL y:z. Ord(y)) & z eqpoll n & z~=0 --> Union(z) : z"; +\ ALL z. (ALL y:z. Ord(y)) & z eqpoll n & z~=0 --> Union(z) : z"; by (etac nat_induct 1); by (fast_tac (AC_cs addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1); by (REPEAT (resolve_tac [allI, impI] 1)); by (etac natE 1); by (fast_tac (AC_cs addSDs [eqpoll_1_iff_singleton RS iffD1] - addSIs [Union_singleton]) 1); + addSIs [Union_singleton]) 1); by (hyp_subst_tac 1); by (REPEAT (eresolve_tac [conjE, not_emptyE] 1)); by (eres_inst_tac [("x","z-{xb}")] allE 1); by (etac impE 1); by (fast_tac (AC_cs addSEs [Diff_sing_eqpoll, - Diff_sing_eqpoll RS eqpoll_succ_imp_not_empty]) 1); + Diff_sing_eqpoll RS eqpoll_succ_imp_not_empty]) 1); by (resolve_tac [Union_eq_Un RSN (2, subst_elem)] 1 THEN (assume_tac 2)); by (forward_tac [bspec] 1 THEN (assume_tac 1)); by (dresolve_tac [Diff_subset RS subsetD RSN (2, bspec)] 1 THEN (assume_tac 1)); @@ -151,36 +151,36 @@ val Union_in_lemma = result(); goal thy "!!z. [| ALL x:z. Ord(x); z eqpoll n; z~=0; n:nat |] \ -\ ==> Union(z) : z"; +\ ==> Union(z) : z"; by (dtac Union_in_lemma 1); by (fast_tac FOL_cs 1); val Union_in = result(); goal thy "!!x. [| InfCard(x); z:Pow(x); z eqpoll n; n:nat |] \ -\ ==> succ(Union(z)) : x"; +\ ==> succ(Union(z)) : x"; by (resolve_tac [Limit_has_succ RS ltE] 1 THEN (assume_tac 3)); by (etac InfCard_is_Limit 1); by (excluded_middle_tac "z=0" 1); by (fast_tac (AC_cs addSIs [InfCard_is_Limit RS Limit_has_0] - addIs [Union_0 RS ssubst]) 2); + addIs [Union_0 RS ssubst]) 2); by (resolve_tac - [PowD RS subsetD RS (InfCard_is_Card RS Card_is_Ord RSN (2, ltI))] 1 - THEN (TRYALL assume_tac)); + [PowD RS subsetD RS (InfCard_is_Card RS Card_is_Ord RSN (2, ltI))] 1 + THEN (TRYALL assume_tac)); by (fast_tac (AC_cs addSIs [Union_in] - addSEs [PowD RS subsetD RSN (2, - InfCard_is_Card RS Card_is_Ord RS Ord_in_Ord)]) 1); + addSEs [PowD RS subsetD RSN (2, + InfCard_is_Card RS Card_is_Ord RS Ord_in_Ord)]) 1); val succ_Union_in_x = result(); goalw thy [lepoll_def] "!!X. [| InfCard(x); n:nat |] ==> \ -\ {y:Pow(x). y eqpoll succ(n)} lepoll \ -\ {y:Pow(x). y eqpoll succ(succ(n))}"; +\ {y:Pow(x). y eqpoll succ(n)} lepoll \ +\ {y:Pow(x). y eqpoll succ(succ(n))}"; by (res_inst_tac [("x","lam z:{y:Pow(x). y eqpoll succ(n)}. \ -\ cons(succ(Union(z)), z)")] exI 1); +\ cons(succ(Union(z)), z)")] exI 1); by (res_inst_tac [("d","%z. z-{Union(z)}")] lam_injective 1); by (resolve_tac [Union_cons_eq_succ_Union RS ssubst] 2); by (rtac cons_Diff_eq 2); by (fast_tac (AC_cs addSDs [InfCard_is_Card RS Card_is_Ord] - addEs [Ord_in_Ord] addSIs [succ_Union_not_mem]) 2); + addEs [Ord_in_Ord] addSIs [succ_Union_not_mem]) 2); by (rtac CollectI 1); by (fast_tac (AC_cs addSEs [cons_eqpoll_succ] addSIs [succ_Union_not_mem] @@ -190,45 +190,45 @@ val succ_lepoll_succ_succ = result(); goal thy "!!X. [| InfCard(X); n:nat |] \ -\ ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X"; +\ ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X"; by (etac nat_induct 1); by (rtac subsets_eqpoll_1_eqpoll 1); by (rtac eqpollI 1); by (resolve_tac [subsets_lepoll_lemma1 RS lepoll_trans] 1 - THEN (REPEAT (assume_tac 1))); + THEN (REPEAT (assume_tac 1))); by (resolve_tac [InfCard_is_Card RS Card_is_Ord RS well_ord_Memrel RS - well_ord_InfCard_square_eq RS eqpoll_imp_lepoll - RSN (2, lepoll_trans)] 1 THEN (assume_tac 2)); + well_ord_InfCard_square_eq RS eqpoll_imp_lepoll + RSN (2, lepoll_trans)] 1 THEN (assume_tac 2)); by (resolve_tac [InfCard_is_Card RS Card_cardinal_eq RS ssubst] 2 - THEN (REPEAT (assume_tac 2))); + THEN (REPEAT (assume_tac 2))); by (eresolve_tac [eqpoll_refl RS prod_eqpoll_cong RS eqpoll_imp_lepoll] 1); by (fast_tac (AC_cs addEs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_trans] - addSIs [succ_lepoll_succ_succ]) 1); + addSIs [succ_lepoll_succ_succ]) 1); val subsets_eqpoll_X = result(); goalw thy [surj_def] "!!f. [| f:surj(A,B); y<=B |] \ -\ ==> f``(converse(f)``y) = y"; +\ ==> f``(converse(f)``y) = y"; by (fast_tac (AC_cs addSIs [equalityI] addDs [apply_equality2] - addEs [apply_iff RS iffD2]) 1); + addEs [apply_iff RS iffD2]) 1); val image_vimage_eq = result(); goal thy "!!f. [| f:inj(A,B); y<=A |] ==> converse(f)``(f``y) = y"; by (fast_tac (AC_cs addSIs [equalityI] addSEs [inj_is_fun RS apply_Pair] - addDs [inj_equality]) 1); + addDs [inj_equality]) 1); val vimage_image_eq = result(); goalw thy [eqpoll_def] "!!A B. A eqpoll B \ -\ ==> {Y:Pow(A). Y eqpoll n} eqpoll {Y:Pow(B). Y eqpoll n}"; +\ ==> {Y:Pow(A). Y eqpoll n} eqpoll {Y:Pow(B). Y eqpoll n}"; by (etac exE 1); by (res_inst_tac [("x","lam X:{Y:Pow(A). EX f. f : bij(Y, n)}. f``X")] exI 1); by (res_inst_tac [("d","%Z. converse(f)``Z")] lam_bijective 1); by (fast_tac (AC_cs - addSIs [bij_is_inj RS restrict_bij RS bij_converse_bij RS comp_bij] + addSIs [bij_is_inj RS restrict_bij RS bij_converse_bij RS comp_bij] addSEs [bij_is_fun RS fun_is_rel RS image_subset RS PowI]) 1); by (fast_tac (AC_cs addSIs [bij_converse_bij RS bij_is_inj RS restrict_bij - RS bij_converse_bij RS comp_bij] + RS bij_converse_bij RS comp_bij] addSEs [bij_converse_bij RS bij_is_fun RS fun_is_rel - RS image_subset RS PowI]) 1); + RS image_subset RS PowI]) 1); by (fast_tac (AC_cs addSEs [bij_is_inj RS vimage_image_eq]) 1); by (fast_tac (AC_cs addSEs [bij_is_surj RS image_vimage_eq]) 1); val subsets_eqpoll = result(); @@ -236,8 +236,8 @@ goalw thy [WO2_def] "!!X. WO2 ==> EX a. Card(a) & X eqpoll a"; by (REPEAT (eresolve_tac [allE,exE,conjE] 1)); by (fast_tac (AC_cs addSEs [well_ord_Memrel RS well_ord_cardinal_eqpoll RS - (eqpoll_sym RSN (2, eqpoll_trans)) RS eqpoll_sym] - addSIs [Card_cardinal]) 1); + (eqpoll_sym RSN (2, eqpoll_trans)) RS eqpoll_sym] + addSIs [Card_cardinal]) 1); val WO2_imp_ex_Card = result(); goal thy "!!X. [| X lepoll Y; ~Finite(X) |] ==> ~Finite(Y)"; @@ -249,7 +249,7 @@ val infinite_Card_is_InfCard = result(); goal thy "!!X n. [| WO2; n:nat; ~Finite(X) |] \ -\ ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X"; +\ ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X"; by (dtac WO2_imp_ex_Card 1); by (REPEAT (eresolve_tac [allE,exE,conjE] 1)); by (forward_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1 THEN (assume_tac 1)); @@ -262,11 +262,11 @@ goal thy "!!X. well_ord(X,R) ==> EX a. Card(a) & X eqpoll a"; by (fast_tac (AC_cs addSEs [well_ord_cardinal_eqpoll RS eqpoll_sym] - addSIs [Card_cardinal]) 1); + addSIs [Card_cardinal]) 1); val well_ord_imp_ex_Card = result(); goal thy "!!X n. [| well_ord(X,R); n:nat; ~Finite(X) |] \ -\ ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X"; +\ ==> {Y:Pow(X). Y eqpoll succ(n)} eqpoll X"; by (dtac well_ord_imp_ex_Card 1); by (REPEAT (eresolve_tac [allE,exE,conjE] 1)); by (forward_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1 THEN (assume_tac 1)); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC17_AC1.ML --- a/src/ZF/AC/AC17_AC1.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC17_AC1.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC17_AC1.ML +(* Title: ZF/AC/AC17_AC1.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proof of AC1 ==> AC17 *) @@ -8,25 +8,25 @@ open AC17_AC1; (* *********************************************************************** *) -(* more properties of HH *) +(* more properties of HH *) (* *********************************************************************** *) goal thy "!!f. [| x - (UN j:LEAST i. HH(lam X:Pow(x)-{0}. {f`X}, x, i) = {x}. \ -\ HH(lam X:Pow(x)-{0}. {f`X}, x, j)) = 0; \ -\ f : Pow(x)-{0} -> x |] \ -\ ==> EX r. well_ord(x,r)"; +\ HH(lam X:Pow(x)-{0}. {f`X}, x, j)) = 0; \ +\ f : Pow(x)-{0} -> x |] \ +\ ==> EX r. well_ord(x,r)"; by (rtac exI 1); by (eresolve_tac [[bij_Least_HH_x RS bij_converse_bij RS bij_is_inj, - Ord_Least RS well_ord_Memrel] MRS well_ord_rvimage] 1); + Ord_Least RS well_ord_Memrel] MRS well_ord_rvimage] 1); by (assume_tac 1); val UN_eq_imp_well_ord = result(); (* *********************************************************************** *) -(* theorems closer to the proof *) +(* theorems closer to the proof *) (* *********************************************************************** *) goalw thy AC_defs "!!Z. ~AC1 ==> \ -\ EX A. ALL f:Pow(A)-{0} -> A. EX u:Pow(A)-{0}. f`u ~: u"; +\ EX A. ALL f:Pow(A)-{0} -> A. EX u:Pow(A)-{0}. f`u ~: u"; by (etac swap 1); by (rtac allI 1); by (etac swap 1); @@ -38,10 +38,10 @@ val not_AC1_imp_ex = result(); goal thy "!!x. [| ALL f:Pow(x) - {0} -> x. EX u: Pow(x) - {0}. f`u~:u; \ -\ EX f: Pow(x)-{0}->x. \ -\ x - (UN a:(LEAST i. HH(lam X:Pow(x)-{0}. {f`X},x,i)={x}). \ -\ HH(lam X:Pow(x)-{0}. {f`X},x,a)) = 0 |] \ -\ ==> P"; +\ EX f: Pow(x)-{0}->x. \ +\ x - (UN a:(LEAST i. HH(lam X:Pow(x)-{0}. {f`X},x,i)={x}). \ +\ HH(lam X:Pow(x)-{0}. {f`X},x,a)) = 0 |] \ +\ ==> P"; by (etac bexE 1); by (eresolve_tac [UN_eq_imp_well_ord RS exE] 1 THEN (assume_tac 1)); by (eresolve_tac [ex_choice_fun_Pow RS exE] 1); @@ -54,20 +54,20 @@ val lemma1 = result(); goal thy "!!x. ~ (EX f: Pow(x)-{0}->x. x - F(f) = 0) \ -\ ==> (lam f: Pow(x)-{0}->x. x - F(f)) \ -\ : (Pow(x) -{0} -> x) -> Pow(x) - {0}"; +\ ==> (lam f: Pow(x)-{0}->x. x - F(f)) \ +\ : (Pow(x) -{0} -> x) -> Pow(x) - {0}"; by (fast_tac (AC_cs addSIs [lam_type] addIs [equalityI] - addSDs [Diff_eq_0_iff RS iffD1]) 1); + addSDs [Diff_eq_0_iff RS iffD1]) 1); val lemma2 = result(); goal thy "!!f. [| f`Z : Z; Z:Pow(x)-{0} |] ==> \ -\ (lam X:Pow(x)-{0}. {f`X})`Z : Pow(Z)-{0}"; +\ (lam X:Pow(x)-{0}. {f`X})`Z : Pow(Z)-{0}"; by (asm_full_simp_tac AC_ss 1); by (fast_tac (AC_cs addSDs [equals0D]) 1); val lemma3 = result(); goal thy "!!z. EX f:F. f`((lam f:F. Q(f))`f) : (lam f:F. Q(f))`f \ -\ ==> EX f:F. f`Q(f) : Q(f)"; +\ ==> EX f:F. f`Q(f) : Q(f)"; by (asm_full_simp_tac AC_ss 1); val lemma4 = result(); @@ -75,9 +75,9 @@ by (rtac classical 1); by (eresolve_tac [not_AC1_imp_ex RS exE] 1); by (excluded_middle_tac - "EX f: Pow(x)-{0}->x. \ -\ x - (UN a:(LEAST i. HH(lam X:Pow(x)-{0}. {f`X},x,i)={x}). \ -\ HH(lam X:Pow(x)-{0}. {f`X},x,a)) = 0" 1); + "EX f: Pow(x)-{0}->x. \ +\ x - (UN a:(LEAST i. HH(lam X:Pow(x)-{0}. {f`X},x,i)={x}). \ +\ HH(lam X:Pow(x)-{0}. {f`X},x,a)) = 0" 1); by (etac lemma1 2 THEN (assume_tac 2)); by (dtac lemma2 1); by (etac allE 1); @@ -89,5 +89,5 @@ by (assume_tac 1); by (dtac lemma3 1 THEN (assume_tac 1)); by (fast_tac (AC_cs addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem), - f_subset_imp_HH_subset] addSEs [mem_irrefl]) 1); + f_subset_imp_HH_subset] addSEs [mem_irrefl]) 1); qed "AC17_AC1"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC18_AC19.ML --- a/src/ZF/AC/AC18_AC19.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC18_AC19.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC18_AC19.ML +(* Title: ZF/AC/AC18_AC19.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proof of AC1 ==> AC18 ==> AC19 ==> AC1 *) @@ -8,11 +8,11 @@ open AC18_AC19; (* ********************************************************************** *) -(* AC1 ==> AC18 *) +(* AC1 ==> AC18 *) (* ********************************************************************** *) goal thy "!!f. [| f: (PROD b:{P(a). a:A}. b); ALL a:A. P(a)<=Q(a) |] \ -\ ==> (lam a:A. f`P(a)):(PROD a:A. Q(a))"; +\ ==> (lam a:A. f`P(a)):(PROD a:A. Q(a))"; by (rtac lam_type 1); by (dtac apply_type 1); by (rtac RepFunI 1 THEN (assume_tac 1)); @@ -26,23 +26,23 @@ by (eres_inst_tac [("x","{{b:B(a). x:X(a,b)}. a:A}")] allE 1); by (etac impE 1); by (fast_tac (AC_cs addSEs [RepFunE] addSDs [INT_E] - addEs [UN_E, sym RS equals0D]) 1); + addEs [UN_E, sym RS equals0D]) 1); by (etac exE 1); by (rtac UN_I 1); by (fast_tac (AC_cs addSEs [PROD_subsets]) 1); by (simp_tac AC_ss 1); by (fast_tac (FOL_cs addSEs [not_emptyE] addDs [RepFunI RSN (2, apply_type)] - addEs [CollectD2] addSIs [INT_I]) 1); + addEs [CollectD2] addSIs [INT_I]) 1); val lemma_AC18 = result(); val [prem] = goalw thy (AC18_def::AC_defs) "AC1 ==> AC18"; by (resolve_tac [prem RS revcut_rl] 1); by (fast_tac (AC_cs addSEs [lemma_AC18, UN_E, not_emptyE, apply_type] - addSIs [equalityI, INT_I, UN_I]) 1); + addSIs [equalityI, INT_I, UN_I]) 1); qed "AC1_AC18"; (* ********************************************************************** *) -(* AC18 ==> AC19 *) +(* AC18 ==> AC19 *) (* ********************************************************************** *) val [prem] = goalw thy [AC18_def, AC19_def] "AC18 ==> AC19"; @@ -52,14 +52,14 @@ qed "AC18_AC19"; (* ********************************************************************** *) -(* AC19 ==> AC1 *) +(* AC19 ==> AC1 *) (* ********************************************************************** *) goalw thy [u_def] - "!!A. [| A ~= 0; 0 ~: A |] ==> {u_(a). a:A} ~= 0 & 0 ~: {u_(a). a:A}"; + "!!A. [| A ~= 0; 0 ~: A |] ==> {u_(a). a:A} ~= 0 & 0 ~: {u_(a). a:A}"; by (fast_tac (AC_cs addSIs [not_emptyI, RepFunI] - addSEs [not_emptyE, RepFunE] - addSDs [sym RS (RepFun_eq_0_iff RS iffD1)]) 1); + addSEs [not_emptyE, RepFunE] + addSDs [sym RS (RepFun_eq_0_iff RS iffD1)]) 1); val RepRep_conj = result(); goal thy "!!c. [|c : a; x = c Un {0}; x ~: a |] ==> x - {0} : a"; @@ -74,16 +74,16 @@ val lemma1_1 = result(); goalw thy [u_def] - "!!a. [| f`(u_(a)) ~: a; f: (PROD B:{u_(a). a:A}. B); a:A |] \ -\ ==> f`(u_(a))-{0} : a"; + "!!a. [| f`(u_(a)) ~: a; f: (PROD B:{u_(a). a:A}. B); a:A |] \ +\ ==> f`(u_(a))-{0} : a"; by (fast_tac (AC_cs addSEs [RepFunI, RepFunE, lemma1_1] - addSDs [apply_type]) 1); + addSDs [apply_type]) 1); val lemma1_2 = result(); goal thy "!!A. EX f. f: (PROD B:{u_(a). a:A}. B) ==> EX f. f: (PROD B:A. B)"; by (etac exE 1); by (res_inst_tac - [("x","lam a:A. if(f`(u_(a)) : a, f`(u_(a)), f`(u_(a))-{0})")] exI 1); + [("x","lam a:A. if(f`(u_(a)) : a, f`(u_(a)), f`(u_(a))-{0})")] exI 1); by (rtac lam_type 1); by (split_tac [expand_if] 1); by (rtac conjI 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC1_AC17.ML --- a/src/ZF/AC/AC1_AC17.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC1_AC17.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC1_AC17.ML +(* Title: ZF/AC/AC1_AC17.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proof of AC1 ==> AC17 *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC1_WO2.ML --- a/src/ZF/AC/AC1_WO2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC1_WO2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC1_WO2.ML +(* Title: ZF/AC/AC1_WO2.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proof of AC1 ==> WO2 *) @@ -8,7 +8,7 @@ open AC1_WO2; val [prem] = goal thy "f : (PROD X:Pow(x) - {0}. X) ==> \ -\ ?g(f) : bij(x, LEAST i. HH(lam X:Pow(x)-{0}. {f`X}, x, i) = {x})"; +\ ?g(f) : bij(x, LEAST i. HH(lam X:Pow(x)-{0}. {f`X}, x, i) = {x})"; by (resolve_tac [bij_Least_HH_x RS bij_converse_bij] 1); by (rtac f_subsets_imp_UN_HH_eq_x 1); by (resolve_tac [lam_type RS apply_type] 1 THEN (assume_tac 2)); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC2_AC6.ML --- a/src/ZF/AC/AC2_AC6.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC2_AC6.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC2_AC6.ML +(* Title: ZF/AC/AC2_AC6.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proofs needed to show that each of AC2, AC3, ..., AC6 is equivalent to AC0 and AC1: @@ -11,16 +11,16 @@ *) (* ********************************************************************** *) -(* AC1 ==> AC2 *) +(* AC1 ==> AC2 *) (* ********************************************************************** *) goal thy "!!B. [| B:A; f:(PROD X:A. X); 0~:A |] \ -\ ==> {f`B} <= B Int {f`C. C:A}"; +\ ==> {f`B} <= B Int {f`C. C:A}"; by (fast_tac (AC_cs addSEs [singletonE, apply_type, RepFunI]) 1); val lemma1 = result(); goalw thy [pairwise_disjoint_def] - "!!A. [| pairwise_disjoint(A); B:A; C:A; D:B; D:C |] ==> f`B = f`C"; + "!!A. [| pairwise_disjoint(A); B:A; C:A; D:B; D:C |] ==> f`B = f`C"; by (fast_tac (ZF_cs addSEs [equals0D]) 1); val lemma2 = result(); @@ -35,28 +35,28 @@ (* ********************************************************************** *) -(* AC2 ==> AC1 *) +(* AC2 ==> AC1 *) (* ********************************************************************** *) goal thy "!!A. 0~:A ==> 0 ~: {B*{B}. B:A}"; by (fast_tac (AC_cs addSDs [sym RS (Sigma_empty_iff RS iffD1)] - addSEs [RepFunE, equals0D]) 1); + addSEs [RepFunE, equals0D]) 1); val lemma1 = result(); goal thy "!!A. [| X*{X} Int C = {y}; X:A |] \ -\ ==> (THE y. X*{X} Int C = {y}): X*A"; +\ ==> (THE y. X*{X} Int C = {y}): X*A"; by (rtac subst_elem 1); by (fast_tac (ZF_cs addSIs [the_equality] - addSEs [sym RS trans RS (singleton_eq_iff RS iffD1)]) 2); + addSEs [sym RS trans RS (singleton_eq_iff RS iffD1)]) 2); by (fast_tac (AC_cs addSEs [equalityE, make_elim singleton_subsetD]) 1); val lemma2 = result(); goal thy "!!A. ALL D:{E*{E}. E:A}. EX y. D Int C = {y} \ -\ ==> (lam x:A. fst(THE z. (x*{x} Int C = {z}))) : \ -\ (PROD X:A. X) "; +\ ==> (lam x:A. fst(THE z. (x*{x} Int C = {z}))) : \ +\ (PROD X:A. X) "; by (fast_tac (FOL_cs addSEs [lemma2] - addSIs [lam_type, RepFunI, fst_type] - addSDs [bspec]) 1); + addSIs [lam_type, RepFunI, fst_type] + addSDs [bspec]) 1); val lemma3 = result(); goalw thy (AC_defs@AC_aux_defs) "!!Z. AC2 ==> AC1"; @@ -68,7 +68,7 @@ (* ********************************************************************** *) -(* AC1 ==> AC4 *) +(* AC1 ==> AC4 *) (* ********************************************************************** *) goal thy "!!R. 0 ~: {R``{x}. x:domain(R)}"; @@ -83,7 +83,7 @@ (* ********************************************************************** *) -(* AC4 ==> AC3 *) +(* AC4 ==> AC3 *) (* ********************************************************************** *) goal thy "!!f. f:A->B ==> (UN z:A. {z}*f`z) <= A*Union(B)"; @@ -92,9 +92,9 @@ goal thy "!!f. domain(UN z:A. {z}*f(z)) = {a:A. f(a)~=0}"; by (fast_tac (ZF_cs addIs [equalityI] - addSEs [not_emptyE] - addSIs [singletonI, not_emptyI] - addDs [range_type]) 1); + addSEs [not_emptyE] + addSIs [singletonI, not_emptyI] + addDs [range_type]) 1); val lemma2 = result(); goal thy "!!f. x:A ==> (UN z:A. {z}*f(z))``{x} = f(x)"; @@ -106,11 +106,11 @@ by (REPEAT (eresolve_tac [allE,impE] 1)); by (etac lemma1 1); by (asm_full_simp_tac (AC_ss addsimps [lemma2, lemma3] - addcongs [Pi_cong]) 1); + addcongs [Pi_cong]) 1); qed "AC4_AC3"; (* ********************************************************************** *) -(* AC3 ==> AC1 *) +(* AC3 ==> AC1 *) (* ********************************************************************** *) goal thy "!!A. b~:A ==> (PROD x:{a:A. id(A)`a~=b}. id(A)`x) = (PROD x:A. x)"; @@ -127,7 +127,7 @@ qed "AC3_AC1"; (* ********************************************************************** *) -(* AC4 ==> AC5 *) +(* AC4 ==> AC5 *) (* ********************************************************************** *) goalw thy (range_def::AC_defs) "!!Z. AC4 ==> AC5"; @@ -138,7 +138,7 @@ by (rtac bexI 1); by (rtac Pi_type 2 THEN (assume_tac 2)); by (fast_tac (ZF_cs addSDs [apply_type] - addSEs [fun_is_rel RS converse_type RS subsetD RS SigmaD2]) 2); + addSEs [fun_is_rel RS converse_type RS subsetD RS SigmaD2]) 2); by (rtac ballI 1); by (rtac apply_equality 1 THEN (assume_tac 2)); by (etac domainE 1); @@ -148,7 +148,7 @@ (* ********************************************************************** *) -(* AC5 ==> AC4, Rubin & Rubin, p. 11 *) +(* AC5 ==> AC4, Rubin & Rubin, p. 11 *) (* ********************************************************************** *) goal thy "!!A. R <= A*B ==> (lam x:R. fst(x)) : R -> A"; @@ -158,8 +158,8 @@ goalw thy [range_def] "!!A. R <= A*B ==> range(lam x:R. fst(x)) = domain(R)"; by (rtac equalityI 1); by (fast_tac (AC_cs addSEs [lamE, Pair_inject] - addEs [subst_elem] - addSDs [converseD, Pair_fst_snd_eq]) 1); + addEs [subst_elem] + addSDs [converseD, Pair_fst_snd_eq]) 1); by (rtac subsetI 1); by (etac domainE 1); by (rtac domainI 1); @@ -173,13 +173,13 @@ val lemma3 = result(); goal thy "!!g. [| R <= A*B; g: C->R; ALL x:C. (lam z:R. fst(z))` (g`x) = x |] \ -\ ==> (lam x:C. snd(g`x)): (PROD x:C. R``{x})"; +\ ==> (lam x:C. snd(g`x)): (PROD x:C. R``{x})"; by (rtac lam_type 1); by (dtac apply_type 1 THEN (assume_tac 1)); by (dtac bspec 1 THEN (assume_tac 1)); by (rtac imageI 1); by (resolve_tac [subsetD RS Pair_fst_snd_eq RSN (2, subst_elem)] 1 - THEN (REPEAT (assume_tac 1))); + THEN (REPEAT (assume_tac 1))); by (asm_full_simp_tac AC_ss 1); val lemma4 = result(); @@ -193,7 +193,7 @@ (* ********************************************************************** *) -(* AC1 <-> AC6 *) +(* AC1 <-> AC6 *) (* ********************************************************************** *) goalw thy AC_defs "AC1 <-> AC6"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC7_AC9.ML --- a/src/ZF/AC/AC7_AC9.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC7_AC9.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,16 +1,16 @@ -(* Title: ZF/AC/AC7-AC9.ML +(* Title: ZF/AC/AC7-AC9.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proofs needed to state that AC7, AC8 and AC9 are equivalent to the previous instances of AC. *) (* ********************************************************************** *) -(* Lemmas used in the proofs AC7 ==> AC6 and AC9 ==> AC1 *) -(* - Sigma_fun_space_not0 *) -(* - all_eqpoll_imp_pair_eqpoll *) -(* - Sigma_fun_space_eqpoll *) +(* Lemmas used in the proofs AC7 ==> AC6 and AC9 ==> AC1 *) +(* - Sigma_fun_space_not0 *) +(* - all_eqpoll_imp_pair_eqpoll *) +(* - Sigma_fun_space_eqpoll *) (* ********************************************************************** *) goal ZF.thy "!!A. [| C~:A; B:A |] ==> B~=C"; @@ -19,25 +19,25 @@ goal thy "!!A. [| 0~:A; B:A |] ==> (nat->Union(A))*B ~= 0"; by (fast_tac (ZF_cs addSDs [Sigma_empty_iff RS iffD1] - addDs [fun_space_emptyD, mem_not_eq_not_mem] - addEs [equals0D] - addSIs [equals0I,UnionI]) 1); + addDs [fun_space_emptyD, mem_not_eq_not_mem] + addEs [equals0D] + addSIs [equals0I,UnionI]) 1); val Sigma_fun_space_not0 = result(); goal thy "!!A C. (ALL B:A. B eqpoll C) ==> (ALL B1:A. ALL B2:A. B1 eqpoll B2)"; by (REPEAT (rtac ballI 1)); by (resolve_tac [bspec RS (bspec RSN (2, eqpoll_sym RSN (2, eqpoll_trans)))] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); val all_eqpoll_imp_pair_eqpoll = result(); goal thy "!!A. [| ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a)); b:A \ -\ |] ==> P(b)=R(b)"; +\ |] ==> P(b)=R(b)"; by (dtac bspec 1 THEN (assume_tac 1)); by (asm_full_simp_tac ZF_ss 1); val if_eqE1 = result(); goal thy "!!A. ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a)) \ -\ ==> ALL a:A. a~=b --> Q(a)=S(a)"; +\ ==> ALL a:A. a~=b --> Q(a)=S(a)"; by (rtac ballI 1); by (rtac impI 1); by (dtac bspec 1 THEN (assume_tac 1)); @@ -46,17 +46,17 @@ goal thy "!!A. [| (lam x:A. f(x))=(lam x:A. g(x)); a:A |] ==> f(a)=g(a)"; by (fast_tac (ZF_cs addDs [subsetD] - addSIs [lamI] - addEs [equalityE, lamE]) 1); + addSIs [lamI] + addEs [equalityE, lamE]) 1); val lam_eqE = result(); goalw thy [inj_def] - "!!A. C:A ==> (lam g:(nat->Union(A))*C. \ -\ (lam n:nat. if(n=0, snd(g), fst(g)`(n #- 1)))) \ -\ : inj((nat->Union(A))*C, (nat->Union(A)) ) "; + "!!A. C:A ==> (lam g:(nat->Union(A))*C. \ +\ (lam n:nat. if(n=0, snd(g), fst(g)`(n #- 1)))) \ +\ : inj((nat->Union(A))*C, (nat->Union(A)) ) "; by (rtac CollectI 1); by (fast_tac (ZF_cs addSIs [lam_type,RepFunI,if_type,snd_type,apply_type, - fst_type,diff_type,nat_succI,nat_0I]) 1); + fst_type,diff_type,nat_succI,nat_0I]) 1); by (REPEAT (resolve_tac [ballI, impI] 1)); by (asm_full_simp_tac ZF_ss 1); by (REPEAT (etac SigmaE 1)); @@ -69,20 +69,20 @@ by (dresolve_tac [nat_succI RSN (2, lam_eqE)] 1 THEN (assume_tac 1)); by (asm_full_simp_tac (AC_ss addsimps [succ_not_0 RS if_not_P]) 1); by (fast_tac (AC_cs addSEs [diff_succ_succ RS (diff_0 RSN (2, trans)) RS subst] - addSIs [nat_0I]) 1); + addSIs [nat_0I]) 1); val lemma = result(); goal thy "!!A. [| C:A; 0~:A |] ==> (nat->Union(A)) * C eqpoll (nat->Union(A))"; by (rtac eqpollI 1); by (fast_tac (ZF_cs addSEs [prod_lepoll_self, not_sym RS not_emptyE, - subst_elem] addEs [swap]) 2); + subst_elem] addEs [swap]) 2); by (rewtac lepoll_def); by (fast_tac (ZF_cs addSIs [lemma]) 1); val Sigma_fun_space_eqpoll = result(); (* ********************************************************************** *) -(* AC6 ==> AC7 *) +(* AC6 ==> AC7 *) (* ********************************************************************** *) goalw thy AC_defs "!!Z. AC6 ==> AC7"; @@ -90,9 +90,9 @@ qed "AC6_AC7"; (* ********************************************************************** *) -(* AC7 ==> AC6, Rubin & Rubin p. 12, Theorem 2.8 *) -(* The case of the empty family of sets added in order to complete *) -(* the proof. *) +(* AC7 ==> AC6, Rubin & Rubin p. 12, Theorem 2.8 *) +(* The case of the empty family of sets added in order to complete *) +(* the proof. *) (* ********************************************************************** *) goal thy "!!y. y: (PROD B:A. Y*B) ==> (lam B:A. snd(y`B)): (PROD B:A. B)"; @@ -100,19 +100,19 @@ val lemma1_1 = result(); goal thy "!!A. y: (PROD B:{Y*C. C:A}. B) \ -\ ==> (lam B:A. y`(Y*B)): (PROD B:A. Y*B)"; +\ ==> (lam B:A. y`(Y*B)): (PROD B:A. Y*B)"; by (fast_tac (ZF_cs addSIs [lam_type, apply_type]) 1); val lemma1_2 = result(); goal thy "!!A. (PROD B:{(nat->Union(A))*C. C:A}. B) ~= 0 \ -\ ==> (PROD B:A. B) ~= 0"; +\ ==> (PROD B:A. B) ~= 0"; by (fast_tac (ZF_cs addSIs [equals0I,lemma1_1, lemma1_2] - addSEs [equals0D]) 1); + addSEs [equals0D]) 1); val lemma1 = result(); goal thy "!!A. 0 ~: A ==> 0 ~: {(nat -> Union(A)) * C. C:A}"; by (fast_tac (ZF_cs addEs [RepFunE, - Sigma_fun_space_not0 RS not_sym RS notE]) 1); + Sigma_fun_space_not0 RS not_sym RS notE]) 1); val lemma2 = result(); goalw thy AC_defs "!!Z. AC7 ==> AC6"; @@ -124,18 +124,18 @@ by (etac allE 1); by (etac impE 1 THEN (assume_tac 2)); by (fast_tac (AC_cs addSEs [RepFunE] - addSIs [lemma2, all_eqpoll_imp_pair_eqpoll, - Sigma_fun_space_eqpoll]) 1); + addSIs [lemma2, all_eqpoll_imp_pair_eqpoll, + Sigma_fun_space_eqpoll]) 1); qed "AC7_AC6"; (* ********************************************************************** *) -(* AC1 ==> AC8 *) +(* AC1 ==> AC8 *) (* ********************************************************************** *) goalw thy [eqpoll_def] - "!!A. ALL B:A. EX B1 B2. B= & B1 eqpoll B2 \ -\ ==> 0 ~: { bij(fst(B),snd(B)). B:A }"; + "!!A. ALL B:A. EX B1 B2. B= & B1 eqpoll B2 \ +\ ==> 0 ~: { bij(fst(B),snd(B)). B:A }"; by (rtac notI 1); by (etac RepFunE 1); by (dtac bspec 1 THEN (assume_tac 1)); @@ -146,7 +146,7 @@ val lemma1 = result(); goal thy "!!A. [| f: (PROD X:RepFun(A,p). X); D:A |] \ -\ ==> (lam x:A. f`p(x))`D : p(D)"; +\ ==> (lam x:A. f`p(x))`D : p(D)"; by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1)); by (fast_tac (AC_cs addSEs [apply_type]) 1); val lemma2 = result(); @@ -162,13 +162,13 @@ (* ********************************************************************** *) -(* AC8 ==> AC9 *) -(* - this proof replaces the following two from Rubin & Rubin: *) -(* AC8 ==> AC1 and AC1 ==> AC9 *) +(* AC8 ==> AC9 *) +(* - this proof replaces the following two from Rubin & Rubin: *) +(* AC8 ==> AC1 and AC1 ==> AC9 *) (* ********************************************************************** *) goal thy "!!A. ALL B1:A. ALL B2:A. B1 eqpoll B2 ==> \ -\ ALL B:A*A. EX B1 B2. B= & B1 eqpoll B2"; +\ ALL B:A*A. EX B1 B2. B= & B1 eqpoll B2"; by (fast_tac ZF_cs 1); val lemma1 = result(); @@ -187,37 +187,37 @@ (* ********************************************************************** *) -(* AC9 ==> AC1 *) +(* AC9 ==> AC1 *) (* The idea of this proof comes from "Equivalents of the Axiom of Choice" *) -(* by Rubin & Rubin. But (x * y) is not necessarily equipollent to *) -(* (x * y) Un {0} when y is a set of total functions acting from nat to *) -(* Union(A) -- therefore we have used the set (y * nat) instead of y. *) +(* by Rubin & Rubin. But (x * y) is not necessarily equipollent to *) +(* (x * y) Un {0} when y is a set of total functions acting from nat to *) +(* Union(A) -- therefore we have used the set (y * nat) instead of y. *) (* ********************************************************************** *) (* Rules nedded to prove lemma1 *) val snd_lepoll_SigmaI = prod_lepoll_self RS ((prod_commute_eqpoll RS eqpoll_imp_lepoll) RSN (2,lepoll_trans)); val lemma1_cs = FOL_cs addSEs [UnE, RepFunE] - addSIs [all_eqpoll_imp_pair_eqpoll, ballI, - nat_cons_eqpoll RS eqpoll_trans] - addEs [Sigma_fun_space_not0 RS not_emptyE] - addIs [snd_lepoll_SigmaI, eqpoll_refl RSN - (2, prod_eqpoll_cong), Sigma_fun_space_eqpoll]; + addSIs [all_eqpoll_imp_pair_eqpoll, ballI, + nat_cons_eqpoll RS eqpoll_trans] + addEs [Sigma_fun_space_not0 RS not_emptyE] + addIs [snd_lepoll_SigmaI, eqpoll_refl RSN + (2, prod_eqpoll_cong), Sigma_fun_space_eqpoll]; goal thy "!!A. [| 0~:A; A~=0 |] \ -\ ==> ALL B1: ({((nat->Union(A))*B)*nat. B:A} \ -\ Un {cons(0,((nat->Union(A))*B)*nat). B:A}). \ -\ ALL B2: ({((nat->Union(A))*B)*nat. B:A} \ -\ Un {cons(0,((nat->Union(A))*B)*nat). B:A}). \ -\ B1 eqpoll B2"; +\ ==> ALL B1: ({((nat->Union(A))*B)*nat. B:A} \ +\ Un {cons(0,((nat->Union(A))*B)*nat). B:A}). \ +\ ALL B2: ({((nat->Union(A))*B)*nat. B:A} \ +\ Un {cons(0,((nat->Union(A))*B)*nat). B:A}). \ +\ B1 eqpoll B2"; by (fast_tac lemma1_cs 1); val lemma1 = result(); goal thy "!!A. ALL B1:{(F*B)*N. B:A} Un {cons(0,(F*B)*N). B:A}. \ -\ ALL B2:{(F*B)*N. B:A} \ -\ Un {cons(0,(F*B)*N). B:A}. f` : bij(B1, B2) \ -\ ==> (lam B:A. snd(fst((f`)`0))) : \ -\ (PROD X:A. X)"; +\ ALL B2:{(F*B)*N. B:A} \ +\ Un {cons(0,(F*B)*N). B:A}. f` : bij(B1, B2) \ +\ ==> (lam B:A. snd(fst((f`)`0))) : \ +\ (PROD X:A. X)"; by (rtac lam_type 1); by (rtac snd_type 1); by (rtac fst_type 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/AC_Equiv.ML --- a/src/ZF/AC/AC_Equiv.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/AC_Equiv.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/AC_Equiv.ML +(* Title: ZF/AC/AC_Equiv.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski *) @@ -45,8 +45,8 @@ (* ********************************************************************** *) (* The following two theorms are useful when rewriting only one instance *) -(* of a definition *) -(* first one for definitions of formulae and the second for terms *) +(* of a definition *) +(* first one for definitions of formulae and the second for terms *) val prems = goal ZF.thy "(A == B) ==> A <-> B"; by (asm_simp_tac (ZF_ss addsimps prems) 1); @@ -101,7 +101,7 @@ (* used only in Hartog.ML *) goal Cardinal.thy - "!!A r. well_ord(A,r) ==> ordertype(A, r Int A*A) = ordertype(A,r)"; + "!!A r. well_ord(A,r) ==> ordertype(A, r Int A*A) = ordertype(A,r)"; by (res_inst_tac [("P","%a. ordertype(A,a)=ordertype(A,r)")] (rvimage_id RS subst) 1); by (eresolve_tac [id_bij RS bij_ordertype_vimage] 1); @@ -109,13 +109,13 @@ (* used only in AC16_lemmas.ML *) goalw CardinalArith.thy [InfCard_def] - "!!i. [| ~Finite(i); Card(i) |] ==> InfCard(i)"; + "!!i. [| ~Finite(i); Card(i) |] ==> InfCard(i)"; by (asm_simp_tac (ZF_ss addsimps [Card_is_Ord RS nat_le_infinite_Ord]) 1); val Inf_Card_is_InfCard = result(); goal thy "(THE z. {x}={z}) = x"; by (fast_tac (AC_cs addSIs [the_equality] - addSEs [singleton_eq_iff RS iffD1 RS sym]) 1); + addSEs [singleton_eq_iff RS iffD1 RS sym]) 1); val the_element = result(); goal thy "(lam x:A. {x}) : bij(A, {{x}. x:A})"; @@ -125,15 +125,15 @@ val lam_sing_bij = result(); val [major,minor] = goal thy - "[| f : Pi(A,B); (!!x. x:A ==> B(x)<=C(x)) |] ==> f : Pi(A,C)"; + "[| f : Pi(A,B); (!!x. x:A ==> B(x)<=C(x)) |] ==> f : Pi(A,C)"; by (fast_tac (AC_cs addSIs [major RS Pi_type, minor RS subsetD, - major RS apply_type]) 1); + major RS apply_type]) 1); val Pi_weaken_type = result(); val [major, minor] = goalw thy [inj_def] - "[| f:inj(A, B); (!!a. a:A ==> f`a : C) |] ==> f:inj(A,C)"; + "[| f:inj(A, B); (!!a. a:A ==> f`a : C) |] ==> f:inj(A,C)"; by (fast_tac (AC_cs addSEs [minor] - addSIs [major RS CollectD1 RS Pi_type, major RS CollectD2]) 1); + addSIs [major RS CollectD1 RS Pi_type, major RS CollectD2]) 1); val inj_strengthen_type = result(); goal thy "A*B=0 <-> A=0 | B=0"; @@ -146,13 +146,13 @@ goalw thy [Finite_def] "~Finite(nat)"; by (fast_tac (AC_cs addSDs [eqpoll_imp_lepoll] - addIs [Ord_nat RSN (2, ltI) RS lt_not_lepoll RS notE]) 1); + addIs [Ord_nat RSN (2, ltI) RS lt_not_lepoll RS notE]) 1); val nat_not_Finite = result(); val le_imp_lepoll = le_imp_subset RS subset_imp_lepoll; (* ********************************************************************** *) -(* Another elimination rule for EX! *) +(* Another elimination rule for EX! *) (* ********************************************************************** *) goal thy "!!x. [| EX! x. P(x); P(x); P(y) |] ==> x=y"; @@ -163,14 +163,14 @@ val ex1_two_eq = result(); (* ********************************************************************** *) -(* image of a surjection *) +(* image of a surjection *) (* ********************************************************************** *) goalw thy [surj_def] "!!f. f : surj(A, B) ==> f``A = B"; by (etac CollectE 1); by (resolve_tac [subset_refl RSN (2, image_fun) RS ssubst] 1 THEN (assume_tac 1)); by (fast_tac (AC_cs addSEs [RepFunE, apply_type] - addSIs [RepFunI] addIs [equalityI]) 1); + addSIs [RepFunI] addIs [equalityI]) 1); val surj_image_eq = result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/Cardinal_aux.ML --- a/src/ZF/AC/Cardinal_aux.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/Cardinal_aux.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/Cardinal_aux.ML +(* Title: ZF/AC/Cardinal_aux.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski Auxiliary lemmas concerning cardinalities *) @@ -8,8 +8,8 @@ open Cardinal_aux; (* ********************************************************************** *) -(* Lemmas involving ordinals and cardinalities used in the proofs *) -(* concerning AC16 and DC *) +(* Lemmas involving ordinals and cardinalities used in the proofs *) +(* concerning AC16 and DC *) (* ********************************************************************** *) val nat_0_in_2 = Ord_0 RS le_refl RS leI RS ltD; @@ -19,7 +19,7 @@ by (rtac Card_cardinal 1); by (resolve_tac [Card_nat RS (Card_def RS def_imp_iff RS iffD1 RS ssubst)] 1); by (resolve_tac [nat_le_infinite_Ord RS le_imp_lepoll - RSN (2, well_ord_Memrel RS well_ord_lepoll_imp_Card_le)] 1 + RSN (2, well_ord_Memrel RS well_ord_lepoll_imp_Card_le)] 1 THEN REPEAT (assume_tac 1)); val Inf_Ord_imp_InfCard_cardinal = result(); @@ -34,46 +34,46 @@ goalw thy [sum_def] "A+A = 2*A"; by (fast_tac (AC_cs addIs [equalityI] - addSIs [singletonI, nat_0_in_2, succI1] - addSEs [succE, singletonE]) 1); + addSIs [singletonI, nat_0_in_2, succI1] + addSEs [succE, singletonE]) 1); val sum_eq_2_times = result(); val sum_lepoll_prod = [sum_eq_2_times RS equalityD1 RS subset_imp_lepoll, - asm_rl, lepoll_refl] MRS (prod_lepoll_mono RSN (2, lepoll_trans)) - |> standard; + asm_rl, lepoll_refl] MRS (prod_lepoll_mono RSN (2, lepoll_trans)) + |> standard; goal thy "!!A. [| A lepoll B; 2 lepoll A |] ==> A+B lepoll A*B"; by (REPEAT (ares_tac [[sum_lepoll_mono, sum_lepoll_prod] - MRS lepoll_trans, lepoll_refl] 1)); + MRS lepoll_trans, lepoll_refl] 1)); val lepoll_imp_sum_lepoll_prod = result(); goal thy "!!A. [| A eqpoll i; B eqpoll i; ~Finite(i); Ord(i) |] \ -\ ==> A Un B eqpoll i"; +\ ==> A Un B eqpoll i"; by (rtac eqpollI 1); by (eresolve_tac [subset_imp_lepoll RSN (2, eqpoll_sym RS eqpoll_imp_lepoll - RS lepoll_trans)] 2); + RS lepoll_trans)] 2); by (fast_tac AC_cs 2); by (resolve_tac [Un_lepoll_sum RS lepoll_trans] 1); by (resolve_tac [lepoll_imp_sum_lepoll_prod RS lepoll_trans] 1); by (eresolve_tac [eqpoll_sym RSN (2, eqpoll_trans) RS eqpoll_imp_lepoll] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (resolve_tac [nat_le_infinite_Ord RS le_imp_lepoll RS - (Ord_nat RS (nat_2I RS OrdmemD) RS subset_imp_lepoll RS lepoll_trans) - RS (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, lepoll_trans))] 1 + (Ord_nat RS (nat_2I RS OrdmemD) RS subset_imp_lepoll RS lepoll_trans) + RS (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, lepoll_trans))] 1 THEN (REPEAT (assume_tac 1))); by (eresolve_tac [prod_eqpoll_cong RS eqpoll_imp_lepoll RS lepoll_trans] 1 THEN (assume_tac 1)); by (resolve_tac [Inf_Ord_imp_InfCard_cardinal RSN (2, well_ord_Memrel RS - well_ord_InfCard_square_eq) RS eqpoll_imp_lepoll] 1 + well_ord_InfCard_square_eq) RS eqpoll_imp_lepoll] 1 THEN REPEAT (assume_tac 1)); val Un_eqpoll_Inf_Ord = result(); goalw thy [lepoll_def] "!!A. [| A lepoll B; C lepoll D; B Int D = 0 |] \ -\ ==> A Un C lepoll B Un D"; +\ ==> A Un C lepoll B Un D"; by (REPEAT (etac exE 1)); by (res_inst_tac [("x","lam a: A Un C. if(a:A, f`a, fa`a)")] exI 1); by (res_inst_tac [("d","%z. if(z:B, converse(f)`z, converse(fa)`z)")] - lam_injective 1); + lam_injective 1); by (split_tac [expand_if] 1); by (etac UnE 1); by (fast_tac (AC_cs addSEs [inj_is_fun RS apply_type]) 1); @@ -83,7 +83,7 @@ by (REPEAT (split_tac [expand_if] 1)); by (rtac conjI 1); by (fast_tac (AC_cs addSEs [left_inverse, inj_is_fun RS apply_type] - addEs [swap]) 1); + addEs [swap]) 1); by (rtac impI 1); by (etac UnE 1); by (contr_tac 1); @@ -102,14 +102,14 @@ by (split_tac [expand_if] 1); by (asm_full_simp_tac (ZF_ss addsimps [double_Diff_sing, Diff_eq_0_iff]) 1); by (fast_tac (ZF_cs addSIs [the_equality, equalityI, equals0I] - addEs [equalityE] addSEs [singletonE]) 1); + addEs [equalityE] addSEs [singletonE]) 1); val paired_bij_lemma = result(); goal thy "(lam y:{{y,z}. y:x}. if(y-{z}=0, z, THE w. y-{z}={w})) \ -\ : bij({{y,z}. y:x}, x)"; +\ : bij({{y,z}. y:x}, x)"; by (res_inst_tac [("d","%a. {a,z}")] lam_bijective 1); by (TRYALL (fast_tac (AC_cs addSEs [RepFunE] addSIs [RepFunI] - addss (AC_ss addsimps [paired_bij_lemma])))); + addss (AC_ss addsimps [paired_bij_lemma])))); val paired_bij = result(); goalw thy [eqpoll_def] "{{y,z}. y:x} eqpoll x"; @@ -121,15 +121,15 @@ val ex_eqpoll_disjoint = result(); goal thy "!!A. [| A lepoll i; B lepoll i; ~Finite(i); Ord(i) |] \ -\ ==> A Un B lepoll i"; +\ ==> A Un B lepoll i"; by (res_inst_tac [("A1","i"), ("C1","i")] (ex_eqpoll_disjoint RS exE) 1); by (etac conjE 1); by (dresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RSN (2, lepoll_trans)] 1); by (assume_tac 1); by (resolve_tac [Un_lepoll_Un RS lepoll_trans] 1 THEN (REPEAT (assume_tac 1))); by (eresolve_tac [eqpoll_refl RSN (2, Un_eqpoll_Inf_Ord) RS - eqpoll_imp_lepoll] 1 - THEN (REPEAT (assume_tac 1))); + eqpoll_imp_lepoll] 1 + THEN (REPEAT (assume_tac 1))); val Un_lepoll_Inf_Ord = result(); goal thy "!!P. [| P(i); i:j; Ord(j) |] ==> (LEAST i. P(i)) : j"; @@ -141,7 +141,7 @@ val Least_in_Ord = result(); goal thy "!!x. [| well_ord(x,r); y<=x; y lepoll succ(n); n:nat |] \ -\ ==> y-{THE b. first(b,y,r)} lepoll n"; +\ ==> y-{THE b. first(b,y,r)} lepoll n"; by (res_inst_tac [("Q","y=0")] (excluded_middle RS disjE) 1); by (fast_tac (AC_cs addSIs [Diff_sing_lepoll, the_first_in]) 1); by (res_inst_tac [("b","y-{THE b. first(b, y, r)}")] subst 1); @@ -161,7 +161,7 @@ val UN_sing_lepoll = result(); goal thy "!!a T. [| well_ord(T, R); ~Finite(a); Ord(a); n:nat |] ==> \ -\ ALL f. (ALL b:a. f`b lepoll n & f`b <= T) --> (UN b:a. f`b) lepoll a"; +\ ALL f. (ALL b:a. f`b lepoll n & f`b <= T) --> (UN b:a. f`b) lepoll a"; by (etac nat_induct 1); by (rtac allI 1); by (rtac impI 1); @@ -183,16 +183,16 @@ val UN_fun_lepoll_lemma = result(); goal thy "!!a f. [| ALL b:a. f`b lepoll n & f`b <= T; well_ord(T, R); \ -\ ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. f`b) lepoll a"; +\ ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. f`b) lepoll a"; by (eresolve_tac [UN_fun_lepoll_lemma RS allE] 1 THEN (REPEAT (assume_tac 1))); by (fast_tac AC_cs 1); val UN_fun_lepoll = result(); goal thy "!!a f. [| ALL b:a. F(b) lepoll n & F(b) <= T; well_ord(T, R); \ -\ ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. F(b)) lepoll a"; +\ ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. F(b)) lepoll a"; by (rtac impE 1 THEN (assume_tac 3)); by (res_inst_tac [("f","lam b:a. F(b)")] (UN_fun_lepoll) 2 - THEN (TRYALL assume_tac)); + THEN (TRYALL assume_tac)); by (simp_tac AC_ss 2); by (asm_full_simp_tac AC_ss 1); val UN_lepoll = result(); @@ -223,7 +223,7 @@ val disj_Un_eqpoll_sum = result(); goalw thy [lepoll_def, eqpoll_def] - "!!X a. a lepoll X ==> EX Y. Y<=X & a eqpoll Y"; + "!!X a. a lepoll X ==> EX Y. Y<=X & a eqpoll Y"; by (etac exE 1); by (forward_tac [subset_refl RSN (2, restrict_bij)] 1); by (res_inst_tac [("x","f``a")] exI 1); @@ -231,7 +231,7 @@ val lepoll_imp_eqpoll_subset = result(); (* ********************************************************************** *) -(* Some other lemmas *) +(* Some other lemmas *) (* ********************************************************************** *) goal thy "!!m n. [| m:nat; n:nat |] ==> m + n eqpoll m #+ n"; @@ -245,14 +245,14 @@ val nat_sum_eqpoll_sum = result(); val eqpoll_ordertype = - ordermap_bij RS (exI RS (eqpoll_def RS (def_imp_iff RS iffD2))); + ordermap_bij RS (exI RS (eqpoll_def RS (def_imp_iff RS iffD2))); goalw thy [lesspoll_def] - "!!a. [| b lesspoll a; well_ord(a,r); well_ord(b,s) |] \ -\ ==> ordertype(b,s) < ordertype(a,r)"; + "!!a. [| b lesspoll a; well_ord(a,r); well_ord(b,s) |] \ +\ ==> ordertype(b,s) < ordertype(a,r)"; by (forw_inst_tac [("A2","b")] - ([Ord_ordertype, Ord_ordertype] MRS Ord_linear) 1 - THEN REPEAT (assume_tac 1)); + ([Ord_ordertype, Ord_ordertype] MRS Ord_linear) 1 + THEN REPEAT (assume_tac 1)); by (etac disjE 1); by (etac ltI 1); by (etac Ord_ordertype 1); @@ -276,18 +276,18 @@ val lesspoll_imp_lepoll = result(); goalw thy [lepoll_def] - "!!A. [| A lepoll B; well_ord(B,r) |] ==> EX s. well_ord(A,s)"; + "!!A. [| A lepoll B; well_ord(B,r) |] ==> EX s. well_ord(A,s)"; by (fast_tac (AC_cs addSEs [well_ord_rvimage]) 1); val lepoll_well_ord = result(); goal thy "!!A a. [| A lesspoll a; Ord(a) |] ==> EX b. b A lesspoll B | A eqpoll B"; @@ -302,7 +302,7 @@ goal thy "!!m. [| m le n; n:nat |] ==> m:nat"; by (fast_tac (AC_cs addSDs [nat_succI RS (Ord_nat RSN (2, OrdmemD))] - addSEs [ltE]) 1); + addSEs [ltE]) 1); val le_in_nat = result(); goalw thy [Finite_def] "!!A. [| Finite(A); Finite(B) |] ==> Finite(A Un B)"; @@ -311,10 +311,10 @@ by (dtac sum_lepoll_mono 1 THEN (assume_tac 1)); by (dresolve_tac [nat_sum_eqpoll_sum RS eqpoll_imp_lepoll RSN (2, Un_lepoll_sum RS lepoll_trans RS lepoll_trans)] 1 - THEN (REPEAT (assume_tac 1))); + THEN (REPEAT (assume_tac 1))); by (forw_inst_tac [("n2","na")] - (add_type RS nat_into_Ord RSN (2, lepoll_imp_ex_le_eqpoll)) 1 - THEN (REPEAT (assume_tac 1))); + (add_type RS nat_into_Ord RSN (2, lepoll_imp_ex_le_eqpoll)) 1 + THEN (REPEAT (assume_tac 1))); by (fast_tac (AC_cs addSEs [add_type RSN (2, le_in_nat)]) 1); val Finite_Un = result(); @@ -328,44 +328,44 @@ val lt_Card_imp_lesspoll = result(); (* ********************************************************************** *) -(* Diff_lesspoll_eqpoll_Card *) +(* Diff_lesspoll_eqpoll_Card *) (* ********************************************************************** *) goal thy "!!A B. [| A eqpoll a; ~Finite(a); Card(a); B lesspoll a; \ -\ A-B lesspoll a |] ==> P"; +\ A-B lesspoll a |] ==> P"; by (REPEAT (eresolve_tac [lesspoll_imp_ex_lt_eqpoll RS exE, - Card_is_Ord, conjE] 1)); + Card_is_Ord, conjE] 1)); by (forw_inst_tac [("j","xa")] (lt_Ord RS (lt_Ord RSN (2, Un_upper1_le))) 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (forw_inst_tac [("j","xa")] (lt_Ord RS (lt_Ord RSN (2, Un_upper2_le))) 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (dtac Un_least_lt 1 THEN (assume_tac 1)); by (dresolve_tac [le_imp_lepoll RSN - (2, eqpoll_imp_lepoll RS lepoll_trans)] 1 - THEN (assume_tac 1)); + (2, eqpoll_imp_lepoll RS lepoll_trans)] 1 + THEN (assume_tac 1)); by (dresolve_tac [le_imp_lepoll RSN - (2, eqpoll_imp_lepoll RS lepoll_trans)] 1 - THEN (assume_tac 1)); + (2, eqpoll_imp_lepoll RS lepoll_trans)] 1 + THEN (assume_tac 1)); by (res_inst_tac [("Q","Finite(x Un xa)")] (excluded_middle RS disjE) 1); by (dresolve_tac [[lepoll_Finite, lepoll_Finite] MRS Finite_Un] 2 - THEN (REPEAT (assume_tac 2))); + THEN (REPEAT (assume_tac 2))); by (dresolve_tac [subset_Un_Diff RS subset_imp_lepoll RS lepoll_Finite] 2); by (fast_tac (AC_cs - addDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite]) 2); + addDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite]) 2); by (dresolve_tac [ Un_lepoll_Inf_Ord] 1 - THEN (REPEAT (assume_tac 1))); + THEN (REPEAT (assume_tac 1))); by (fast_tac (AC_cs addSEs [ltE, Ord_in_Ord]) 1); by (dresolve_tac [subset_Un_Diff RS subset_imp_lepoll RS lepoll_trans RSN - (3, lt_Card_imp_lesspoll RS lepoll_lesspoll_lesspoll)] 1 - THEN (TRYALL assume_tac)); + (3, lt_Card_imp_lesspoll RS lepoll_lesspoll_lesspoll)] 1 + THEN (TRYALL assume_tac)); by (fast_tac (AC_cs addSDs [lesspoll_def RS def_imp_iff RS iffD1]) 1); val Diff_lesspoll_eqpoll_Card_lemma = result(); goal thy "!!A. [| A eqpoll a; ~Finite(a); Card(a); B lesspoll a |] \ -\ ==> A - B eqpoll a"; +\ ==> A - B eqpoll a"; by (rtac swap 1 THEN (fast_tac AC_cs 1)); by (rtac Diff_lesspoll_eqpoll_Card_lemma 1 THEN (REPEAT (assume_tac 1))); by (fast_tac (AC_cs addSIs [lesspoll_def RS def_imp_iff RS iffD2, - subset_imp_lepoll RS (eqpoll_imp_lepoll RSN (2, lepoll_trans))]) 1); + subset_imp_lepoll RS (eqpoll_imp_lepoll RSN (2, lepoll_trans))]) 1); val Diff_lesspoll_eqpoll_Card = result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/DC.ML --- a/src/ZF/AC/DC.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/DC.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/DC.ML - ID: $Id$ - Author: Krzysztof Grabczewski +(* Title: ZF/AC/DC.ML + ID: $Id$ + Author: Krzysztof Grabczewski The proofs concerning the Axiom of Dependent Choice *) @@ -8,41 +8,41 @@ open DC; (* ********************************************************************** *) -(* DC ==> DC(omega) *) -(* *) -(* The scheme of the proof: *) -(* *) -(* Assume DC. Let R and x satisfy the premise of DC(omega). *) -(* *) -(* Define XX and RR as follows: *) -(* *) -(* XX = (UN n:nat. {f:n->X. ALL k:n. : R}) *) -(* f RR g iff domain(g)=succ(domain(f)) & *) -(* restrict(g, domain(f)) = f *) -(* *) -(* Then RR satisfies the hypotheses of DC. *) -(* So applying DC: *) -(* *) -(* EX f:nat->XX. ALL n:nat. f`n RR f`succ(n) *) -(* *) -(* Thence *) -(* *) -(* ff = {. n:nat} *) -(* *) -(* is the desired function. *) -(* *) +(* DC ==> DC(omega) *) +(* *) +(* The scheme of the proof: *) +(* *) +(* Assume DC. Let R and x satisfy the premise of DC(omega). *) +(* *) +(* Define XX and RR as follows: *) +(* *) +(* XX = (UN n:nat. {f:n->X. ALL k:n. : R}) *) +(* f RR g iff domain(g)=succ(domain(f)) & *) +(* restrict(g, domain(f)) = f *) +(* *) +(* Then RR satisfies the hypotheses of DC. *) +(* So applying DC: *) +(* *) +(* EX f:nat->XX. ALL n:nat. f`n RR f`succ(n) *) +(* *) +(* Thence *) +(* *) +(* ff = {. n:nat} *) +(* *) +(* is the desired function. *) +(* *) (* ********************************************************************** *) goal thy "{z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \ -\ & restrict(snd(z), domain(fst(z))) = fst(z)} <= XX*XX"; +\ & restrict(snd(z), domain(fst(z))) = fst(z)} <= XX*XX"; by (fast_tac AC_cs 1); val lemma1_1 = result(); goal thy "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. : R) \ -\ ==> {z: (UN n:nat. {f:n->X. ALL k:n. : R}) * \ -\ (UN n:nat. {f:n->X. ALL k:n. : R}). \ -\ domain(snd(z))=succ(domain(fst(z))) \ -\ & restrict(snd(z), domain(fst(z))) = fst(z)} ~= 0"; +\ ==> {z: (UN n:nat. {f:n->X. ALL k:n. : R}) * \ +\ (UN n:nat. {f:n->X. ALL k:n. : R}). \ +\ domain(snd(z))=succ(domain(fst(z))) \ +\ & restrict(snd(z), domain(fst(z))) = fst(z)} ~= 0"; by (etac ballE 1); by (eresolve_tac [empty_subsetI RS PowI RSN (2, notE)] 2); by (eresolve_tac [nat_0I RS n_lesspoll_nat RSN (2, impE)] 1); @@ -52,31 +52,31 @@ by (rtac SigmaI 1); by (fast_tac (AC_cs addSIs [nat_0I RS UN_I, empty_fun]) 1); by (fast_tac (AC_cs addSIs [nat_1I RS UN_I, singleton_fun RS Pi_type] - addss (AC_ss addsimps [[lepoll_refl, succI1] MRS lepoll_1_is_sing, - apply_singleton_eq, image_0])) 1); + addss (AC_ss addsimps [[lepoll_refl, succI1] MRS lepoll_1_is_sing, + apply_singleton_eq, image_0])) 1); by (asm_full_simp_tac (AC_ss addsimps [domain_0, restrict_0, domain_sing, - [lepoll_refl, succI1] MRS lepoll_1_is_sing]) 1); + [lepoll_refl, succI1] MRS lepoll_1_is_sing]) 1); val lemma1_2 = result(); goal thy "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. : R) \ -\ ==> range({z: (UN n:nat. {f:n->X. ALL k:n. : R}) * \ -\ (UN n:nat. {f:n->X. ALL k:n. : R}). \ -\ domain(snd(z))=succ(domain(fst(z))) \ -\ & restrict(snd(z), domain(fst(z))) = fst(z)}) \ -\ <= domain({z: (UN n:nat. {f:n->X. ALL k:n. : R}) * \ -\ (UN n:nat. {f:n->X. ALL k:n. : R}). \ -\ domain(snd(z))=succ(domain(fst(z))) \ -\ & restrict(snd(z), domain(fst(z))) = fst(z)})"; +\ ==> range({z: (UN n:nat. {f:n->X. ALL k:n. : R}) * \ +\ (UN n:nat. {f:n->X. ALL k:n. : R}). \ +\ domain(snd(z))=succ(domain(fst(z))) \ +\ & restrict(snd(z), domain(fst(z))) = fst(z)}) \ +\ <= domain({z: (UN n:nat. {f:n->X. ALL k:n. : R}) * \ +\ (UN n:nat. {f:n->X. ALL k:n. : R}). \ +\ domain(snd(z))=succ(domain(fst(z))) \ +\ & restrict(snd(z), domain(fst(z))) = fst(z)})"; by (rtac range_subset_domain 1); by (rtac subsetI 2); by (etac CollectD1 2); by (etac UN_E 1); by (etac CollectE 1); by (dresolve_tac [fun_is_rel RS image_subset RS PowI RSN (2, bspec)] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (eresolve_tac [[n_lesspoll_nat, nat_into_Ord RSN (2, image_Ord_lepoll)] - MRS lepoll_lesspoll_lesspoll RSN (2, impE)] 1 - THEN REPEAT (assume_tac 1)); + MRS lepoll_lesspoll_lesspoll RSN (2, impE)] 1 + THEN REPEAT (assume_tac 1)); by (etac bexE 1); by (res_inst_tac [("x","cons(, g)")] exI 1); by (rtac CollectI 1); @@ -87,23 +87,23 @@ by (rtac CollectI 1); by (etac cons_fun_type2 1 THEN (assume_tac 1)); by (fast_tac (AC_cs addSEs [succE] addss (AC_ss - addsimps [cons_image_n, cons_val_n, cons_image_k, cons_val_k])) 1); + addsimps [cons_image_n, cons_val_n, cons_image_k, cons_val_k])) 1); by (asm_full_simp_tac (AC_ss - addsimps [domain_cons, domain_of_fun, succ_def, restrict_cons_eq]) 1); + addsimps [domain_cons, domain_of_fun, succ_def, restrict_cons_eq]) 1); val lemma1_3 = result(); goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. : R}); \ -\ RR = {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \ -\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \ -\ ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. : R) \ -\ |] ==> RR <= XX*XX & RR ~= 0 & range(RR) <= domain(RR)"; +\ RR = {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \ +\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \ +\ ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. : R) \ +\ |] ==> RR <= XX*XX & RR ~= 0 & range(RR) <= domain(RR)"; by (fast_tac (AC_cs addSIs [lemma1_1] addSEs [lemma1_2, lemma1_3]) 1); val lemma1 = result(); goal thy "!!f. [| : {z:XX*XX. \ -\ domain(snd(z))=succ(domain(fst(z))) & Q(z)}; \ -\ XX = (UN n:nat. {f:n->X. ALL k:n. : R}); f:k->X \ -\ |] ==> g:succ(k)->X"; +\ domain(snd(z))=succ(domain(fst(z))) & Q(z)}; \ +\ XX = (UN n:nat. {f:n->X. ALL k:n. : R}); f:k->X \ +\ |] ==> g:succ(k)->X"; by (etac CollectE 1); by (dtac SigmaD2 1); by (hyp_subst_tac 1); @@ -117,12 +117,12 @@ val lemma2_1 = result(); goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. : R}); \ -\ ALL n:nat. : \ -\ {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \ -\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \ -\ f: nat -> XX; n:nat \ -\ |] ==> EX k:nat. f`succ(n) : k -> X & n:k \ -\ & : R"; +\ ALL n:nat. : \ +\ {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \ +\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \ +\ f: nat -> XX; n:nat \ +\ |] ==> EX k:nat. f`succ(n) : k -> X & n:k \ +\ & : R"; by (etac nat_induct 1); by (dresolve_tac [nat_1I RSN (2, apply_type)] 1); by (dresolve_tac [nat_0I RSN (2, bspec)] 1); @@ -131,28 +131,28 @@ by (etac CollectE 1); by (rtac bexI 1 THEN (assume_tac 2)); by (fast_tac (AC_cs addSEs [nat_0_le RS leE, ltD, ltD RSN (2, bspec)] - addEs [sym RS trans RS succ_neq_0, domain_of_fun]) 1); + addEs [sym RS trans RS succ_neq_0, domain_of_fun]) 1); by (etac bexE 1); by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1)); by (etac conjE 1); by (dtac lemma2_1 1 THEN REPEAT (assume_tac 1)); by (hyp_subst_tac 1); by (dresolve_tac [nat_succI RS nat_succI RSN (2, apply_type)] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (etac UN_E 1); by (etac CollectE 1); by (dresolve_tac [[domain_of_fun RS sym, domain_of_fun] MRS trans] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (fast_tac (AC_cs addSEs [nat_succI, nat_into_Ord RS succ_in_succ] - addSDs [nat_into_Ord RS succ_in_succ RSN (2, bspec)]) 1); + addSDs [nat_into_Ord RS succ_in_succ RSN (2, bspec)]) 1); val lemma2 = result(); goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. : R}); \ -\ ALL n:nat. : \ -\ {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \ -\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \ -\ f: nat -> XX; n:nat \ -\ |] ==> ALL x:n. f`succ(n)`x = f`succ(x)`x"; +\ ALL n:nat. : \ +\ {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \ +\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \ +\ f: nat -> XX; n:nat \ +\ |] ==> ALL x:n. f`succ(n)`x = f`succ(x)`x"; by (etac nat_induct 1); by (fast_tac AC_cs 1); by (rtac ballI 1); @@ -169,31 +169,31 @@ by (asm_full_simp_tac AC_ss 1); by (dtac lemma2 1 THEN REPEAT (assume_tac 1)); by (fast_tac (FOL_cs addSDs [domain_of_fun] - addSEs [bexE, nat_into_Ord RSN (2, OrdmemD) RS subsetD]) 1); + addSEs [bexE, nat_into_Ord RSN (2, OrdmemD) RS subsetD]) 1); val lemma3_1 = result(); goal thy "!!n. ALL x:n. f`succ(n)`x = f`succ(x)`x \ -\ ==> {f`succ(x)`x. x:n} = {f`succ(n)`x. x:n}"; +\ ==> {f`succ(x)`x. x:n} = {f`succ(n)`x. x:n}"; by (asm_full_simp_tac AC_ss 1); val lemma3_2 = result(); goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. : R}); \ -\ ALL n:nat. : \ -\ {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \ -\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \ -\ f: nat -> XX; n:nat \ -\ |] ==> (lam x:nat. f`succ(x)`x) `` n = f`succ(n)``n"; +\ ALL n:nat. : \ +\ {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \ +\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \ +\ f: nat -> XX; n:nat \ +\ |] ==> (lam x:nat. f`succ(x)`x) `` n = f`succ(n)``n"; by (etac natE 1); by (asm_full_simp_tac (AC_ss addsimps [image_0]) 1); by (resolve_tac [image_lam RS ssubst] 1); by (fast_tac (AC_cs addSEs [[nat_succI, Ord_nat] MRS OrdmemD]) 1); by (resolve_tac [lemma3_1 RS lemma3_2 RS ssubst] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (fast_tac (AC_cs addSEs [nat_succI]) 1); by (dresolve_tac [nat_succI RSN (4, lemma2)] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (fast_tac (AC_cs addSEs [conjE, image_fun RS sym, - nat_into_Ord RSN (2, OrdmemD)]) 1); + nat_into_Ord RSN (2, OrdmemD)]) 1); val lemma3 = result(); goal thy "!!f. [| f:A->B; B<=C |] ==> f:A->C"; @@ -205,67 +205,67 @@ by (REPEAT (resolve_tac [allI, impI] 1)); by (REPEAT (eresolve_tac [conjE, allE] 1)); by (eresolve_tac [[refl, refl] MRS lemma1 RSN (2, impE)] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (etac bexE 1); by (res_inst_tac [("x","lam n:nat. f`succ(n)`n")] bexI 1); by (fast_tac (AC_cs addSIs [lam_type] addSDs [refl RS lemma2] - addSEs [fun_type_gen, apply_type]) 2); + addSEs [fun_type_gen, apply_type]) 2); by (rtac oallI 1); by (forward_tac [ltD RSN (3, refl RS lemma2)] 1 - THEN assume_tac 2); + THEN assume_tac 2); by (fast_tac (AC_cs addSEs [fun_type_gen]) 1); by (eresolve_tac [ltD RSN (3, refl RS lemma3) RS ssubst] 1 - THEN assume_tac 2); + THEN assume_tac 2); by (fast_tac (AC_cs addSEs [fun_type_gen]) 1); by (fast_tac (AC_cs addss AC_ss) 1); qed "DC0_DC_nat"; (* ********************************************************************** *) -(* DC(omega) ==> DC *) -(* *) -(* The scheme of the proof: *) -(* *) -(* Assume DC(omega). Let R and x satisfy the premise of DC. *) -(* *) -(* Define XX and RR as follows: *) -(* *) -(* XX = (UN n:nat. *) -(* {f:succ(n)->domain(R). ALL k:n. : R}) *) -(* RR = {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) *) -(* & (ALL f:fst(z). restrict(snd(z), domain(f)) = f)) | *) -(* (~ (EX g:XX. domain(g)=succ(UN f:fst(z). domain(f)) *) -(* & (ALL f:fst(z). restrict(g, domain(f)) = f)) & *) -(* snd(z)={<0,x>})} *) -(* *) -(* Then XX and RR satisfy the hypotheses of DC(omega). *) -(* So applying DC: *) -(* *) -(* EX f:nat->XX. ALL n:nat. f``n RR f`n *) -(* *) -(* Thence *) -(* *) -(* ff = {. n:nat} *) -(* *) -(* is the desired function. *) -(* *) +(* DC(omega) ==> DC *) +(* *) +(* The scheme of the proof: *) +(* *) +(* Assume DC(omega). Let R and x satisfy the premise of DC. *) +(* *) +(* Define XX and RR as follows: *) +(* *) +(* XX = (UN n:nat. *) +(* {f:succ(n)->domain(R). ALL k:n. : R}) *) +(* RR = {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) *) +(* & (ALL f:fst(z). restrict(snd(z), domain(f)) = f)) | *) +(* (~ (EX g:XX. domain(g)=succ(UN f:fst(z). domain(f)) *) +(* & (ALL f:fst(z). restrict(g, domain(f)) = f)) & *) +(* snd(z)={<0,x>})} *) +(* *) +(* Then XX and RR satisfy the hypotheses of DC(omega). *) +(* So applying DC: *) +(* *) +(* EX f:nat->XX. ALL n:nat. f``n RR f`n *) +(* *) +(* Thence *) +(* *) +(* ff = {. n:nat} *) +(* *) +(* is the desired function. *) +(* *) (* ********************************************************************** *) goalw thy [lesspoll_def, Finite_def] - "!!A. A lesspoll nat ==> Finite(A)"; + "!!A. A lesspoll nat ==> Finite(A)"; by (fast_tac (AC_cs addSDs [ltD, lepoll_imp_ex_le_eqpoll] - addSIs [Ord_nat]) 1); + addSIs [Ord_nat]) 1); val lesspoll_nat_is_Finite = result(); goal thy "!!n. n:nat ==> ALL A. (A eqpoll n & A <= X) --> A : Fin(X)"; by (etac nat_induct 1); by (rtac allI 1); by (fast_tac (AC_cs addSIs [Fin.emptyI] - addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1); + addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1); by (rtac allI 1); by (rtac impI 1); by (etac conjE 1); by (resolve_tac [eqpoll_succ_imp_not_empty RS not_emptyE] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (forward_tac [Diff_sing_eqpoll] 1 THEN (assume_tac 1)); by (etac allE 1); by (etac impE 1); @@ -285,72 +285,72 @@ val Finite_Fin = result(); goal thy "!!x. x: X ==> {<0,x>}: (UN n:nat. \ -\ {f:succ(n)->X. ALL k:n. : R})"; +\ {f:succ(n)->X. ALL k:n. : R})"; by (fast_tac (AC_cs addSIs [nat_0I RS UN_I, singleton_fun RS Pi_type] - addss (AC_ss addsimps [[lepoll_refl, succI1] MRS lepoll_1_is_sing, - apply_singleton_eq])) 1); + addss (AC_ss addsimps [[lepoll_refl, succI1] MRS lepoll_1_is_sing, + apply_singleton_eq])) 1); val singleton_in_funs = result(); goal thy - "!!X. [| XX = (UN n:nat. \ -\ {f:succ(n)->domain(R). ALL k:n. : R}); \ -\ RR = {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \ -\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f)) | \ -\ (~ (EX g:XX. domain(g)=succ(UN f:fst(z). domain(f)) \ -\ & (ALL f:fst(z). restrict(g, domain(f)) = f)) & snd(z)={<0,x>})}; \ -\ range(R) <= domain(R); x:domain(R) \ -\ |] ==> RR <= Pow(XX)*XX & \ -\ (ALL Y:Pow(XX). Y lesspoll nat --> (EX x:XX. :RR))"; + "!!X. [| XX = (UN n:nat. \ +\ {f:succ(n)->domain(R). ALL k:n. : R}); \ +\ RR = {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \ +\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f)) | \ +\ (~ (EX g:XX. domain(g)=succ(UN f:fst(z). domain(f)) \ +\ & (ALL f:fst(z). restrict(g, domain(f)) = f)) & snd(z)={<0,x>})}; \ +\ range(R) <= domain(R); x:domain(R) \ +\ |] ==> RR <= Pow(XX)*XX & \ +\ (ALL Y:Pow(XX). Y lesspoll nat --> (EX x:XX. :RR))"; by (rtac conjI 1); by (fast_tac (FOL_cs addSEs [FinD RS PowI, SigmaE, CollectE] - addSIs [subsetI, SigmaI]) 1); + addSIs [subsetI, SigmaI]) 1); by (rtac ballI 1); by (rtac impI 1); by (dresolve_tac [[lesspoll_nat_is_Finite, PowD] MRS Finite_Fin] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (excluded_middle_tac "EX g:XX. domain(g)=succ(UN f:Y. domain(f)) \ -\ & (ALL f:Y. restrict(g, domain(f)) = f)" 1); +\ & (ALL f:Y. restrict(g, domain(f)) = f)" 1); by (fast_tac (AC_cs addss AC_ss) 2); by (fast_tac (FOL_cs addSEs [CollectE, singleton_in_funs] - addSIs [SigmaI] addIs [bexI] addss AC_ss) 1); + addSIs [SigmaI] addIs [bexI] addss AC_ss) 1); val lemma1 = result(); goal thy "!!f. [| f:nat->X; n:nat |] ==> \ -\ (UN x:f``succ(n). P(x)) = P(f`n) Un (UN x:f``n. P(x))"; +\ (UN x:f``succ(n). P(x)) = P(f`n) Un (UN x:f``n. P(x))"; by (asm_full_simp_tac (AC_ss - addsimps [Ord_nat RSN (2, OrdmemD) RSN (2, image_fun), - [nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1); + addsimps [Ord_nat RSN (2, OrdmemD) RSN (2, image_fun), + [nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1); val UN_image_succ_eq = result(); goal thy "!!f. [| (UN x:f``n. P(x)) = y; P(f`n) = succ(y); \ -\ f:nat -> X; n:nat |] ==> (UN x:f``succ(n). P(x)) = succ(y)"; +\ f:nat -> X; n:nat |] ==> (UN x:f``succ(n). P(x)) = succ(y)"; by (asm_full_simp_tac (AC_ss addsimps [UN_image_succ_eq]) 1); by (fast_tac (AC_cs addSIs [equalityI]) 1); val UN_image_succ_eq_succ = result(); goal thy "!!f. [| f: (UN n:nat. {g:succ(n) -> D. P(g, n)}); \ -\ domain(f)=succ(x); x=y |] ==> f`y : D"; +\ domain(f)=succ(x); x=y |] ==> f`y : D"; by (fast_tac (AC_cs addEs [apply_type] - addSDs [[sym, domain_of_fun] MRS trans]) 1); + addSDs [[sym, domain_of_fun] MRS trans]) 1); val apply_UN_type = result(); goal thy "!!f. [| f : nat -> X; n:nat |] ==> f``succ(n) = cons(f`n, f``n)"; by (asm_full_simp_tac (AC_ss - addsimps [nat_succI, Ord_nat RSN (2, OrdmemD), image_fun]) 1); + addsimps [nat_succI, Ord_nat RSN (2, OrdmemD), image_fun]) 1); val image_fun_succ = result(); goal thy "!!f. [| domain(f`n) = succ(u); f : nat -> (UN n:nat. \ -\ {f:succ(n)->domain(R). ALL k:n. : R}); \ -\ u=k; n:nat \ -\ |] ==> f`n : succ(k) -> domain(R)"; +\ {f:succ(n)->domain(R). ALL k:n. : R}); \ +\ u=k; n:nat \ +\ |] ==> f`n : succ(k) -> domain(R)"; by (dtac apply_type 1 THEN (assume_tac 1)); by (fast_tac (AC_cs addEs [UN_E, domain_eq_imp_fun_type]) 1); val f_n_type = result(); goal thy "!!f. [| f : nat -> (UN n:nat. \ -\ {f:succ(n)->domain(R). ALL k:n. : R}); \ -\ domain(f`n) = succ(k); n:nat \ -\ |] ==> ALL i:k. : R"; +\ {f:succ(n)->domain(R). ALL k:n. : R}); \ +\ domain(f`n) = succ(k); n:nat \ +\ |] ==> ALL i:k. : R"; by (dtac apply_type 1 THEN (assume_tac 1)); by (etac UN_E 1); by (etac CollectE 1); @@ -360,8 +360,8 @@ val f_n_pairs_in_R = result(); goalw thy [restrict_def] - "!!f. [| restrict(f, domain(x))=x; f:n->X; domain(x) <= n \ -\ |] ==> restrict(cons(, f), domain(x)) = x"; + "!!f. [| restrict(f, domain(x))=x; f:n->X; domain(x) <= n \ +\ |] ==> restrict(cons(, f), domain(x)) = x"; by (eresolve_tac [sym RS trans RS sym] 1); by (rtac fun_extension 1); by (fast_tac (AC_cs addSIs [lam_type]) 1); @@ -370,11 +370,11 @@ val restrict_cons_eq_restrict = result(); goal thy "!!f. [| ALL x:f``n. restrict(f`n, domain(x))=x; \ -\ f : nat -> (UN n:nat. \ -\ {f:succ(n)->domain(R). ALL k:n. : R}); \ -\ n:nat; domain(f`n) = succ(n); \ -\ (UN x:f``n. domain(x)) <= n |] \ -\ ==> ALL x:f``succ(n). restrict(cons(, f`n), domain(x))=x"; +\ f : nat -> (UN n:nat. \ +\ {f:succ(n)->domain(R). ALL k:n. : R}); \ +\ n:nat; domain(f`n) = succ(n); \ +\ (UN x:f``n. domain(x)) <= n |] \ +\ ==> ALL x:f``succ(n). restrict(cons(, f`n), domain(x))=x"; by (rtac ballI 1); by (asm_full_simp_tac (AC_ss addsimps [image_fun_succ]) 1); by (dtac f_n_type 1 THEN REPEAT (ares_tac [refl] 1)); @@ -385,29 +385,29 @@ val all_in_image_restrict_eq = result(); goal thy "!!X. [| ALL b : \ -\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \ -\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f)) | \ -\ (~ (EX g:XX. domain(g)=succ(UN f:fst(z). domain(f)) \ -\ & (ALL f:fst(z). restrict(g, domain(f)) = f)) & snd(z)={<0,x>})}; \ -\ XX = (UN n:nat. \ -\ {f:succ(n)->domain(R). ALL k:n. : R}); \ -\ f: nat -> XX; range(R) <= domain(R); x:domain(R) \ -\ |] ==> ALL b : \ -\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \ -\ & (UN f:fst(z). domain(f)) = b \ -\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f))}"; +\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \ +\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f)) | \ +\ (~ (EX g:XX. domain(g)=succ(UN f:fst(z). domain(f)) \ +\ & (ALL f:fst(z). restrict(g, domain(f)) = f)) & snd(z)={<0,x>})}; \ +\ XX = (UN n:nat. \ +\ {f:succ(n)->domain(R). ALL k:n. : R}); \ +\ f: nat -> XX; range(R) <= domain(R); x:domain(R) \ +\ |] ==> ALL b : \ +\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \ +\ & (UN f:fst(z). domain(f)) = b \ +\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f))}"; by (rtac oallI 1); by (dtac ltD 1); by (etac nat_induct 1); by (dresolve_tac [[nat_0I, Ord_nat] MRS ltI RSN (2, ospec)] 1); by (fast_tac (FOL_cs addss - (ZF_ss addsimps [image_0, singleton_fun RS domain_of_fun, - [lepoll_refl, succI1] MRS lepoll_1_is_sing, singleton_in_funs])) 1); + (ZF_ss addsimps [image_0, singleton_fun RS domain_of_fun, + [lepoll_refl, succI1] MRS lepoll_1_is_sing, singleton_in_funs])) 1); by (dresolve_tac [[nat_succI, Ord_nat] MRS ltI RSN (2, ospec)] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (REPEAT (eresolve_tac [conjE, CollectE, disjE] 1)); by (fast_tac (FOL_cs addSEs [trans, subst_context] - addSIs [UN_image_succ_eq_succ] addss AC_ss) 1); + addSIs [UN_image_succ_eq_succ] addss AC_ss) 1); by (etac conjE 1); by (etac notE 1); by (asm_full_simp_tac (AC_ss addsimps [UN_image_succ_eq_succ]) 1); @@ -418,100 +418,100 @@ by (forward_tac [f_n_type] 1 THEN REPEAT (assume_tac 1)); by (res_inst_tac [("x","cons(, f`xa)")] bexI 1); by (fast_tac (FOL_cs - addEs [subst_context RSN (2, trans) RS domain_cons_eq_succ, - subst_context, all_in_image_restrict_eq, trans, equalityD1]) 1); + addEs [subst_context RSN (2, trans) RS domain_cons_eq_succ, + subst_context, all_in_image_restrict_eq, trans, equalityD1]) 1); by (rtac UN_I 1); by (etac nat_succI 1); by (rtac CollectI 1); by (eresolve_tac [rangeI RSN (2, subsetD) RSN (2, cons_fun_type2)] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (rtac ballI 1); by (etac succE 1); by (asm_full_simp_tac (AC_ss addsimps [cons_val_n, cons_val_k]) 1); by (dresolve_tac [domain_of_fun RSN (2, f_n_pairs_in_R)] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (dtac bspec 1 THEN (assume_tac 1)); by (asm_full_simp_tac (AC_ss - addsimps [nat_into_Ord RS succ_in_succ, succI2, cons_val_k]) 1); + addsimps [nat_into_Ord RS succ_in_succ, succI2, cons_val_k]) 1); val simplify_recursion = result(); goal thy "!!X. [| XX = (UN n:nat. \ -\ {f:succ(n)->domain(R). ALL k:n. : R}); \ -\ ALL b : \ -\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \ -\ & (UN f:fst(z). domain(f)) = b \ -\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f))}; \ -\ f: nat -> XX; range(R) <= domain(R); x:domain(R); n:nat \ -\ |] ==> f`n : succ(n) -> domain(R) \ -\ & (ALL i:n. :R)"; +\ {f:succ(n)->domain(R). ALL k:n. : R}); \ +\ ALL b : \ +\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \ +\ & (UN f:fst(z). domain(f)) = b \ +\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f))}; \ +\ f: nat -> XX; range(R) <= domain(R); x:domain(R); n:nat \ +\ |] ==> f`n : succ(n) -> domain(R) \ +\ & (ALL i:n. :R)"; by (dtac ospec 1); by (eresolve_tac [Ord_nat RSN (2, ltI)] 1); by (etac CollectE 1); by (asm_full_simp_tac AC_ss 1); by (rtac conjI 1); by (fast_tac (AC_cs - addSEs [trans RS domain_eq_imp_fun_type, subst_context]) 1); + addSEs [trans RS domain_eq_imp_fun_type, subst_context]) 1); by (fast_tac (FOL_cs - addSEs [conjE, f_n_pairs_in_R, trans, subst_context]) 1); + addSEs [conjE, f_n_pairs_in_R, trans, subst_context]) 1); val lemma2 = result(); goal thy "!!n. [| XX = (UN n:nat. \ -\ {f:succ(n)->domain(R). ALL k:n. : R}); \ -\ ALL b : \ -\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \ -\ & (UN f:fst(z). domain(f)) = b \ -\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f))}; \ -\ f : nat -> XX; n:nat; range(R) <= domain(R); x:domain(R) \ -\ |] ==> f`n`n = f`succ(n)`n"; +\ {f:succ(n)->domain(R). ALL k:n. : R}); \ +\ ALL b : \ +\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \ +\ & (UN f:fst(z). domain(f)) = b \ +\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f))}; \ +\ f : nat -> XX; n:nat; range(R) <= domain(R); x:domain(R) \ +\ |] ==> f`n`n = f`succ(n)`n"; by (forward_tac [lemma2 RS conjunct1 RS domain_of_fun] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (dresolve_tac [[nat_succI, Ord_nat] MRS ltI RSN (2, ospec)] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (asm_full_simp_tac AC_ss 1); by (REPEAT (etac conjE 1)); by (etac ballE 1); by (eresolve_tac [restrict_eq_imp_val_eq RS sym] 1); by (fast_tac (AC_cs addSEs [ssubst]) 1); by (asm_full_simp_tac (AC_ss - addsimps [[nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1); + addsimps [[nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1); val lemma3 = result(); goalw thy [DC_def, DC0_def] "!!Z. DC(nat) ==> DC0"; by (REPEAT (resolve_tac [allI, impI] 1)); by (REPEAT (eresolve_tac [asm_rl, conjE, ex_in_domain RS exE, allE] 1)); by (eresolve_tac [[refl, refl] MRS lemma1 RSN (2, impE)] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (etac bexE 1); by (dresolve_tac [refl RSN (2, simplify_recursion)] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (res_inst_tac [("x","lam n:nat. f`n`n")] bexI 1); by (rtac lam_type 2); by (eresolve_tac [[refl RS lemma2 RS conjunct1, succI1] MRS apply_type] 2 - THEN REPEAT (assume_tac 2)); + THEN REPEAT (assume_tac 2)); by (rtac ballI 1); by (forward_tac [refl RS (nat_succI RSN (6, lemma2)) RS conjunct2] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (dresolve_tac [refl RS lemma3] 1 THEN REPEAT (assume_tac 1)); by (asm_full_simp_tac (AC_ss addsimps [nat_succI]) 1); qed "DC_nat_DC0"; (* ********************************************************************** *) -(* ALL K. Card(K) --> DC(K) ==> WO3 *) +(* ALL K. Card(K) --> DC(K) ==> WO3 *) (* ********************************************************************** *) goalw thy [lesspoll_def] - "!!A. [| ~ A lesspoll B; C lesspoll B |] ==> A - C ~= 0"; + "!!A. [| ~ A lesspoll B; C lesspoll B |] ==> A - C ~= 0"; by (fast_tac (AC_cs addSDs [Diff_eq_0_iff RS iffD1 RS subset_imp_lepoll] - addSIs [eqpollI] addEs [notE] addSEs [eqpollE, lepoll_trans]) 1); + addSIs [eqpollI] addEs [notE] addSEs [eqpollE, lepoll_trans]) 1); val lemma1 = result(); val [f_type, Ord_a, not_eq] = goalw thy [inj_def] - "[| f:a->X; Ord(a); (!!b c. [| b f`b~=f`c) \ -\ |] ==> f:inj(a, X)"; + "[| f:a->X; Ord(a); (!!b c. [| b f`b~=f`c) \ +\ |] ==> f:inj(a, X)"; by (resolve_tac [f_type RS CollectI] 1); by (REPEAT (resolve_tac [ballI,impI] 1)); by (resolve_tac [Ord_a RS Ord_in_Ord RS Ord_linear_lt] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (eres_inst_tac [("j","x")] (Ord_a RS Ord_in_Ord) 1); by (REPEAT (fast_tac (AC_cs addDs [not_eq, not_eq RS not_sym]) 1)); val fun_Ord_inj = result(); @@ -521,42 +521,42 @@ val value_in_image = result(); goalw thy [DC_def, WO3_def] - "!!Z. ALL K. Card(K) --> DC(K) ==> WO3"; + "!!Z. ALL K. Card(K) --> DC(K) ==> WO3"; by (rtac allI 1); by (excluded_middle_tac "A lesspoll Hartog(A)" 1); by (fast_tac (AC_cs addSDs [lesspoll_imp_ex_lt_eqpoll] - addSIs [Ord_Hartog, leI RS le_imp_subset]) 2); + addSIs [Ord_Hartog, leI RS le_imp_subset]) 2); by (REPEAT (eresolve_tac [allE, impE] 1)); by (rtac Card_Hartog 1); by (eres_inst_tac [("x","A")] allE 1); by (eres_inst_tac [("x","{z:Pow(A)*A . fst(z) \ -\ lesspoll Hartog(A) & snd(z) ~: fst(z)}")] allE 1); +\ lesspoll Hartog(A) & snd(z) ~: fst(z)}")] allE 1); by (asm_full_simp_tac AC_ss 1); by (etac impE 1); by (fast_tac (AC_cs addEs [lemma1 RS not_emptyE]) 1); by (etac bexE 1); by (resolve_tac [exI RS (lepoll_def RS (def_imp_iff RS iffD2)) - RS (HartogI RS notE)] 1); + RS (HartogI RS notE)] 1); by (resolve_tac [Ord_Hartog RSN (2, fun_Ord_inj)] 1 THEN (assume_tac 1)); by (dresolve_tac [Ord_Hartog RSN (2, OrdmemD) RSN (2, - ltD RSN (3, value_in_image))] 1 - THEN REPEAT (assume_tac 1)); + ltD RSN (3, value_in_image))] 1 + THEN REPEAT (assume_tac 1)); by (fast_tac (AC_cs addSDs [Ord_Hartog RSN (2, ltI) RSN (2, ospec)] - addEs [subst]) 1); + addEs [subst]) 1); qed "DC_WO3"; (* ********************************************************************** *) -(* WO1 ==> ALL K. Card(K) --> DC(K) *) +(* WO1 ==> ALL K. Card(K) --> DC(K) *) (* ********************************************************************** *) goal thy - "!!a. [| Ord(a); b:a |] ==> (lam x:a. P(x))``b = (lam x:b. P(x))``b"; + "!!a. [| Ord(a); b:a |] ==> (lam x:a. P(x))``b = (lam x:b. P(x))``b"; by (rtac images_eq 1); by (REPEAT (fast_tac (AC_cs addSEs [RepFunI, OrdmemD] - addSIs [lam_type]) 2)); + addSIs [lam_type]) 2)); by (rtac ballI 1); by (dresolve_tac [OrdmemD RS subsetD] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (asm_full_simp_tac AC_ss 1); val lam_images_eq = result(); @@ -570,40 +570,40 @@ val lam_type_RepFun = result(); goal thy "!!Z. [| ALL Y:Pow(X). Y lesspoll a --> (EX x:X. : R); \ -\ b:a; Z:Pow(X); Z lesspoll a |] \ -\ ==> {x:X. : R} ~= 0"; +\ b:a; Z:Pow(X); Z lesspoll a |] \ +\ ==> {x:X. : R} ~= 0"; by (fast_tac (FOL_cs addEs [bexE, equals0D] - addSDs [bspec] addIs [CollectI]) 1); + addSDs [bspec] addIs [CollectI]) 1); val lemma_ = result(); goal thy "!!K. [| Card(K); well_ord(X,Q); \ -\ ALL Y:Pow(X). Y lesspoll K --> (EX x:X. : R); b:K |] \ -\ ==> ff(b, X, Q, R) : {x:X. <(lam c:b. ff(c, X, Q, R))``b, x> : R}"; +\ ALL Y:Pow(X). Y lesspoll K --> (EX x:X. : R); b:K |] \ +\ ==> ff(b, X, Q, R) : {x:X. <(lam c:b. ff(c, X, Q, R))``b, x> : R}"; by (res_inst_tac [("P","b:K")] impE 1 THEN TRYALL assume_tac); by (res_inst_tac [("i","b")] (Card_is_Ord RS Ord_in_Ord RS trans_induct) 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (rtac impI 1); by (resolve_tac [ff_def RS def_transrec RS ssubst] 1); by (etac the_first_in 1); by (fast_tac AC_cs 1); by (asm_full_simp_tac (AC_ss - addsimps [[lam_type_RepFun, subset_refl] MRS image_fun]) 1); + addsimps [[lam_type_RepFun, subset_refl] MRS image_fun]) 1); by (etac lemma_ 1 THEN (assume_tac 1)); by (fast_tac (AC_cs addSEs [RepFunE, impE, notE] - addEs [Card_is_Ord RSN (2, OrdmemD) RS subsetD]) 1); + addEs [Card_is_Ord RSN (2, OrdmemD) RS subsetD]) 1); by (fast_tac (AC_cs addSEs [[in_Card_imp_lesspoll, RepFun_lepoll] - MRS lepoll_lesspoll_lesspoll]) 1); + MRS lepoll_lesspoll_lesspoll]) 1); val lemma = result(); goalw thy [DC_def, WO1_def] - "!!Z. WO1 ==> ALL K. Card(K) --> DC(K)"; + "!!Z. WO1 ==> ALL K. Card(K) --> DC(K)"; by (REPEAT (resolve_tac [allI,impI] 1)); by (REPEAT (eresolve_tac [allE,exE,conjE] 1)); by (res_inst_tac [("x","lam b:K. ff(b, X, Ra, R)")] bexI 1); by (rtac lam_type 2); by (resolve_tac [lemma RS CollectD1] 2 THEN REPEAT (assume_tac 2)); by (asm_full_simp_tac (AC_ss - addsimps [[Card_is_Ord, ltD] MRS lam_images_eq]) 1); + addsimps [[Card_is_Ord, ltD] MRS lam_images_eq]) 1); by (fast_tac (AC_cs addSEs [ltE, lemma RS CollectD2]) 1); qed" WO1_DC_Card"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/DC_lemmas.ML --- a/src/ZF/AC/DC_lemmas.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/DC_lemmas.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,13 +1,13 @@ -(* Title: ZF/AC/DC_lemmas.ML +(* Title: ZF/AC/DC_lemmas.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski More general lemmas used in the proofs concerning DC *) val [prem] = goalw thy [lepoll_def] - "Ord(a) ==> {P(b). b:a} lepoll a"; + "Ord(a) ==> {P(b). b:a} lepoll a"; by (res_inst_tac [("x","lam z:RepFun(a,P). LEAST i. z=P(i)")] exI 1); by (res_inst_tac [("d","%z. P(z)")] (sym RSN (2, lam_injective)) 1); by (fast_tac (AC_cs addSEs [RepFunE] addSIs [Least_in_Ord, prem]) 1); @@ -21,12 +21,12 @@ by (etac eqpollE 1); by (rtac succ_lepoll_natE 1 THEN (assume_tac 2)); by (eresolve_tac [nat_succI RS (Ord_nat RSN (2, OrdmemD) RS - subset_imp_lepoll) RS lepoll_trans] 1 - THEN (assume_tac 1)); + subset_imp_lepoll) RS lepoll_trans] 1 + THEN (assume_tac 1)); val n_lesspoll_nat = result(); goalw thy [lepoll_def] - "!!f. [| f:X->Y; Ord(X) |] ==> f``X lepoll X"; + "!!f. [| f:X->Y; Ord(X) |] ==> f``X lepoll X"; by (res_inst_tac [("x","lam x:f``X. LEAST y. f`y = x")] exI 1); by (res_inst_tac [("d","%z. f`z")] lam_injective 1); by (fast_tac (AC_cs addSIs [Least_in_Ord, apply_equality]) 1); @@ -42,8 +42,8 @@ val restrict_0 = result(); val [major, minor] = goal thy - "[| (!!g. g:X ==> EX u. :R); R<=X*X \ -\ |] ==> range(R) <= domain(R)"; + "[| (!!g. g:X ==> EX u. :R); R<=X*X \ +\ |] ==> range(R) <= domain(R)"; by (rtac subsetI 1); by (etac rangeE 1); by (dresolve_tac [minor RS subsetD RS SigmaD2 RS major] 1); @@ -66,7 +66,7 @@ by (resolve_tac [Pi_iff_old RS iffD2] 1); by (rtac conjI 1); by (fast_tac (AC_cs addSIs [cons_subset, succI1] - addEs [fun_is_rel RS subset_trans RS subsetD, succI2]) 1); + addEs [fun_is_rel RS subset_trans RS subsetD, succI2]) 1); by (rtac ballI 1); by (etac succE 1); by (fast_tac (AC_cs addDs [domain_of_fun, domainI] addSEs [mem_irrefl]) 1); @@ -114,12 +114,12 @@ goal thy "!!k. [| Ord(k); i:k |] ==> succ(i) : succ(k)"; by (resolve_tac [Ord_linear RS disjE] 1 THEN (assume_tac 3)); by (REPEAT (fast_tac (AC_cs addSIs [Ord_succ] - addEs [Ord_in_Ord, mem_irrefl, mem_asym] - addSDs [succ_inject]) 1)); + addEs [Ord_in_Ord, mem_irrefl, mem_asym] + addSDs [succ_inject]) 1)); val succ_in_succ = result(); goalw thy [restrict_def] - "!!f. [| restrict(f, domain(g)) = g; x: domain(g) |] ==> f`x = g`x"; + "!!f. [| restrict(f, domain(g)) = g; x: domain(g) |] ==> f`x = g`x"; by (etac subst 1); by (asm_full_simp_tac AC_ss 1); val restrict_eq_imp_val_eq = result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/HH.ML --- a/src/ZF/AC/HH.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/HH.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/HH.ML +(* Title: ZF/AC/HH.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski Some properties of the recursive definition of HH used in the proofs of AC17 ==> AC1 @@ -11,12 +11,12 @@ open HH; (* ********************************************************************** *) -(* Lemmas useful in each of the three proofs *) +(* Lemmas useful in each of the three proofs *) (* ********************************************************************** *) goal thy "HH(f,x,a) = \ -\ (let z = x - (UN b:a. HH(f,x,b)) \ -\ in if(f`z:Pow(z)-{0}, f`z, {x}))"; +\ (let z = x - (UN b:a. HH(f,x,b)) \ +\ in if(f`z:Pow(z)-{0}, f`z, {x}))"; by (resolve_tac [HH_def RS def_transrec RS trans] 1); by (simp_tac ZF_ss 1); val HH_def_satisfies_eq = result(); @@ -42,13 +42,13 @@ val prems = goal thy "(!!y. y:A ==> P(y) = {x}) ==> x - (UN y:A. P(y)) = x"; by (asm_full_simp_tac (AC_ss addsimps prems) 1); by (fast_tac (AC_cs addSIs [equalityI] addSDs [prem] - addSEs [RepFunE, mem_irrefl]) 1); + addSEs [RepFunE, mem_irrefl]) 1); val Diff_UN_eq_self = result(); goal thy "!!a. x - (UN b:a. HH(f,x,b)) = x - (UN b:a1. HH(f,x,b)) \ -\ ==> HH(f,x,a) = HH(f,x,a1)"; +\ ==> HH(f,x,a) = HH(f,x,a1)"; by (resolve_tac [HH_def_satisfies_eq RS - (HH_def_satisfies_eq RS sym RSN (3, trans RS trans))] 1); + (HH_def_satisfies_eq RS sym RSN (3, trans RS trans))] 1); by (etac subst_context 1); val HH_eq = result(); @@ -58,7 +58,7 @@ by (rtac impI 1); by (resolve_tac [HH_eq RS trans] 1 THEN (assume_tac 2)); by (resolve_tac [leI RS le_imp_subset RS subset_imp_Diff_eq RS ssubst] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (res_inst_tac [("t","%z. z-?X")] subst_context 1); by (rtac Diff_UN_eq_self 1); by (dtac Ord_DiffE 1 THEN (assume_tac 1)); @@ -73,7 +73,7 @@ val HH_subset_x_lt_too = result(); goal thy "!!a. HH(f,x,a) : Pow(x)-{0} \ -\ ==> HH(f,x,a) : Pow(x - (UN b:a. HH(f,x,b)))-{0}"; +\ ==> HH(f,x,a) : Pow(x - (UN b:a. HH(f,x,b)))-{0}"; by (dresolve_tac [HH_def_satisfies_eq RS subst] 1); by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); by (asm_full_simp_tac (AC_ss addsimps [Let_def, Diff_subset RS PowI]) 1); @@ -90,26 +90,26 @@ val HH_eq_arg_lt = result(); goal thy "!!x. [| HH(f,x,v)=HH(f,x,w); HH(f,x,w): Pow(x)-{0}; \ -\ Ord(v); Ord(w) |] ==> v=w"; +\ Ord(v); Ord(w) |] ==> v=w"; by (res_inst_tac [("j","w")] Ord_linear_lt 1 THEN TRYALL assume_tac); by (resolve_tac [sym RS (ltD RSN (3, HH_eq_arg_lt))] 2 - THEN REPEAT (assume_tac 2)); + THEN REPEAT (assume_tac 2)); by (dtac subst_elem 1 THEN (assume_tac 1)); by (fast_tac (FOL_cs addDs [ltD] addSEs [HH_eq_arg_lt]) 1); val HH_eq_imp_arg_eq = result(); goalw thy [lepoll_def, inj_def] - "!!i. [| HH(f, x, i) : Pow(x)-{0}; Ord(i) |] ==> i lepoll Pow(x)-{0}"; + "!!i. [| HH(f, x, i) : Pow(x)-{0}; Ord(i) |] ==> i lepoll Pow(x)-{0}"; by (res_inst_tac [("x","lam j:i. HH(f, x, j)")] exI 1); by (asm_simp_tac AC_ss 1); by (fast_tac (FOL_cs addSEs [HH_eq_imp_arg_eq, Ord_in_Ord, HH_subset_x_lt_too] - addSIs [lam_type, ballI, ltI] addIs [bexI]) 1); + addSIs [lam_type, ballI, ltI] addIs [bexI]) 1); val HH_subset_x_imp_lepoll = result(); goal thy "HH(f, x, Hartog(Pow(x)-{0})) = {x}"; by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 2)); by (fast_tac (FOL_cs addSDs [HH_subset_x_imp_lepoll] - addSIs [Ord_Hartog] addSEs [HartogE]) 1); + addSIs [Ord_Hartog] addSEs [HartogE]) 1); val HH_Hartog_is_x = result(); goal thy "HH(f, x, LEAST i. HH(f, x, i) = {x}) = {x}"; @@ -124,20 +124,20 @@ val less_Least_subset_x = result(); (* ********************************************************************** *) -(* Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1 *) +(* Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1 *) (* ********************************************************************** *) goalw thy [inj_def] - "(lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a)) : \ -\ inj(LEAST i. HH(f,x,i)={x}, Pow(x)-{0})"; + "(lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a)) : \ +\ inj(LEAST i. HH(f,x,i)={x}, Pow(x)-{0})"; by (asm_full_simp_tac AC_ss 1); by (fast_tac (AC_cs addSIs [lam_type] addDs [less_Least_subset_x] - addSEs [HH_eq_imp_arg_eq, Ord_Least RS Ord_in_Ord]) 1); + addSEs [HH_eq_imp_arg_eq, Ord_Least RS Ord_in_Ord]) 1); val lam_Least_HH_inj_Pow = result(); goal thy "!!x. ALL a:(LEAST i. HH(f,x,i)={x}). EX z:x. HH(f,x,a) = {z} \ -\ ==> (lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \ -\ : inj(LEAST i. HH(f,x,i)={x}, {{y}. y:x})"; +\ ==> (lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \ +\ : inj(LEAST i. HH(f,x,i)={x}, {{y}. y:x})"; by (resolve_tac [lam_Least_HH_inj_Pow RS inj_strengthen_type] 1); by (asm_full_simp_tac AC_ss 1); by (fast_tac (AC_cs addSEs [RepFun_eqI]) 1); @@ -148,9 +148,9 @@ val elem_of_sing_eq = result(); goalw thy [surj_def] - "!!x. [| x - (UN a:A. F(a)) = 0; \ -\ ALL a:A. EX z:x. F(a) = {z} |] \ -\ ==> (lam a:A. F(a)) : surj(A, {{y}. y:x})"; + "!!x. [| x - (UN a:A. F(a)) = 0; \ +\ ALL a:A. EX z:x. F(a) = {z} |] \ +\ ==> (lam a:A. F(a)) : surj(A, {{y}. y:x})"; by (asm_full_simp_tac (AC_ss addsimps [Diff_eq_0_iff]) 1); by (rtac conjI 1); by (fast_tac (AC_cs addSIs [lam_type] addSEs [RepFun_eqI]) 1); @@ -167,42 +167,42 @@ goal thy "!!x. y:Pow(x)-{0} ==> x ~= 0"; by (fast_tac (AC_cs addSIs [equals0I, singletonI RS subst_elem] - addSDs [equals0D]) 1); + addSDs [equals0D]) 1); val not_emptyI2 = result(); goal thy "!!f. f`(x - (UN j:i. HH(f,x,j))): Pow(x - (UN j:i. HH(f,x,j)))-{0} \ -\ ==> HH(f, x, i) : Pow(x) - {0}"; +\ ==> HH(f, x, i) : Pow(x) - {0}"; by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); by (asm_full_simp_tac (AC_ss addsimps [Let_def, Diff_subset RS PowI, - not_emptyI2 RS if_P]) 1); + not_emptyI2 RS if_P]) 1); by (fast_tac AC_cs 1); val f_subset_imp_HH_subset = result(); val [prem] = goal thy "(!!z. z:Pow(x)-{0} ==> f`z : Pow(z)-{0}) ==> \ -\ x - (UN j: (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0"; +\ x - (UN j: (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0"; by (excluded_middle_tac "?P : {0}" 1); by (fast_tac AC_cs 2); by (dresolve_tac [Diff_subset RS PowI RS DiffI RS prem RS - f_subset_imp_HH_subset] 1); + f_subset_imp_HH_subset] 1); by (fast_tac (AC_cs addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem)] - addSEs [mem_irrefl]) 1); + addSEs [mem_irrefl]) 1); val f_subsets_imp_UN_HH_eq_x = result(); goal thy "HH(f,x,i)=f`(x - (UN j:i. HH(f,x,j))) | HH(f,x,i)={x}"; by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1); by (simp_tac (ZF_ss addsimps [Let_def, Diff_subset RS PowI] - setloop split_tac [expand_if]) 1); + setloop split_tac [expand_if]) 1); val HH_values2 = result(); goal thy "!!f. HH(f,x,i): Pow(x)-{0} ==> HH(f,x,i)=f`(x - (UN j:i. HH(f,x,j)))"; by (resolve_tac [HH_values2 RS disjE] 1 THEN (assume_tac 1)); by (fast_tac (AC_cs addSEs [equalityE, mem_irrefl] - addSDs [singleton_subsetD]) 1); + addSDs [singleton_subsetD]) 1); val HH_subset_imp_eq = result(); goal thy "!!f. [| f : (PROD X:Pow(x)-{0}. {{z}. z:x}); \ -\ a:(LEAST i. HH(f,x,i)={x}) |] ==> EX z:x. HH(f,x,a) = {z}"; +\ a:(LEAST i. HH(f,x,i)={x}) |] ==> EX z:x. HH(f,x,a) = {z}"; by (dtac less_Least_subset_x 1); by (forward_tac [HH_subset_imp_eq] 1); by (dtac apply_type 1); @@ -212,19 +212,19 @@ val f_sing_imp_HH_sing = result(); goalw thy [bij_def] - "!!f. [| x - (UN j: (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0; \ -\ f : (PROD X:Pow(x)-{0}. {{z}. z:x}) |] \ -\ ==> (lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \ -\ : bij(LEAST i. HH(f,x,i)={x}, {{y}. y:x})"; + "!!f. [| x - (UN j: (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0; \ +\ f : (PROD X:Pow(x)-{0}. {{z}. z:x}) |] \ +\ ==> (lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \ +\ : bij(LEAST i. HH(f,x,i)={x}, {{y}. y:x})"; by (fast_tac (AC_cs addSIs [lam_Least_HH_inj, lam_surj_sing, - f_sing_imp_HH_sing]) 1); + f_sing_imp_HH_sing]) 1); val f_sing_lam_bij = result(); goal thy "!!f. f: (PROD X: Pow(x)-{0}. F(X)) \ -\ ==> (lam X:Pow(x)-{0}. {f`X}) : (PROD X: Pow(x)-{0}. {{z}. z:F(X)})"; +\ ==> (lam X:Pow(x)-{0}. {f`X}) : (PROD X: Pow(x)-{0}. {{z}. z:F(X)})"; by (fast_tac (FOL_cs addSIs [lam_type, RepFun_eqI, singleton_eq_iff RS iffD2] - addDs [apply_type]) 1); + addDs [apply_type]) 1); val lam_singI = result(); val bij_Least_HH_x = standard (lam_singI RSN (2, - [f_sing_lam_bij, lam_sing_bij RS bij_converse_bij] MRS comp_bij)); + [f_sing_lam_bij, lam_sing_bij RS bij_converse_bij] MRS comp_bij)); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/Hartog.ML --- a/src/ZF/AC/Hartog.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/Hartog.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/Hartog.ML +(* Title: ZF/AC/Hartog.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski Some proofs on the Hartogs function. *) @@ -13,23 +13,23 @@ qed "Ords_in_set"; goalw thy [lepoll_def] "!!X. [| Ord(a); a lepoll X |] ==> \ -\ EX Y. Y<=X & (EX R. well_ord(Y,R) & ordertype(Y,R)=a)"; +\ EX Y. Y<=X & (EX R. well_ord(Y,R) & ordertype(Y,R)=a)"; by (etac exE 1); by (REPEAT (resolve_tac [exI, conjI] 1)); by (eresolve_tac [inj_is_fun RS fun_is_rel RS image_subset] 1); by (rtac exI 1); by (rtac conjI 1); by (eresolve_tac [well_ord_Memrel RSN (2, subset_refl RSN (2, - restrict_bij RS bij_converse_bij) RS bij_is_inj RS well_ord_rvimage)] 1 - THEN (assume_tac 1)); + restrict_bij RS bij_converse_bij) RS bij_is_inj RS well_ord_rvimage)] 1 + THEN (assume_tac 1)); by (resolve_tac [subset_refl RSN (2, restrict_bij RS bij_converse_bij) RS - (well_ord_Memrel RSN (2, bij_ordertype_vimage)) RS - (ordertype_Memrel RSN (2, trans))] 1 - THEN (REPEAT (assume_tac 1))); + (well_ord_Memrel RSN (2, bij_ordertype_vimage)) RS + (ordertype_Memrel RSN (2, trans))] 1 + THEN (REPEAT (assume_tac 1))); qed "Ord_lepoll_imp_ex_well_ord"; goal thy "!!X. [| Ord(a); a lepoll X |] ==> \ -\ EX Y. Y<=X & (EX R. R<=X*X & ordertype(Y,R)=a)"; +\ EX Y. Y<=X & (EX R. R<=X*X & ordertype(Y,R)=a)"; by (dtac Ord_lepoll_imp_ex_well_ord 1 THEN (assume_tac 1)); by (step_tac ZF_cs 1); by (REPEAT (ares_tac [exI, conjI] 1)); @@ -38,8 +38,8 @@ qed "Ord_lepoll_imp_eq_ordertype"; goal thy "!!X. ALL a. Ord(a) --> a lepoll X ==> \ -\ ALL a. Ord(a) --> \ -\ a:{a. Z:Pow(X)*Pow(X*X), EX Y R. Z= & ordertype(Y,R)=a}"; +\ ALL a. Ord(a) --> \ +\ a:{a. Z:Pow(X)*Pow(X*X), EX Y R. Z= & ordertype(Y,R)=a}"; by (REPEAT (resolve_tac [allI,impI] 1)); by (REPEAT (eresolve_tac [allE, impE] 1)); by (assume_tac 1); @@ -76,7 +76,7 @@ goal thy "!!i. [| i < Hartog(A); i eqpoll Hartog(A) |] ==> P"; by (fast_tac (ZF_cs addEs [less_HartogE1, eqpoll_sym RS eqpoll_imp_lepoll - RS lepoll_trans RS HartogE]) 1); + RS lepoll_trans RS HartogE]) 1); qed "less_HartogE"; goal thy "Card(Hartog(A))"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/OrdQuant.ML --- a/src/ZF/AC/OrdQuant.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/OrdQuant.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/OrdQuant.thy +(* Title: ZF/AC/OrdQuant.thy ID: $Id$ - Authors: Krzysztof Grabczewski and L C Paulson + Authors: Krzysztof Grabczewski and L C Paulson Quantifiers and union operator for ordinals. *) @@ -91,14 +91,14 @@ simp_tac (ZF_ss addcongs [oex_cong] addsimps (OUN_iff::prems)) 1 ]); val OrdQuant_cs = ZF_cs - addSIs [oallI] - addIs [oexI, OUN_I] - addSEs [oexE, OUN_E] - addEs [rev_oallE]; + addSIs [oallI] + addIs [oexI, OUN_I] + addSEs [oexE, OUN_E] + addEs [rev_oallE]; val Ord_atomize = atomize (("oall", [ospec])::ZF_conn_pairs, - ZF_mem_pairs); + ZF_mem_pairs); val OrdQuant_ss = ZF_ss setmksimps (map mk_meta_eq o Ord_atomize o gen_all) - addsimps [oall_simp, ltD RS beta] + addsimps [oall_simp, ltD RS beta] addcongs [oall_cong, oex_cong, OUN_cong]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/ROOT.ML --- a/src/ZF/AC/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,12 +1,12 @@ -(* Title: ZF/AC/ROOT +(* Title: ZF/AC/ROOT ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge Executes the proofs of the AC-equivalences, due to Krzysztof Grabczewski *) -ZF_build_completed; (*Make examples fail if ZF did*) +ZF_build_completed; (*Make examples fail if ZF did*) writeln"Root file for ZF/AC"; proof_timing := true; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/Transrec2.ML --- a/src/ZF/AC/Transrec2.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/Transrec2.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/Transrec2.ML +(* Title: ZF/AC/Transrec2.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski Transfinite recursion introduced to handle definitions based on the three cases of ordinals. @@ -20,7 +20,7 @@ goal thy "transrec2(succ(i),a,b) = b(i, transrec2(i,a,b))"; by (rtac (transrec2_def RS def_transrec RS trans) 1); by (simp_tac (ZF_ss addsimps [succ_not_0, THE_succ_eq, if_P] - setsolver K (fast_tac FOL_cs)) 1); + setsolver K (fast_tac FOL_cs)) 1); val transrec2_succ = result(); goal thy "!!i. Limit(i) ==> transrec2(i,a,b) = (UN j AC1 and WO1 ==> AC10(n) for n >= 1 @@ -28,7 +28,7 @@ open WO1_AC; (* ********************************************************************** *) -(* WO1 ==> AC1 *) +(* WO1 ==> AC1 *) (* ********************************************************************** *) goalw thy [AC1_def, WO1_def] "!!Z. WO1 ==> AC1"; @@ -36,11 +36,11 @@ qed "WO1_AC1"; (* ********************************************************************** *) -(* WO1 ==> AC10(n) (n >= 1) *) +(* WO1 ==> AC10(n) (n >= 1) *) (* ********************************************************************** *) goalw thy [WO1_def] "!!A. [| WO1; ALL B:A. EX C:D(B). P(C,B) |] \ -\ ==> EX f. ALL B:A. P(f`B,B)"; +\ ==> EX f. ALL B:A. P(f`B,B)"; by (eres_inst_tac [("x","Union({{C:D(B). P(C,B)}. B:A})")] allE 1); by (etac exE 1); by (dtac ex_choice_fun 1); @@ -49,7 +49,7 @@ by (res_inst_tac [("x","lam x:A. f`{C:D(x). P(C,x)}")] exI 1); by (asm_full_simp_tac AC_ss 1); by (fast_tac (AC_cs addSDs [RepFunI RSN (2, apply_type)] - addSEs [CollectD2]) 1); + addSEs [CollectD2]) 1); val lemma1 = result(); goalw thy [WO1_def] "!!A. [| ~Finite(B); WO1 |] ==> |B| + |B| eqpoll B"; @@ -59,18 +59,18 @@ by (fast_tac (AC_cs addSEs [well_ord_cardinal_eqpoll]) 1); by (resolve_tac [cadd_def RS def_imp_eq RS subst] 1); by (resolve_tac [Card_cardinal RSN (2, Inf_Card_is_InfCard) RS - InfCard_cdouble_eq RS ssubst] 1); + InfCard_cdouble_eq RS ssubst] 1); by (rtac eqpoll_refl 2); by (rtac notI 1); by (etac notE 1); by (resolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite] 1 - THEN (assume_tac 2)); + THEN (assume_tac 2)); by (fast_tac (AC_cs addSEs [well_ord_cardinal_eqpoll]) 1); val lemma2_1 = result(); goal thy "!!f. f : bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i:D} : Pow(Pow(B))"; by (fast_tac (AC_cs addSIs [InlI, InrI] - addSEs [RepFunE, bij_is_fun RS apply_type]) 1); + addSEs [RepFunE, bij_is_fun RS apply_type]) 1); val lemma2_2 = result(); goal thy "!!f. [| f:inj(A,B); f`a = f`b; a:A; b:A |] ==> a=b"; @@ -79,40 +79,40 @@ val lemma = result(); goalw thy AC_aux_defs - "!!f. f : bij(D+D, B) ==> \ -\ pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i:D})"; + "!!f. f : bij(D+D, B) ==> \ +\ pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i:D})"; by (fast_tac (AC_cs addSEs [RepFunE, not_emptyE] - addDs [bij_is_inj RS lemma, Inl_iff RS iffD1, - Inr_iff RS iffD1, Inl_Inr_iff RS iffD1 RS FalseE, - Inr_Inl_iff RS iffD1 RS FalseE] - addSIs [InlI, InrI]) 1); + addDs [bij_is_inj RS lemma, Inl_iff RS iffD1, + Inr_iff RS iffD1, Inl_Inr_iff RS iffD1 RS FalseE, + Inr_Inl_iff RS iffD1 RS FalseE] + addSIs [InlI, InrI]) 1); val lemma2_3 = result(); val [major, minor] = goalw thy AC_aux_defs - "[| f : bij(D+D, B); 1 le n |] ==> \ -\ sets_of_size_between({{f`Inl(i), f`Inr(i)}. i:D}, 2, succ(n))"; + "[| f : bij(D+D, B); 1 le n |] ==> \ +\ sets_of_size_between({{f`Inl(i), f`Inr(i)}. i:D}, 2, succ(n))"; by (rewtac succ_def); by (fast_tac (AC_cs addSIs [cons_lepoll_cong, minor, lepoll_refl, InlI, InrI] - addIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans, - le_imp_subset RS subset_imp_lepoll] - addDs [major RS bij_is_inj RS lemma, Inl_Inr_iff RS iffD1 RS FalseE] - addSEs [RepFunE, not_emptyE, mem_irrefl]) 1); + addIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans, + le_imp_subset RS subset_imp_lepoll] + addDs [major RS bij_is_inj RS lemma, Inl_Inr_iff RS iffD1 RS FalseE] + addSEs [RepFunE, not_emptyE, mem_irrefl]) 1); val lemma2_4 = result(); goalw thy [bij_def, surj_def] - "!!f. f : bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i:D})=B"; + "!!f. f : bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i:D})=B"; by (fast_tac (AC_cs addSEs [inj_is_fun RS apply_type, CollectE, sumE] - addSIs [InlI, InrI, equalityI]) 1); + addSIs [InlI, InrI, equalityI]) 1); val lemma2_5 = result(); goal thy "!!A. [| WO1; ~Finite(B); 1 le n |] \ -\ ==> EX C:Pow(Pow(B)). pairwise_disjoint(C) & \ -\ sets_of_size_between(C, 2, succ(n)) & \ -\ Union(C)=B"; +\ ==> EX C:Pow(Pow(B)). pairwise_disjoint(C) & \ +\ sets_of_size_between(C, 2, succ(n)) & \ +\ Union(C)=B"; by (eresolve_tac [lemma2_1 RS (eqpoll_def RS def_imp_iff RS iffD1 RS exE)] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (fast_tac (FOL_cs addSIs [bexI] - addSEs [lemma2_2, lemma2_3, lemma2_4, lemma2_5]) 1); + addSEs [lemma2_2, lemma2_3, lemma2_4, lemma2_5]) 1); val lemma2 = result(); goalw thy AC_defs "!!n. [| WO1; 1 le n |] ==> AC10(n)"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/WO1_WO6.ML --- a/src/ZF/AC/WO1_WO6.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/WO1_WO6.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/WO1-WO6.ML +(* Title: ZF/AC/WO1-WO6.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski Proofs needed to state that formulations WO1,...,WO6 are all equivalent. All but one WO6 ==> WO1 (placed in separate file WO6_WO1.ML) @@ -21,7 +21,7 @@ goalw thy (eqpoll_def::WO_defs) "!!Z. WO3 ==> WO1"; by (fast_tac (ZF_cs addSEs [bij_is_inj RS well_ord_rvimage, - well_ord_Memrel RS well_ord_subset]) 1); + well_ord_Memrel RS well_ord_subset]) 1); qed "WO3_WO1"; (* ********************************************************************** *) @@ -42,7 +42,7 @@ goal thy "!!f. [| f:surj(A,B); Ord(A) |] ==> (UN a WO4(1)"; @@ -53,10 +53,10 @@ by (etac Ord_ordertype 1); by (rtac conjI 1); by (eresolve_tac [ordermap_bij RS bij_converse_bij RS bij_is_fun RS - lam_sets RS domain_of_fun] 1); + lam_sets RS domain_of_fun] 1); by (asm_simp_tac (AC_ss addsimps [singleton_eqpoll_1 RS eqpoll_imp_lepoll, Ord_ordertype RSN (2, ordermap_bij RS bij_converse_bij RS - bij_is_surj RS surj_imp_eq)]) 1); + bij_is_surj RS surj_imp_eq)]) 1); qed "WO1_WO4"; (* ********************************************************************** *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/WO1_WO7.ML --- a/src/ZF/AC/WO1_WO7.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/WO1_WO7.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,46 +1,46 @@ -(* Title: ZF/AC/WO1_WO7.ML +(* Title: ZF/AC/WO1_WO7.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski WO7 <-> LEMMA <-> WO1 (Rubin & Rubin p. 5) LEMMA is the sentence denoted by (**) *) (* ********************************************************************** *) -(* It is easy to see, that WO7 is equivallent to (**) *) +(* It is easy to see, that WO7 is equivallent to (**) *) (* ********************************************************************** *) goalw thy [WO7_def] "WO7 <-> (ALL X. ~Finite(X) --> \ -\ (EX R. well_ord(X,R) & ~well_ord(X,converse(R))))"; +\ (EX R. well_ord(X,R) & ~well_ord(X,converse(R))))"; by (fast_tac (ZF_cs addSEs [Finite_well_ord_converse]) 1); val WO7_iff_LEMMA = result(); (* ********************************************************************** *) -(* It is also easy to show that LEMMA implies WO1. *) +(* It is also easy to show that LEMMA implies WO1. *) (* ********************************************************************** *) goalw thy [WO1_def] "!!Z. ALL X. ~Finite(X) --> \ -\ (EX R. well_ord(X,R) & ~well_ord(X,converse(R))) ==> WO1"; +\ (EX R. well_ord(X,R) & ~well_ord(X,converse(R))) ==> WO1"; by (rtac allI 1); by (etac allE 1); by (excluded_middle_tac "Finite(A)" 1); by (fast_tac AC_cs 1); by (rewrite_goals_tac [Finite_def, eqpoll_def]); by (fast_tac (ZF_cs addSIs [[bij_is_inj, nat_implies_well_ord] MRS - well_ord_rvimage]) 1); + well_ord_rvimage]) 1); val LEMMA_imp_WO1 = result(); (* ********************************************************************** *) -(* The Rubins' proof of the other implication is contained within the *) -(* following sentence : *) -(* "... each infinite ordinal is well ordered by < but not by >." *) -(* This statement can be proved by the following two theorems. *) -(* But moreover we need to show similar property for any well ordered *) +(* The Rubins' proof of the other implication is contained within the *) +(* following sentence : *) +(* "... each infinite ordinal is well ordered by < but not by >." *) +(* This statement can be proved by the following two theorems. *) +(* But moreover we need to show similar property for any well ordered *) (* infinite set. It is not very difficult thanks to Isabelle order types *) -(* We show that if a set is well ordered by some relation and by it's *) +(* We show that if a set is well ordered by some relation and by it's *) (* converse, then apropriate order type is well ordered by the converse *) -(* of it's membership relation, which in connection with the previous *) -(* gives the conclusion. *) +(* of it's membership relation, which in connection with the previous *) +(* gives the conclusion. *) (* ********************************************************************** *) goalw thy [wf_on_def, wf_def] @@ -63,19 +63,19 @@ val converse_Memrel_not_well_ord = result(); goal thy "!!A. [| well_ord(A,r); well_ord(A,converse(r)) |] \ -\ ==> well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r))))"; +\ ==> well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r))))"; by (rtac ([ordertype_ord_iso RS ord_iso_sym RS ord_iso_rvimage_eq, - Memrel_type RS (subset_Int_iff RS iffD1)] - MRS trans RS subst) 1 - THEN (assume_tac 1)); + Memrel_type RS (subset_Int_iff RS iffD1)] + MRS trans RS subst) 1 + THEN (assume_tac 1)); by (rtac (rvimage_converse RS subst) 1); by (etac (ordertype_ord_iso RS ord_iso_sym RS ord_iso_is_bij RS - bij_is_inj RS well_ord_rvimage) 1 - THEN (assume_tac 1)); + bij_is_inj RS well_ord_rvimage) 1 + THEN (assume_tac 1)); val well_ord_converse_Memrel = result(); goalw thy [WO1_def] "!!Z. WO1 ==> ALL X. ~Finite(X) --> \ -\ (EX R. well_ord(X,R) & ~well_ord(X,converse(R)))"; +\ (EX R. well_ord(X,R) & ~well_ord(X,converse(R)))"; by (REPEAT (resolve_tac [allI,impI] 1)); by (REPEAT (eresolve_tac [allE,exE] 1)); by (REPEAT (ares_tac [exI,conjI,notI] 1)); @@ -83,8 +83,8 @@ by (forward_tac [Ord_ordertype RS converse_Memrel_not_well_ord] 1); by (contr_tac 2); by (fast_tac (empty_cs addSEs [ordertype_ord_iso RS ord_iso_is_bij RS - bij_is_inj RS (exI RS (lepoll_def RS def_imp_iff RS iffD2)) - RS lepoll_Finite] - addSIs [notI] addEs [notE]) 1); + bij_is_inj RS (exI RS (lepoll_def RS def_imp_iff RS iffD2)) + RS lepoll_Finite] + addSIs [notI] addEs [notE]) 1); val WO1_imp_LEMMA = result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/WO1_WO8.ML --- a/src/ZF/AC/WO1_WO8.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/WO1_WO8.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,12 +1,12 @@ -(* Title: ZF/AC/WO1_WO8.ML +(* Title: ZF/AC/WO1_WO8.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proof of WO8 <-> WO1 (Rubin & Rubin p. 6) *) (* ********************************************************************** *) -(* Trivial implication "WO1 ==> WO8" *) +(* Trivial implication "WO1 ==> WO8" *) (* ********************************************************************** *) goalw thy WO_defs "!!Z. WO1 ==> WO8"; @@ -14,7 +14,7 @@ qed "WO1_WO8"; (* ********************************************************************** *) -(* The proof of "WO8 ==> WO1" - faithful image of Rubin & Rubin's proof *) +(* The proof of "WO8 ==> WO1" - faithful image of Rubin & Rubin's proof *) (* ********************************************************************** *) goalw thy WO_defs "!!Z. WO8 ==> WO1"; @@ -22,9 +22,9 @@ by (eres_inst_tac [("x","{{x}. x:A}")] allE 1); by (etac impE 1); by (fast_tac (AC_cs addSEs [lam_sing_bij RS bij_is_inj RS - well_ord_rvimage]) 2); + well_ord_rvimage]) 2); by (res_inst_tac [("x","lam a:{{x}. x:A}. THE x. a={x}")] exI 1); by (fast_tac (ZF_cs addSEs [RepFunE, singleton_eq_iff RS iffD1 RS sym] - addSIs [lam_type, singletonI] - addIs [the_equality RS ssubst]) 1); + addSIs [lam_type, singletonI] + addIs [the_equality RS ssubst]) 1); qed "WO8_WO1"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/WO2_AC16.ML --- a/src/ZF/AC/WO2_AC16.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/WO2_AC16.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/WO2_AC16.ML +(* Title: ZF/AC/WO2_AC16.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski The proof of WO2 ==> AC16(k #+ m, k) @@ -15,14 +15,14 @@ *) (* ********************************************************************** *) -(* case of limit ordinal *) +(* case of limit ordinal *) (* ********************************************************************** *) goal thy "!!Z. [| ALL y (EX! Y. Y:F(y) & f(z)<=Y); \ -\ ALL i j. i le j --> F(i) <= F(j); j le i; i V = W"; +\ --> (EX! Y. Y:F(y) & f(z)<=Y); \ +\ ALL i j. i le j --> F(i) <= F(j); j le i; i V = W"; by (REPEAT (eresolve_tac [asm_rl, allE, impE] 1)); by (dtac subsetD 1 THEN (assume_tac 1)); by (REPEAT (dtac ospec 1 THEN (assume_tac 1))); @@ -32,21 +32,21 @@ val lemma3_1 = result(); goal thy "!!Z. [| ALL y (EX! Y. Y:F(y) & f(z)<=Y); \ -\ ALL i j. i le j --> F(i) <= F(j); i V = W"; +\ --> (EX! Y. Y:F(y) & f(z)<=Y); \ +\ ALL i j. i le j --> F(i) <= F(j); i V = W"; by (res_inst_tac [("j","j")] (lt_Ord RS (lt_Ord RSN (2, Ord_linear_le))) 1 - THEN (REPEAT (assume_tac 1))); + THEN (REPEAT (assume_tac 1))); by (eresolve_tac [lemma3_1 RS sym] 1 THEN (REPEAT (assume_tac 1))); by (etac lemma3_1 1 THEN (REPEAT (assume_tac 1))); val lemma3 = result(); goal thy "!!a. [| ALL y:x. y < a --> F(y) <= X & \ -\ (ALL x \ -\ (EX! Y. Y : F(y) & fa(x) <= Y)); x < a |] \ -\ ==> ALL y \ -\ (EX! Y. Y : F(y) & fa(z) <= Y)"; +\ (ALL x \ +\ (EX! Y. Y : F(y) & fa(x) <= Y)); x < a |] \ +\ ==> ALL y \ +\ (EX! Y. Y : F(y) & fa(z) <= Y)"; by (REPEAT (resolve_tac [oallI, impI] 1)); by (dresolve_tac [ltD RSN (2, bspec)] 1 THEN (assume_tac 1)); by (eresolve_tac [lt_trans RSN (2, impE)] 1 THEN (REPEAT (assume_tac 1))); @@ -54,12 +54,12 @@ val lemma4 = result(); goal thy "!!a. [| ALL y:x. y < a --> F(y) <= X & \ -\ (ALL x \ -\ (EX! Y. Y : F(y) & fa(x) <= Y)); \ -\ x < a; Limit(x); ALL i j. i le j --> F(i) <= F(j) |] \ -\ ==> (UN x (EX! Y. Y : (UN x \ +\ (EX! Y. Y : F(y) & fa(x) <= Y)); \ +\ x < a; Limit(x); ALL i j. i le j --> F(i) <= F(j) |] \ +\ ==> (UN x (EX! Y. Y : (UN x A - B - C eqpoll a"; +\ C lesspoll a |] ==> A - B - C eqpoll a"; by (rtac Diff_lesspoll_eqpoll_Card 1 THEN (REPEAT (assume_tac 1))); by (rtac Diff_lesspoll_eqpoll_Card 1 THEN (REPEAT (assume_tac 1))); val dbl_Diff_eqpoll_Card = result(); (* ********************************************************************** *) -(* Case of finite ordinals *) +(* Case of finite ordinals *) (* ********************************************************************** *) goalw thy [lesspoll_def] - "!!X. [| Finite(X); ~Finite(a); Ord(a) |] ==> X lesspoll a"; + "!!X. [| Finite(X); ~Finite(a); Ord(a) |] ==> X lesspoll a"; by (rtac conjI 1); by (dresolve_tac [nat_le_infinite_Ord RS le_imp_lepoll] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (rewtac Finite_def); by (fast_tac (AC_cs addSEs [eqpoll_sym RS eqpoll_trans]) 2); by (rtac lepoll_trans 1 THEN (assume_tac 2)); by (fast_tac (AC_cs addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_subset RS - subset_imp_lepoll RSN (2, eqpoll_imp_lepoll RS lepoll_trans)]) 1); + subset_imp_lepoll RSN (2, eqpoll_imp_lepoll RS lepoll_trans)]) 1); val Finite_lesspoll_infinite_Ord = result(); goal thy "!!x. x:X ==> Union(X) = Union(X-{x}) Un x"; by (fast_tac (AC_cs addIs [equalityI] addSIs [singletonI] - addSEs [singletonE]) 1); + addSEs [singletonE]) 1); val Union_eq_Un_Diff = result(); goal thy "!!n. n:nat ==> ALL X. X eqpoll n --> (ALL x:X. Finite(x)) \ -\ --> Finite(Union(X))"; +\ --> Finite(Union(X))"; by (etac nat_induct 1); by (fast_tac (AC_cs addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0] - addSIs [nat_0I RS nat_into_Finite] addIs [Union_0 RS ssubst]) 1); + addSIs [nat_0I RS nat_into_Finite] addIs [Union_0 RS ssubst]) 1); by (REPEAT (resolve_tac [allI, impI] 1)); by (resolve_tac [eqpoll_succ_imp_not_empty RS not_emptyE] 1 THEN (assume_tac 1)); by (res_inst_tac [("P","%z. Finite(z)")] (Union_eq_Un_Diff RS ssubst) 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (rtac Finite_Un 1); by (fast_tac AC_cs 2); by (fast_tac (AC_cs addSIs [Diff_sing_eqpoll]) 1); @@ -152,41 +152,41 @@ goalw thy [Finite_def] "!!x. [| x lepoll n; n:nat |] ==> Finite(x)"; by (fast_tac (AC_cs - addEs [nat_into_Ord RSN (2, lepoll_imp_ex_le_eqpoll) RS exE, - Ord_nat RSN (2, ltI) RSN (2, lt_trans1) RS ltD]) 1); + addEs [nat_into_Ord RSN (2, lepoll_imp_ex_le_eqpoll) RS exE, + Ord_nat RSN (2, ltI) RSN (2, lt_trans1) RS ltD]) 1); val lepoll_nat_num_imp_Finite = result(); goal thy "!!X. [| ALL x:X. x lepoll n & x <= T; well_ord(T, R); X lepoll b; \ -\ b Union(X) lesspoll a"; +\ b Union(X) lesspoll a"; by (excluded_middle_tac "Finite(X)" 1); by (resolve_tac [Card_is_Ord RSN (3, Finite_lesspoll_infinite_Ord)] 2 - THEN (REPEAT (assume_tac 3))); + THEN (REPEAT (assume_tac 3))); by (fast_tac (AC_cs addSEs [lepoll_nat_num_imp_Finite] - addSIs [Finite_Union]) 2); + addSIs [Finite_Union]) 2); by (dresolve_tac [lt_Ord RSN (2, lepoll_imp_ex_le_eqpoll)] 1 THEN (assume_tac 1)); by (REPEAT (eresolve_tac [exE, conjE] 1)); by (forward_tac [eqpoll_imp_lepoll RS lepoll_infinite] 1 THEN (assume_tac 1)); by (eresolve_tac [eqpoll_sym RS (eqpoll_def RS def_imp_iff RS iffD1) RS - exE] 1); + exE] 1); by (forward_tac [bij_is_surj RS surj_image_eq] 1); by (dresolve_tac [[bij_is_fun, subset_refl] MRS image_fun] 1); by (dresolve_tac [sym RS trans] 1 THEN (assume_tac 1)); by (hyp_subst_tac 1); by (rtac lepoll_lesspoll_lesspoll 1); by (eresolve_tac [lt_trans1 RSN (2, lt_Card_imp_lesspoll)] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (rtac UN_lepoll 1 - THEN (TRYALL (fast_tac (AC_cs addSEs [lt_Ord])))); + THEN (TRYALL (fast_tac (AC_cs addSEs [lt_Ord])))); val Union_lesspoll = result(); (* ********************************************************************** *) -(* recfunAC16_lepoll_index *) +(* recfunAC16_lepoll_index *) (* ********************************************************************** *) goal thy "A Un {a} = cons(a, A)"; by (fast_tac (AC_cs addSIs [singletonI] - addSEs [singletonE] addIs [equalityI]) 1); + addSEs [singletonE] addIs [equalityI]) 1); val Un_sing_eq_cons = result(); goal thy "!!A. A lepoll B ==> A Un {a} lepoll succ(B)"; @@ -203,28 +203,28 @@ val Diff_UN_succ_subset = result(); goal thy "!!x. Ord(x) ==> \ -\ recfunAC16(f, g, x, a) - (UN i z:F(LEAST i. z:F(i)) - (UN j<(LEAST i. z:F(i)). F(j))"; +\ ==> z:F(LEAST i. z:F(i)) - (UN j<(LEAST i. z:F(i)). F(j))"; by (fast_tac (AC_cs addEs [less_LeastE] addSEs [OUN_E, LeastI]) 1); val in_Least_Diff = result(); goal thy "!!w. [| (LEAST i. w:F(i)) = (LEAST i. z:F(i)); \ -\ w:(UN i EX b EX b a=b"; @@ -240,7 +240,7 @@ val two_in_lepoll_1 = result(); goal thy "!!a. [| ALL i (UN x (UN x Union(recfunAC16(f,g,y,a)) lesspoll a"; +\ A eqpoll a; y Union(recfunAC16(f,g,y,a)) lesspoll a"; by (eresolve_tac [eqpoll_def RS def_imp_iff RS iffD1 RS exE] 1); by (rtac Union_lesspoll 1 THEN (TRYALL assume_tac)); by (eresolve_tac [lt_Ord RS recfunAC16_lepoll_index] 3); by (eresolve_tac [[bij_is_inj, Card_is_Ord RS well_ord_Memrel] MRS - well_ord_rvimage] 2 THEN (assume_tac 2)); + well_ord_rvimage] 2 THEN (assume_tac 2)); by (fast_tac (AC_cs addSEs [eqpoll_imp_lepoll]) 1); val Union_recfunAC16_lesspoll = result(); goal thy "!!a. [| recfunAC16(f, fa, y, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)}; \ -\ Card(a); ~ Finite(a); A eqpoll a; \ -\ k : nat; m : nat; y A - Union(recfunAC16(f, fa, y, a)) - fa`y eqpoll a"; +\ Card(a); ~ Finite(a); A eqpoll a; \ +\ k : nat; m : nat; y A - Union(recfunAC16(f, fa, y, a)) - fa`y eqpoll a"; by (rtac dbl_Diff_eqpoll_Card 1 THEN (TRYALL assume_tac)); by (rtac Union_recfunAC16_lesspoll 1 THEN (REPEAT (assume_tac 1))); by (eresolve_tac [add_type RS nat_succI] 1 THEN (assume_tac 1)); by (resolve_tac [nat_succI RSN (2, bexI RS (Finite_def RS def_imp_iff RS - iffD2)) RS (Card_is_Ord RSN (3, Finite_lesspoll_infinite_Ord))] 1 - THEN (TRYALL assume_tac)); + iffD2)) RS (Card_is_Ord RSN (3, Finite_lesspoll_infinite_Ord))] 1 + THEN (TRYALL assume_tac)); by (eresolve_tac [ltD RSN (2, bij_is_fun RS apply_type) RS CollectE] 1 - THEN (TRYALL assume_tac)); + THEN (TRYALL assume_tac)); val dbl_Diff_eqpoll = result(); (* back to the proof *) @@ -307,49 +307,49 @@ eqpoll_trans RS eqpoll_trans))) |> standard; goal thy "!!x. [| x : Pow(A - B - fa`i); x eqpoll m; \ -\ fa : bij(a, {x: Pow(A) . x eqpoll k}); i fa ` i Un x : {x: Pow(A) . x eqpoll k #+ m}"; +\ fa : bij(a, {x: Pow(A) . x eqpoll k}); i fa ` i Un x : {x: Pow(A) . x eqpoll k #+ m}"; by (rtac CollectI 1); by (fast_tac (AC_cs addSIs [PowD RS (PowD RSN (2, Un_least RS PowI))] - addSEs [ltD RSN (2, bij_is_fun RS apply_type RS CollectE)]) 1); + addSEs [ltD RSN (2, bij_is_fun RS apply_type RS CollectE)]) 1); by (rtac disj_Un_eqpoll_nat_sum 1 - THEN (TRYALL assume_tac)); + THEN (TRYALL assume_tac)); by (fast_tac (AC_cs addSIs [equals0I]) 1); by (eresolve_tac [ltD RSN (2, bij_is_fun RS apply_type RS CollectE)] 1 - THEN (REPEAT (assume_tac 1))); + THEN (REPEAT (assume_tac 1))); val Un_in_Collect = result(); (* ********************************************************************** *) -(* Lemmas simplifying assumptions *) +(* Lemmas simplifying assumptions *) (* ********************************************************************** *) goal thy "!!j. [| ALL y:succ(j). y F(y)<=X & (ALL x Q(x,y)); succ(j) F(j)<=X & (ALL x Q(x,j))"; +\ --> Q(x,y)); succ(j) F(j)<=X & (ALL x Q(x,j))"; by (dresolve_tac [succI1 RSN (2, bspec)] 1); by (etac impE 1); by (resolve_tac [lt_Ord RS (succI1 RS ltI RS lt_Ord RS le_refl) RS lt_trans] 1 - THEN (REPEAT (assume_tac 1))); + THEN (REPEAT (assume_tac 1))); val lemma6 = result(); goal thy "!!j. [| F(j)<=X; (ALL x Q(x,j)); succ(j) P(j,j) --> F(j) <= X & (ALL x Q(x,j))"; +\ ==> P(j,j) --> F(j) <= X & (ALL x Q(x,j))"; by (fast_tac (AC_cs addSEs [leE]) 1); val lemma7 = result(); (* ********************************************************************** *) (* Lemmas needded to prove ex_next_set which means that for any successor *) -(* ordinal there is a set satisfying certain properties *) +(* ordinal there is a set satisfying certain properties *) (* ********************************************************************** *) goal thy "!!A. [| A eqpoll a; ~ Finite(a); Ord(a); m:nat |] \ -\ ==> EX X:Pow(A). X eqpoll m"; +\ ==> EX X:Pow(A). X eqpoll m"; by (eresolve_tac [Ord_nat RSN (2, ltI) RS - (nat_le_infinite_Ord RSN (2, lt_trans2)) RS - leI RS le_imp_lepoll RS - ((eqpoll_sym RS eqpoll_imp_lepoll) RSN (2, lepoll_trans)) RS - lepoll_imp_eqpoll_subset RS exE] 1 - THEN REPEAT (assume_tac 1)); + (nat_le_infinite_Ord RSN (2, lt_trans2)) RS + leI RS le_imp_lepoll RS + ((eqpoll_sym RS eqpoll_imp_lepoll) RSN (2, lepoll_trans)) RS + lepoll_imp_eqpoll_subset RS exE] 1 + THEN REPEAT (assume_tac 1)); by (fast_tac (AC_cs addSEs [eqpoll_sym]) 1); val ex_subset_eqpoll = result(); @@ -362,12 +362,12 @@ val Int_empty = result(); (* ********************************************************************** *) -(* equipollent subset (and finite) is the whole set *) +(* equipollent subset (and finite) is the whole set *) (* ********************************************************************** *) goal thy "!!A. [| A <= B; a : A; A - {a} = B - {a} |] ==> A = B"; by (fast_tac (AC_cs addSEs [equalityE, singletonE] - addSIs [equalityI, singletonI]) 1); + addSIs [equalityI, singletonI]) 1); val Diffs_eq_imp_eq = result(); goal thy "!!A. m:nat ==> ALL A B. A <= B & m lepoll A & B lepoll m --> A=B"; @@ -376,16 +376,16 @@ by (REPEAT (resolve_tac [allI, impI] 1)); by (REPEAT (etac conjE 1)); by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1 - THEN (assume_tac 1)); + THEN (assume_tac 1)); by (forward_tac [subsetD RS Diff_sing_lepoll] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (forward_tac [lepoll_Diff_sing] 1); by (REPEAT (eresolve_tac [allE, impE] 1)); by (rtac conjI 1); by (fast_tac AC_cs 2); by (fast_tac (AC_cs addSEs [singletonE] addSIs [singletonI]) 1); by (etac Diffs_eq_imp_eq 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); val subset_imp_eq_lemma = result(); goal thy "!!A. [| A <= B; m lepoll A; B lepoll m; m:nat |] ==> A=B"; @@ -393,28 +393,28 @@ val subset_imp_eq = result(); goal thy "!!f. [| f:bij(a, {Y:X. Y eqpoll succ(k)}); k:nat; f`b<=f`y; b b=y"; +\ y b=y"; by (dtac subset_imp_eq 1); by (etac nat_succI 3); by (fast_tac (AC_cs addSEs [bij_is_fun RS (ltD RSN (2, apply_type)) RS - CollectE, eqpoll_sym RS eqpoll_imp_lepoll]) 1); + CollectE, eqpoll_sym RS eqpoll_imp_lepoll]) 1); by (fast_tac (AC_cs addSEs [bij_is_fun RS (ltD RSN (2, apply_type)) RS - CollectE, eqpoll_imp_lepoll]) 1); + CollectE, eqpoll_imp_lepoll]) 1); by (rewrite_goals_tac [bij_def, inj_def]); by (fast_tac (AC_cs addSDs [ltD]) 1); val bij_imp_arg_eq = result(); goal thy "!!a. [| recfunAC16(f, fa, y, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)}; \ -\ Card(a); ~ Finite(a); A eqpoll a; \ -\ k : nat; m : nat; y EX X:{Y: Pow(A). Y eqpoll succ(k #+ m)}. fa`y <= X & \ -\ (ALL b \ -\ (ALL T:recfunAC16(f, fa, y, a). ~ fa`b <= T))"; +\ Card(a); ~ Finite(a); A eqpoll a; \ +\ k : nat; m : nat; y EX X:{Y: Pow(A). Y eqpoll succ(k #+ m)}. fa`y <= X & \ +\ (ALL b \ +\ (ALL T:recfunAC16(f, fa, y, a). ~ fa`b <= T))"; by (eresolve_tac [dbl_Diff_eqpoll RS ex_subset_eqpoll RS bexE] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (etac Card_is_Ord 1); by (forward_tac [Un_in_Collect] 2 THEN REPEAT (assume_tac 2)); by (etac CollectE 4); @@ -429,33 +429,33 @@ by (fast_tac AC_cs 1); by (REPEAT (resolve_tac [ballI, impI, oallI, notI] 1)); by (dresolve_tac [Int_empty RSN (2, subset_Un_disjoint)] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (dtac bij_imp_arg_eq 1 THEN REPEAT (assume_tac 1)); by (hyp_subst_tac 1); by (eresolve_tac [bexI RSN (2, notE)] 1 THEN TRYALL assume_tac); val ex_next_set = result(); (* ********************************************************************** *) -(* Lemma ex_next_Ord states that for any successor *) -(* ordinal there is a number of the set satisfying certain properties *) +(* Lemma ex_next_Ord states that for any successor *) +(* ordinal there is a number of the set satisfying certain properties *) (* ********************************************************************** *) goal thy "!!a. [| recfunAC16(f, fa, y, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)}; \ -\ Card(a); ~ Finite(a); A eqpoll a; \ -\ k : nat; m : nat; y EX c \ -\ (ALL T:recfunAC16(f, fa, y, a). ~ fa`b <= T))"; +\ Card(a); ~ Finite(a); A eqpoll a; \ +\ k : nat; m : nat; y EX c \ +\ (ALL T:recfunAC16(f, fa, y, a). ~ fa`b <= T))"; by (dtac ex_next_set 1 THEN REPEAT (assume_tac 1)); by (etac bexE 1); by (resolve_tac [bij_converse_bij RS bij_is_fun RS apply_type RS ltI RSN - (2, oexI)] 1); + (2, oexI)] 1); by (resolve_tac [right_inverse_bij RS ssubst] 1 - THEN REPEAT (ares_tac [Card_is_Ord] 1)); + THEN REPEAT (ares_tac [Card_is_Ord] 1)); val ex_next_Ord = result(); goal thy "!!Z. [| EX! Y. Y:Z & P(Y); ~P(W) |] ==> EX! Y. Y:Z Un {W} & P(Y)"; @@ -463,15 +463,15 @@ val ex1_in_Un_sing = result(); (* ********************************************************************** *) -(* Lemma simplifying assumptions *) +(* Lemma simplifying assumptions *) (* ********************************************************************** *) goal thy "!!j. [| ALL x (EX! Y. Y:F(j) & P(x, Y)); F(j) <= X; \ -\ L : X; P(j, L) & (ALL x (ALL xa:F(j). ~P(x, xa))) |] \ -\ ==> F(j) Un {L} <= X & \ -\ (ALL x \ -\ (EX! Y. Y:F(j) Un {L} & P(x, Y)))"; +\ --> (EX! Y. Y:F(j) & P(x, Y)); F(j) <= X; \ +\ L : X; P(j, L) & (ALL x (ALL xa:F(j). ~P(x, xa))) |] \ +\ ==> F(j) Un {L} <= X & \ +\ (ALL x \ +\ (EX! Y. Y:F(j) Un {L} & P(x, Y)))"; by (rtac conjI 1); by (fast_tac (AC_cs addSIs [singleton_subsetI]) 1); by (rtac oallI 1); @@ -492,16 +492,16 @@ (* ********************************************************************** *) (* The main part of the proof: inductive proof of the property of T_gamma *) -(* lemma main_induct *) +(* lemma main_induct *) (* ********************************************************************** *) goal thy - "!!a. [| b < a; f : bij(a, {Y: Pow(A) . Y eqpoll succ(k #+ m)}); \ -\ fa : bij(a, {Y: Pow(A) . Y eqpoll succ(k)}); \ -\ ~Finite(a); Card(a); A eqpoll a; k : nat; m : nat |] \ -\ ==> recfunAC16(f, fa, b, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)} & \ -\ (ALL x \ -\ (EX! Y. Y:recfunAC16(f, fa, b, a) & fa ` x <= Y))"; + "!!a. [| b < a; f : bij(a, {Y: Pow(A) . Y eqpoll succ(k #+ m)}); \ +\ fa : bij(a, {Y: Pow(A) . Y eqpoll succ(k)}); \ +\ ~Finite(a); Card(a); A eqpoll a; k : nat; m : nat |] \ +\ ==> recfunAC16(f, fa, b, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)} & \ +\ (ALL x \ +\ (EX! Y. Y:recfunAC16(f, fa, b, a) & fa ` x <= Y))"; by (rtac impE 1 THEN (REPEAT (assume_tac 2))); by (eresolve_tac [lt_Ord RS trans_induct] 1); by (rtac impI 1); @@ -521,25 +521,25 @@ by (etac lemma7 1 THEN (REPEAT (assume_tac 1))); by (rtac impI 1); by (resolve_tac [ex_next_Ord RS oexE] 1 - THEN REPEAT (ares_tac [le_refl RS lt_trans] 1)); + THEN REPEAT (ares_tac [le_refl RS lt_trans] 1)); by (etac lemma8 1 THEN (assume_tac 1)); by (resolve_tac [bij_is_fun RS apply_type] 1 THEN (assume_tac 1)); by (eresolve_tac [Least_le RS lt_trans2 RS ltD] 1 - THEN REPEAT (ares_tac [lt_Ord, succ_leI] 1)); + THEN REPEAT (ares_tac [lt_Ord, succ_leI] 1)); by (eresolve_tac [lt_Ord RSN (2, LeastI)] 1 THEN (assume_tac 1)); val main_induct = result(); (* ********************************************************************** *) -(* Lemma to simplify the inductive proof *) +(* Lemma to simplify the inductive proof *) (* - the desired property is a consequence of the inductive assumption *) (* ********************************************************************** *) val [prem1, prem2, prem3, prem4] = goal thy - "[| (!!b. b F(b) <= S & (ALL x (EX! Y. Y : F(b) & f`x <= Y))); \ -\ f:a->f``(a); Limit(a); (!!i j. i le j ==> F(i) <= F(j)) |] \ -\ ==> (UN j F(b) <= S & (ALL x (EX! Y. Y : F(b) & f`x <= Y))); \ +\ f:a->f``(a); Limit(a); (!!i j. i le j ==> F(i) <= F(j)) |] \ +\ ==> (UN j AC16(k #+ m,k)"; + "!!n k. [| WO2; 0 AC16(k #+ m,k)"; by (rtac allI 1); by (rtac impI 1); by (forward_tac [WO2_infinite_subsets_eqpoll_X] 1 THEN (REPEAT (assume_tac 1))); by (forw_inst_tac [("n","k #+ m")] (WO2_infinite_subsets_eqpoll_X) 1 - THEN (REPEAT (ares_tac [add_type] 1))); + THEN (REPEAT (ares_tac [add_type] 1))); by (forward_tac [WO2_imp_ex_Card] 1); by (REPEAT (eresolve_tac [exE,conjE] 1)); by (dresolve_tac [eqpoll_trans RS eqpoll_sym RS (eqpoll_def RS - def_imp_iff RS iffD1)] 1 THEN (assume_tac 1)); + def_imp_iff RS iffD1)] 1 THEN (assume_tac 1)); by (dresolve_tac [eqpoll_trans RS eqpoll_sym RS (eqpoll_def RS - def_imp_iff RS iffD1)] 1 THEN (assume_tac 1)); + def_imp_iff RS iffD1)] 1 THEN (assume_tac 1)); by (REPEAT (etac exE 1)); by (res_inst_tac [("x","UN j WO1". Simplified by L C Paulson. @@ -25,11 +25,11 @@ (lt_oadd_odiff_disj RS disjE); (* ********************************************************************** *) -(* The most complicated part of the proof - lemma ii - p. 2-4 *) +(* The most complicated part of the proof - lemma ii - p. 2-4 *) (* ********************************************************************** *) (* ********************************************************************** *) -(* some properties of relation uu(beta, gamma, delta) - p. 2 *) +(* some properties of relation uu(beta, gamma, delta) - p. 2 *) (* ********************************************************************** *) goalw thy [uu_def] "domain(uu(f,b,g,d)) <= f`b"; @@ -39,7 +39,7 @@ goal thy "!! a. ALL b \ \ ALL b uu(f,b,g,d) lepoll m"; by (fast_tac (AC_cs - addSEs [uu_subset2 RS subset_imp_lepoll RS lepoll_trans]) 1); + addSEs [uu_subset2 RS subset_imp_lepoll RS lepoll_trans]) 1); val uu_lepoll_m = result(); (* ********************************************************************** *) -(* Two cases for lemma ii *) +(* Two cases for lemma ii *) (* ********************************************************************** *) goalw thy [lesspoll_def] "!! a f u. ALL b \ \ (ALL b (EX g \ -\ u(f,b,g,d) eqpoll m))"; +\ u(f,b,g,d) eqpoll m))"; by (asm_simp_tac OrdQuant_ss 1); by (fast_tac AC_cs 1); val cases = result(); (* ********************************************************************** *) -(* Lemmas used in both cases *) +(* Lemmas used in both cases *) (* ********************************************************************** *) goal thy "!!a C. Ord(a) ==> (UN b \ -\ (UN b \ +\ (UN b P(Least_a, LEAST b. P(Least_a, b))"; +\ Least_a = (LEAST a. EX x. Ord(x) & P(a, x)) |] \ +\ ==> P(Least_a, LEAST b. P(Least_a, b))"; by (etac ssubst 1); by (res_inst_tac [("Q","%z. P(z, LEAST b. P(z, b))")] LeastI2 1); by (REPEAT (fast_tac (ZF_cs addSEs [LeastI]) 1)); @@ -119,24 +119,24 @@ val nested_Least_instance = standard (read_instantiate - [("P","%g d. domain(uu(f,b,g,d)) ~= 0 & \ -\ domain(uu(f,b,g,d)) lepoll m")] nested_LeastI); + [("P","%g d. domain(uu(f,b,g,d)) ~= 0 & \ +\ domain(uu(f,b,g,d)) lepoll m")] nested_LeastI); goalw thy [gg1_def] "!!a. [| Ord(a); m:nat; \ -\ ALL b \ -\ (EX g gg1(f,a,m)`b lepoll m"; +\ ALL b \ +\ (EX g gg1(f,a,m)`b lepoll m"; by (asm_simp_tac OrdQuant_ss 1); by (safe_tac (OrdQuant_cs addSEs [lt_oadd_odiff_cases])); (*Case b EX d EX d uu(f,b,g,LEAST d. (uu(f,b,g,d) ~= 0)) ~= 0"; +\ y*y<=y; (UN b uu(f,b,g,LEAST d. (uu(f,b,g,d) ~= 0)) ~= 0"; by (dtac ex_d_uu_not_empty 1 THEN REPEAT (assume_tac 1)); by (fast_tac (AC_cs addSEs [LeastI, lt_Ord]) 1); val uu_not_empty = result(); goal ZF.thy "!!r. [| r<=A*B; r~=0 |] ==> domain(r)~=0"; by (REPEAT (eresolve_tac [asm_rl, not_emptyE, subsetD RS SigmaE, - sym RSN (2, subst_elem) RS domainI RS not_emptyI] 1)); + sym RSN (2, subst_elem) RS domainI RS not_emptyI] 1)); val not_empty_rel_imp_domain = result(); goal thy "!!f. [| b (LEAST d. uu(f,b,g,d) ~= 0) < a"; +\ y*y <= y; (UN b (LEAST d. uu(f,b,g,d) ~= 0) < a"; by (eresolve_tac [ex_d_uu_not_empty RS oexE] 1 - THEN REPEAT (assume_tac 1)); + THEN REPEAT (assume_tac 1)); by (resolve_tac [Least_le RS lt_trans1] 1 - THEN (REPEAT (ares_tac [lt_Ord] 1))); + THEN (REPEAT (ares_tac [lt_Ord] 1))); val Least_uu_not_empty_lt_a = result(); goal ZF.thy "!!B. [| B<=A; a~:B |] ==> B <= A-{a}"; @@ -202,56 +202,56 @@ by (rtac subsetI 1); by (excluded_middle_tac "?P" 1 THEN (assume_tac 2)); by (eresolve_tac [subset_Diff_sing RS subset_imp_lepoll RSN (2, - Diff_sing_lepoll RSN (3, lepoll_trans RS lepoll_trans)) RS - succ_lepoll_natE] 1 - THEN REPEAT (assume_tac 1)); + Diff_sing_lepoll RSN (3, lepoll_trans RS lepoll_trans)) RS + succ_lepoll_natE] 1 + THEN REPEAT (assume_tac 1)); val supset_lepoll_imp_eq = result(); goal thy - "!!a. [| ALL g \ -\ domain(uu(f, b, g, d)) eqpoll succ(m); \ -\ ALL b \ +\ domain(uu(f, b, g, d)) eqpoll succ(m); \ +\ ALL b uu(f, b, g, LEAST d. uu(f,b,g,d)~=0) : f`b -> f`g"; by (dres_inst_tac [("x2","g")] (ospec RS ospec RS mp) 1); by (rtac ([uu_subset1, uu_not_empty] MRS not_empty_rel_imp_domain) 3); by (rtac Least_uu_not_empty_lt_a 2 THEN TRYALL assume_tac); by (resolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS - (Least_uu_not_empty_lt_a RSN (2, uu_lepoll_m) RSN (2, - uu_subset1 RSN (4, rel_is_fun)))] 1 - THEN TRYALL assume_tac); + (Least_uu_not_empty_lt_a RSN (2, uu_lepoll_m) RSN (2, + uu_subset1 RSN (4, rel_is_fun)))] 1 + THEN TRYALL assume_tac); by (rtac (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, supset_lepoll_imp_eq)) 1); by (REPEAT (fast_tac (AC_cs addSIs [domain_uu_subset, nat_succI]) 1)); val uu_Least_is_fun = result(); goalw thy [vv2_def] - "!!a. [| ALL g \ -\ domain(uu(f, b, g, d)) eqpoll succ(m); \ -\ ALL b vv2(f,b,g,s) <= f`g"; + "!!a. [| ALL g \ +\ domain(uu(f, b, g, d)) eqpoll succ(m); \ +\ ALL b vv2(f,b,g,s) <= f`g"; by (split_tac [expand_if] 1); by (fast_tac (FOL_cs addSEs [uu_Least_is_fun] - addSIs [empty_subsetI, not_emptyI, - singleton_subsetI, apply_type]) 1); + addSIs [empty_subsetI, not_emptyI, + singleton_subsetI, apply_type]) 1); val vv2_subset = result(); (* ********************************************************************** *) -(* Case 2 : Union of images is the whole "y" *) +(* Case 2 : Union of images is the whole "y" *) (* ********************************************************************** *) goalw thy [gg2_def] - "!!a. [| ALL g \ -\ domain(uu(f,b,g,d)) eqpoll succ(m); \ -\ ALL b (UN g \ +\ domain(uu(f,b,g,d)) eqpoll succ(m); \ +\ ALL b (UN g ww2(f,b,g,d) lepoll m"; +\ |] ==> ww2(f,b,g,d) lepoll m"; by (excluded_middle_tac "f`g = 0" 1); by (asm_simp_tac (OrdQuant_ss addsimps [empty_lepollI]) 2); by (dtac ospec 1 THEN (assume_tac 1)); by (rtac Diff_lepoll 1 - THEN (TRYALL assume_tac)); + THEN (TRYALL assume_tac)); by (asm_simp_tac (OrdQuant_ss addsimps [vv2_def, expand_if, not_emptyI]) 1); val ww2_lepoll = result(); goalw thy [gg2_def] - "!!a. [| ALL g \ -\ domain(uu(f,b,g,d)) eqpoll succ(m); \ -\ ALL b \ +\ domain(uu(f,b,g,d)) eqpoll succ(m); \ +\ ALL b gg2(f,a,b,s) ` g lepoll m"; by (asm_simp_tac OrdQuant_ss 1); by (safe_tac (OrdQuant_cs addSEs [lt_oadd_odiff_cases, lt_Ord2])); @@ -297,10 +297,10 @@ val gg2_lepoll_m = result(); (* ********************************************************************** *) -(* lemma ii *) +(* lemma ii *) (* ********************************************************************** *) goalw thy [NN_def] - "!!y. [| succ(m) : NN(y); y*y <= y; m:nat; m~=0 |] ==> m : NN(y)"; + "!!y. [| succ(m) : NN(y); y*y <= y; m:nat; m~=0 |] ==> m : NN(y)"; by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1)); by (resolve_tac [quant_domain_uu_lepoll_m RS cases RS disjE] 1 THEN (assume_tac 1)); @@ -308,7 +308,7 @@ by (asm_full_simp_tac (ZF_ss addsimps [lesspoll_succ_iff]) 1); by (res_inst_tac [("x","a++a")] exI 1); by (fast_tac (OrdQuant_cs addSIs [Ord_oadd, domain_gg1, UN_gg1_eq, - gg1_lepoll_m]) 1); + gg1_lepoll_m]) 1); (* case 2 *) by (REPEAT (eresolve_tac [oexE, conjE] 1)); by (res_inst_tac [("A","f`?B")] not_emptyE 1 THEN (assume_tac 1)); @@ -319,7 +319,7 @@ (*Calling fast_tac might get rid of the res_inst_tac calls, but it is just too slow.*) by (asm_simp_tac (OrdQuant_ss addsimps - [Ord_oadd, domain_gg2, UN_gg2_eq, gg2_lepoll_m]) 1); + [Ord_oadd, domain_gg2, UN_gg2_eq, gg2_lepoll_m]) 1); val lemma_ii = result(); @@ -344,7 +344,7 @@ val le_subsets = result(); goal thy "!!n m. [| n le m; m:nat |] ==> \ -\ rec(n, x, %k r. r Un r*r) <= rec(m, x, %k r. r Un r*r)"; +\ rec(n, x, %k r. r Un r*r) <= rec(m, x, %k r. r Un r*r)"; by (resolve_tac [z_n_subset_z_succ_n RS le_subsets] 1 THEN (TRYALL assume_tac)); by (eresolve_tac [Ord_nat RSN (2, ltI) RSN (2, lt_trans1) RS ltD] 1 @@ -358,23 +358,23 @@ by (res_inst_tac [("a","succ(n Un na)")] UN_I 1); by (eresolve_tac [Un_nat_type RS nat_succI] 1 THEN (assume_tac 1)); by (fast_tac (ZF_cs addIs [le_imp_rec_subset RS subsetD] - addSIs [Un_upper1_le, Un_upper2_le, Un_nat_type] - addSEs [nat_into_Ord] addss nat_ss) 1); + addSIs [Un_upper1_le, Un_upper2_le, Un_nat_type] + addSEs [nat_into_Ord] addss nat_ss) 1); val lemma_iv = result(); (* ********************************************************************** *) -(* Rubin & Rubin wrote : *) +(* Rubin & Rubin wrote : *) (* "It follows from (ii) and mathematical induction that if y*y <= y then *) -(* y can be well-ordered" *) +(* y can be well-ordered" *) -(* In fact we have to prove : *) -(* * WO6 ==> NN(y) ~= 0 *) -(* * reverse induction which lets us infer that 1 : NN(y) *) -(* * 1 : NN(y) ==> y can be well-ordered *) +(* In fact we have to prove : *) +(* * WO6 ==> NN(y) ~= 0 *) +(* * reverse induction which lets us infer that 1 : NN(y) *) +(* * 1 : NN(y) ==> y can be well-ordered *) (* ********************************************************************** *) (* ********************************************************************** *) -(* WO6 ==> NN(y) ~= 0 *) +(* WO6 ==> NN(y) ~= 0 *) (* ********************************************************************** *) goalw thy [WO6_def, NN_def] "!!y. WO6 ==> NN(y) ~= 0"; @@ -382,16 +382,16 @@ val WO6_imp_NN_not_empty = result(); (* ********************************************************************** *) -(* 1 : NN(y) ==> y can be well-ordered *) +(* 1 : NN(y) ==> y can be well-ordered *) (* ********************************************************************** *) goal thy "!!f. [| (UN b EX c EX c f` (LEAST i. f`i = {x}) = {x}"; +\ ==> f` (LEAST i. f`i = {x}) = {x}"; by (dtac lemma1 1 THEN REPEAT (assume_tac 1)); by (fast_tac (AC_cs addSEs [lt_Ord] addIs [LeastI]) 1); val lemma2 = result(); @@ -415,12 +415,12 @@ val y_well_ord = result(); (* ********************************************************************** *) -(* reverse induction which lets us infer that 1 : NN(y) *) +(* reverse induction which lets us infer that 1 : NN(y) *) (* ********************************************************************** *) val [prem1, prem2] = goal thy - "[| n:nat; !!m. [| m:nat; m~=0; P(succ(m)) |] ==> P(m) |] \ -\ ==> n~=0 --> P(n) --> P(1)"; + "[| n:nat; !!m. [| m:nat; m~=0; P(succ(m)) |] ==> P(m) |] \ +\ ==> n~=0 --> P(n) --> P(1)"; by (res_inst_tac [("n","n")] nat_induct 1); by (rtac prem1 1); by (fast_tac ZF_cs 1); @@ -430,9 +430,9 @@ val rev_induct_lemma = result(); val prems = goal thy - "[| P(n); n:nat; n~=0; \ -\ !!m. [| m:nat; m~=0; P(succ(m)) |] ==> P(m) |] \ -\ ==> P(1)"; + "[| P(n); n:nat; n~=0; \ +\ !!m. [| m:nat; m~=0; P(succ(m)) |] ==> P(m) |] \ +\ ==> P(1)"; by (resolve_tac [rev_induct_lemma RS impE] 1); by (etac impE 4 THEN (assume_tac 5)); by (REPEAT (ares_tac prems 1)); @@ -448,7 +448,7 @@ val lemma3 = result(); (* ********************************************************************** *) -(* Main theorem "WO6 ==> WO1" *) +(* Main theorem "WO6 ==> WO1" *) (* ********************************************************************** *) (* another helpful lemma *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/WO_AC.ML --- a/src/ZF/AC/WO_AC.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/WO_AC.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/WO_AC.ML +(* Title: ZF/AC/WO_AC.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski Lemmas used in the proofs like WO? ==> AC? *) @@ -8,9 +8,9 @@ open WO_AC; goal thy "!!A. [| well_ord(Union(A),r); 0~:A; B:A |] \ -\ ==> (THE b. first(b,B,r)) : B"; +\ ==> (THE b. first(b,B,r)) : B"; by (fast_tac (AC_cs addSEs [well_ord_imp_ex1_first RS theI RS - (first_def RS def_imp_iff RS iffD1 RS conjunct1)]) 1); + (first_def RS def_imp_iff RS iffD1 RS conjunct1)]) 1); val first_in_B = result(); goal thy "!!A. [| well_ord(Union(A), R); 0~:A |] ==> EX f. f:(PROD X:A. X)"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/first.ML --- a/src/ZF/AC/first.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/first.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/first.ML +(* Title: ZF/AC/first.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski Lemmas involving the first element of a well ordered set *) @@ -11,8 +11,8 @@ by (fast_tac AC_cs 1); val first_is_elem = result(); -goalw thy [well_ord_def, wf_on_def, wf_def, first_def] - "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (EX! b. first(b,B,r))"; +goalw thy [well_ord_def, wf_on_def, wf_def, first_def] + "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (EX! b. first(b,B,r))"; by (REPEAT (eresolve_tac [conjE,allE,disjE] 1)); by (contr_tac 1); by (etac bexE 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/recfunAC16.ML --- a/src/ZF/AC/recfunAC16.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/recfunAC16.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/recfunAC16.ML +(* Title: ZF/AC/recfunAC16.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski Properties of the recursive definition used in the proof of WO2 ==> AC16 *) @@ -8,7 +8,7 @@ open recfunAC16; (* ********************************************************************** *) -(* Basic properties of recfunAC16 *) +(* Basic properties of recfunAC16 *) (* ********************************************************************** *) goalw thy [recfunAC16_def] "recfunAC16(f,fa,0,a) = 0"; @@ -18,47 +18,47 @@ goalw thy [recfunAC16_def] "recfunAC16(f,fa,succ(i),a) = \ \ if (EX y:recfunAC16(f,fa,i,a). fa ` i <= y, recfunAC16(f,fa,i,a), \ \ recfunAC16(f,fa,i,a) Un {f ` (LEAST j. fa ` i <= f ` j & \ -\ (ALL b (ALL t:recfunAC16(f,fa,i,a). ~ fa`b <= t))))})"; +\ (ALL b (ALL t:recfunAC16(f,fa,i,a). ~ fa`b <= t))))})"; by (rtac transrec2_succ 1); val recfunAC16_succ = result(); goalw thy [recfunAC16_def] "!!i. Limit(i) \ -\ ==> recfunAC16(f,fa,i,a) = (UN j recfunAC16(f,fa,i,a) = (UN j j transrec2(j, 0, B) <= transrec2(i, 0, B)"; +\ ==> j transrec2(j, 0, B) <= transrec2(i, 0, B)"; by (resolve_tac [prem2 RS trans_induct] 1); by (rtac Ord_cases 1 THEN (REPEAT (assume_tac 1))); by (fast_tac lt_cs 1); by (asm_simp_tac (AC_ss addsimps [transrec2_succ]) 1); by (fast_tac (FOL_cs addSIs [succI1, prem1] - addSEs [ballE, leE, prem1 RSN (2, subset_trans)]) 1); + addSEs [ballE, leE, prem1 RSN (2, subset_trans)]) 1); by (fast_tac (AC_cs addIs [OUN_I, ltI] - addSEs [Limit_has_succ RS ltE, succI1 RSN (2, Ord_in_Ord) RS le_refl, - transrec2_Limit RS ssubst]) 1); + addSEs [Limit_has_succ RS ltE, succI1 RSN (2, Ord_in_Ord) RS le_refl, + transrec2_Limit RS ssubst]) 1); val transrec2_mono_lemma = result(); val [prem1, prem2] = goal thy "[| !!g r. r <= B(g,r); j le i |] \ -\ ==> transrec2(j, 0, B) <= transrec2(i, 0, B)"; +\ ==> transrec2(j, 0, B) <= transrec2(i, 0, B)"; by (resolve_tac [prem2 RS leE] 1); by (resolve_tac [transrec2_mono_lemma RS impE] 1); by (TRYALL (fast_tac (AC_cs addSIs [prem1, prem2, lt_Ord2]))); val transrec2_mono = result(); (* ********************************************************************** *) -(* Monotonicity of recfunAC16 *) +(* Monotonicity of recfunAC16 *) (* ********************************************************************** *) goalw thy [recfunAC16_def] - "!!i. i le j ==> recfunAC16(f, g, i, a) <= recfunAC16(f, g, j, a)"; + "!!i. i le j ==> recfunAC16(f, g, i, a) <= recfunAC16(f, g, j, a)"; by (rtac transrec2_mono 1); by (REPEAT (fast_tac (AC_cs addIs [expand_if RS iffD2]) 1)); val recfunAC16_mono = result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/AC/rel_is_fun.ML --- a/src/ZF/AC/rel_is_fun.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/AC/rel_is_fun.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/AC/rel_is_fun.ML +(* Title: ZF/AC/rel_is_fun.ML ID: $Id$ - Author: Krzysztof Grabczewski + Author: Krzysztof Grabczewski Lemmas needed to state when a finite relation is a function. @@ -13,15 +13,15 @@ "!!m. [| m:nat; u lepoll m |] ==> domain(u) lepoll m"; by (etac exE 1); by (res_inst_tac [("x", - "lam x:domain(u). LEAST i. EX y. : u & f` = i")] exI 1); + "lam x:domain(u). LEAST i. EX y. : u & f` = i")] exI 1); by (res_inst_tac [("d","%y. fst(converse(f)`y)")] lam_injective 1); by (fast_tac (ZF_cs addIs [LeastI2, nat_into_Ord RS Ord_in_Ord, - inj_is_fun RS apply_type]) 1); + inj_is_fun RS apply_type]) 1); by (etac domainE 1); by (forward_tac [inj_is_fun RS apply_type] 1 THEN (atac 1)); by (rtac LeastI2 1); by (REPEAT (fast_tac (ZF_cs addIs [fst_conv, left_inverse RS ssubst] - addSEs [nat_into_Ord RS Ord_in_Ord]) 1)); + addSEs [nat_into_Ord RS Ord_in_Ord]) 1)); val lepoll_m_imp_domain_lepoll_m = result(); goal ZF.thy "!!r. [| : r; c~=b |] ==> domain(r-{}) = domain(r)"; @@ -34,17 +34,17 @@ goalw Cardinal.thy [function_def] "!!r. [| succ(m) lepoll domain(r); r lepoll succ(m); m:nat |] ==> \ -\ function(r)"; +\ function(r)"; by (safe_tac ZF_cs); by (resolve_tac [excluded_middle RS disjE] 1 THEN (atac 2)); by (fast_tac (ZF_cs addSEs [lepoll_trans RS succ_lepoll_natE, - Diff_sing_lepoll RSN (2, lepoll_m_imp_domain_lepoll_m)] - addEs [not_sym RSN (2, domain_Diff_eq_domain) RS subst]) 1); + Diff_sing_lepoll RSN (2, lepoll_m_imp_domain_lepoll_m)] + addEs [not_sym RSN (2, domain_Diff_eq_domain) RS subst]) 1); val rel_domain_ex1 = result(); goal Cardinal.thy "!!r. [| succ(m) lepoll domain(r); r lepoll succ(m); m:nat; \ -\ r<=A*B; A=domain(r) |] ==> r: A->B"; +\ r<=A*B; A=domain(r) |] ==> r: A->B"; by (hyp_subst_tac 1); by (asm_simp_tac (ZF_ss addsimps [Pi_iff, rel_domain_ex1]) 1); val rel_is_fun = result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Arith.ML --- a/src/ZF/Arith.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Arith.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/arith.ML +(* Title: ZF/arith.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge For arith.thy. Arithmetic operators and their definitions @@ -47,7 +47,7 @@ val nat_typechecks = [rec_type, nat_0I, nat_1I, nat_succI, Ord_nat]; val nat_simps = [rec_0, rec_succ, not_lt0, nat_0_le, le0_iff, succ_le_iff, - nat_le_refl]; + nat_le_refl]; val nat_ss = ZF_ss addsimps (nat_simps @ nat_typechecks); @@ -113,15 +113,15 @@ by (ALLGOALS (asm_simp_tac (nat_ss addsimps (prems @ [le_iff, diff_0, diff_0_eq_0, - diff_succ_succ, nat_into_Ord])))); + diff_succ_succ, nat_into_Ord])))); qed "diff_le_self"; (*** Simplification over add, mult, diff ***) val arith_typechecks = [add_type, mult_type, diff_type]; val arith_simps = [add_0, add_succ, - mult_0, mult_succ, - diff_0, diff_0_eq_0, diff_succ_succ]; + mult_0, mult_succ, + diff_0, diff_0_eq_0, diff_succ_succ]; val arith_ss = nat_ss addsimps (arith_simps@arith_typechecks); @@ -195,7 +195,7 @@ (fn prems=> [ (nat_ind_tac "m" prems 1), (ALLGOALS (asm_simp_tac - (arith_ss addsimps (prems@[mult_0_right, mult_succ_right])))) ]); + (arith_ss addsimps (prems@[mult_0_right, mult_succ_right])))) ]); (*addition distributes over multiplication*) qed_goal "add_mult_distrib" Arith.thy @@ -224,8 +224,8 @@ "!!m n k. [| m:nat; n:nat; k:nat |] ==> m #* (n #* k) = n #* (m #* k)" (fn _ => [rtac (mult_commute RS trans) 1, rtac (mult_assoc RS trans) 3, - rtac (mult_commute RS subst_context) 6, - REPEAT (ares_tac [mult_type] 1)]); + rtac (mult_commute RS subst_context) 6, + REPEAT (ares_tac [mult_type] 1)]); val mult_ac = [mult_assoc,mult_commute,mult_left_commute]; @@ -277,13 +277,13 @@ by (ALLGOALS (asm_simp_tac (nat_ss addsimps [diff_le_self,diff_succ_succ]))); qed "div_termination"; -val div_rls = (*for mod and div*) +val div_rls = (*for mod and div*) nat_typechecks @ [Ord_transrec_type, apply_type, div_termination RS ltD, if_type, nat_into_Ord, not_lt_iff_le RS iffD1]; val div_ss = ZF_ss addsimps [nat_into_Ord, div_termination RS ltD, - not_lt_iff_le RS iffD2]; + not_lt_iff_le RS iffD2]; (*Type checking depends upon termination!*) goalw Arith.thy [mod_def] "!!m n. [| 0 m mod n : nat"; @@ -331,8 +331,8 @@ (*case n le x*) by (asm_full_simp_tac (arith_ss addsimps [not_lt_iff_le, nat_into_Ord, - mod_geq, div_geq, add_assoc, - div_termination RS ltD, add_diff_inverse]) 1); + mod_geq, div_geq, add_assoc, + div_termination RS ltD, add_diff_inverse]) 1); qed "mod_div_equality"; @@ -363,16 +363,16 @@ by (rtac (add_lt_mono1 RS lt_trans) 1); by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1)); by (EVERY [rtac (add_commute RS ssubst) 1, - rtac (add_commute RS ssubst) 3, - rtac add_lt_mono1 5]); + rtac (add_commute RS ssubst) 3, + rtac add_lt_mono1 5]); by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1)); qed "add_lt_mono"; (*A [clumsy] way of lifting < monotonicity to le monotonicity *) val lt_mono::ford::prems = goal Ordinal.thy - "[| !!i j. [| i f(i) < f(j); \ -\ !!i. i:k ==> Ord(f(i)); \ -\ i le j; j:k \ + "[| !!i j. [| i f(i) < f(j); \ +\ !!i. i:k ==> Ord(f(i)); \ +\ i le j; j:k \ \ |] ==> f(i) le f(j)"; by (cut_facts_tac prems 1); by (fast_tac (lt_cs addSIs [lt_mono,ford] addSEs [leE]) 1); @@ -391,7 +391,7 @@ by (rtac (add_le_mono1 RS le_trans) 1); by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); by (EVERY [rtac (add_commute RS ssubst) 1, - rtac (add_commute RS ssubst) 3, - rtac add_le_mono1 5]); + rtac (add_commute RS ssubst) 3, + rtac add_le_mono1 5]); by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); qed "add_le_mono"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Bool.ML --- a/src/ZF/Bool.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Bool.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/bool +(* Title: ZF/bool ID: $Id$ - Author: Martin D Coen, Cambridge University Computer Laboratory + Author: Martin D Coen, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge For ZF/bool.thy. Booleans in Zermelo-Fraenkel Set Theory @@ -89,7 +89,7 @@ (fn prems=> [ (typechk_tac(prems@[bool_1I, bool_0I, cond_type, not_type])) ]); val bool_typechecks = [bool_1I, bool_0I, cond_type, not_type, and_type, - or_type, xor_type] + or_type, xor_type] val bool_simps = [cond_1,cond_0,not_1,not_0,and_1,and_0,or_1,or_0,xor_1,xor_0]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Cardinal.ML --- a/src/ZF/Cardinal.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Cardinal.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Cardinal.ML +(* Title: ZF/Cardinal.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Cardinals in Zermelo-Fraenkel Set Theory @@ -20,13 +20,13 @@ qed "decomp_bnd_mono"; val [gfun] = goal Cardinal.thy - "g: Y->X ==> \ -\ g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) = \ + "g: Y->X ==> \ +\ g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) = \ \ X - lfp(X, %W. X - g``(Y - f``W)) "; by (res_inst_tac [("P", "%u. ?v = X-u")] (decomp_bnd_mono RS lfp_Tarski RS ssubst) 1); by (simp_tac (ZF_ss addsimps [subset_refl, double_complement, - gfun RS fun_is_rel RS image_subset]) 1); + gfun RS fun_is_rel RS image_subset]) 1); qed "Banach_last_equation"; val prems = goal Cardinal.thy @@ -192,9 +192,9 @@ qed_goal "LeastI2" Cardinal.thy "[| P(i); Ord(i); !!j. P(j) ==> Q(j) |] ==> Q(LEAST j. P(j))" (fn prems => [ resolve_tac prems 1, - rtac LeastI 1, - resolve_tac prems 1, - resolve_tac prems 1 ]); + rtac LeastI 1, + resolve_tac prems 1, + resolve_tac prems 1 ]); (*If there is no such P then LEAST is vacuously 0*) goalw Cardinal.thy [Least_def] @@ -281,8 +281,8 @@ by (safe_tac (ZF_cs addSIs [CardI, Card_is_Ord])); by (fast_tac ZF_cs 2); by (rewrite_goals_tac [Card_def, cardinal_def]); -by (resolve_tac [less_LeastE] 1); -by (eresolve_tac [subst] 2); +by (rtac less_LeastE 1); +by (etac subst 2); by (ALLGOALS assume_tac); qed "Card_iff_initial"; @@ -344,7 +344,7 @@ goal Cardinal.thy "!!i j. [| |i| < K; Ord(i); Card(K) |] ==> i < K"; by (asm_simp_tac (ZF_ss addsimps - [cardinal_lt_imp_lt, Card_is_Ord, Card_cardinal_eq]) 1); + [cardinal_lt_imp_lt, Card_is_Ord, Card_cardinal_eq]) 1); qed "Card_lt_imp_lt"; goal Cardinal.thy "!!i j. [| Ord(i); Card(K) |] ==> (|i| < K) <-> (i < K)"; @@ -353,8 +353,8 @@ goal Cardinal.thy "!!i j. [| Ord(i); Card(K) |] ==> (K le |i|) <-> (K le i)"; by (asm_simp_tac (ZF_ss addsimps - [Card_lt_iff, Card_is_Ord, Ord_cardinal, - not_lt_iff_le RS iff_sym]) 1); + [Card_lt_iff, Card_is_Ord, Ord_cardinal, + not_lt_iff_le RS iff_sym]) 1); qed "Card_le_iff"; @@ -393,7 +393,7 @@ by (rtac ballI 1); by (eres_inst_tac [("n","n")] natE 1); by (asm_simp_tac (ZF_ss addsimps [lepoll_def, inj_def, - succI1 RS Pi_empty2]) 1); + succI1 RS Pi_empty2]) 1); by (fast_tac (ZF_cs addSIs [succ_leI] addSDs [succ_lepoll_succD]) 1); val nat_lepoll_imp_le_lemma = result(); @@ -442,7 +442,7 @@ goal Cardinal.thy "!!m. m:nat ==> A lesspoll succ(m) <-> A lepoll m"; by (fast_tac (ZF_cs addSIs [lepoll_imp_lesspoll_succ, - lesspoll_succ_imp_lepoll]) 1); + lesspoll_succ_imp_lepoll]) 1); qed "lesspoll_succ_iff"; goal Cardinal.thy "!!A m. [| A lepoll succ(m); m:nat |] ==> \ @@ -549,7 +549,7 @@ "!!f. [| f: inj(A,B); A Int B = 0 |] ==> A Un (B - range(f)) eqpoll B"; by (rtac exI 1); by (res_inst_tac [("c", "%x. if(x:A, f`x, x)"), - ("d", "%y. if(y: range(f), converse(f)`y, y)")] + ("d", "%y. if(y: range(f), converse(f)`y, y)")] lam_bijective 1); by (fast_tac (ZF_cs addSIs [if_type, apply_type] addIs [inj_is_fun]) 1); by (asm_simp_tac @@ -612,7 +612,7 @@ "!!X. [| Y lepoll X; Finite(X) |] ==> Finite(Y)"; by (fast_tac (ZF_cs addSEs [eqpollE] addEs [lepoll_trans RS - rewrite_rule [Finite_def] lepoll_nat_imp_Finite]) 1); + rewrite_rule [Finite_def] lepoll_nat_imp_Finite]) 1); qed "lepoll_Finite"; goalw Cardinal.thy [Finite_def] "!!x. Finite(x) ==> Finite(cons(y,x))"; @@ -662,7 +662,7 @@ qed "nat_well_ord_converse_Memrel"; goal Cardinal.thy - "!!A. [| well_ord(A,r); \ + "!!A. [| well_ord(A,r); \ \ well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r)))) \ \ |] ==> well_ord(A,converse(r))"; by (resolve_tac [well_ord_Int_iff RS iffD1] 1); @@ -670,7 +670,7 @@ by (assume_tac 1); by (asm_full_simp_tac (ZF_ss addsimps [rvimage_converse, converse_Int, converse_prod, - ordertype_ord_iso RS ord_iso_rvimage_eq]) 1); + ordertype_ord_iso RS ord_iso_rvimage_eq]) 1); qed "well_ord_converse"; goal Cardinal.thy diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/CardinalArith.ML --- a/src/ZF/CardinalArith.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/CardinalArith.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/CardinalArith.ML +(* Title: ZF/CardinalArith.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Cardinal arithmetic -- WITHOUT the Axiom of Choice @@ -39,7 +39,7 @@ by (etac Fin_induct 1); by (fast_tac (ZF_cs addIs [eqpoll_refl, nat_0I]) 1); by (fast_tac (ZF_cs addSIs [cons_eqpoll_cong, - rewrite_rule [succ_def] nat_succI] + rewrite_rule [succ_def] nat_succI] addSEs [mem_irrefl]) 1); qed "Fin_eqpoll"; @@ -63,19 +63,19 @@ goalw CardinalArith.thy [eqpoll_def] "(A+B)+C eqpoll A+(B+C)"; by (rtac exI 1); -by (resolve_tac [sum_assoc_bij] 1); +by (rtac sum_assoc_bij 1); qed "sum_assoc_eqpoll"; (*Unconditional version requires AC*) goalw CardinalArith.thy [cadd_def] - "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ + "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ \ (i |+| j) |+| k = i |+| (j |+| k)"; by (rtac cardinal_cong 1); by (rtac ([well_ord_cardinal_eqpoll, eqpoll_refl] MRS sum_eqpoll_cong RS - eqpoll_trans) 1); + eqpoll_trans) 1); by (rtac (sum_assoc_eqpoll RS eqpoll_trans) 2); by (rtac ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS sum_eqpoll_cong RS - eqpoll_sym) 2); + eqpoll_sym) 2); by (REPEAT (ares_tac [well_ord_radd] 1)); qed "well_ord_cadd_assoc"; @@ -88,7 +88,7 @@ goalw CardinalArith.thy [cadd_def] "!!K. Card(K) ==> 0 |+| K = K"; by (asm_simp_tac (ZF_ss addsimps [sum_0_eqpoll RS cardinal_cong, - Card_cardinal_eq]) 1); + Card_cardinal_eq]) 1); qed "cadd_0"; (** Addition by another cardinal **) @@ -134,11 +134,11 @@ goalw CardinalArith.thy [eqpoll_def] "succ(A)+B eqpoll succ(A+B)"; by (rtac exI 1); by (res_inst_tac [("c", "%z.if(z=Inl(A),A+B,z)"), - ("d", "%z.if(z=A+B,Inl(A),z)")] + ("d", "%z.if(z=A+B,Inl(A),z)")] lam_bijective 1); by (ALLGOALS (asm_simp_tac (case_ss addsimps [succI2, mem_imp_not_eq] - setloop eresolve_tac [sumE,succE]))); + setloop eresolve_tac [sumE,succE]))); qed "sum_succ_eqpoll"; (*Pulling the succ(...) outside the |...| requires m, n: nat *) @@ -157,7 +157,7 @@ by (nat_ind_tac "m" [mnat] 1); by (asm_simp_tac (arith_ss addsimps [nat_into_Card RS cadd_0]) 1); by (asm_simp_tac (arith_ss addsimps [nat_into_Ord, cadd_succ_lemma, - nat_into_Card RS Card_cardinal_eq]) 1); + nat_into_Card RS Card_cardinal_eq]) 1); qed "nat_cadd_eq_add"; @@ -182,19 +182,19 @@ goalw CardinalArith.thy [eqpoll_def] "(A*B)*C eqpoll A*(B*C)"; by (rtac exI 1); -by (resolve_tac [prod_assoc_bij] 1); +by (rtac prod_assoc_bij 1); qed "prod_assoc_eqpoll"; (*Unconditional version requires AC*) goalw CardinalArith.thy [cmult_def] - "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ + "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ \ (i |*| j) |*| k = i |*| (j |*| k)"; by (rtac cardinal_cong 1); by (rtac ([well_ord_cardinal_eqpoll, eqpoll_refl] MRS prod_eqpoll_cong RS - eqpoll_trans) 1); + eqpoll_trans) 1); by (rtac (prod_assoc_eqpoll RS eqpoll_trans) 2); by (rtac ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS prod_eqpoll_cong RS - eqpoll_sym) 2); + eqpoll_sym) 2); by (REPEAT (ares_tac [well_ord_rmult] 1)); qed "well_ord_cmult_assoc"; @@ -202,18 +202,18 @@ goalw CardinalArith.thy [eqpoll_def] "(A+B)*C eqpoll (A*C)+(B*C)"; by (rtac exI 1); -by (resolve_tac [sum_prod_distrib_bij] 1); +by (rtac sum_prod_distrib_bij 1); qed "sum_prod_distrib_eqpoll"; goalw CardinalArith.thy [cadd_def, cmult_def] - "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ + "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ \ (i |+| j) |*| k = (i |*| k) |+| (j |*| k)"; by (rtac cardinal_cong 1); by (rtac ([well_ord_cardinal_eqpoll, eqpoll_refl] MRS prod_eqpoll_cong RS - eqpoll_trans) 1); + eqpoll_trans) 1); by (rtac (sum_prod_distrib_eqpoll RS eqpoll_trans) 2); by (rtac ([well_ord_cardinal_eqpoll, well_ord_cardinal_eqpoll] MRS sum_eqpoll_cong RS - eqpoll_sym) 2); + eqpoll_sym) 2); by (REPEAT (ares_tac [well_ord_rmult, well_ord_radd] 1)); qed "well_ord_cadd_cmult_distrib"; @@ -227,7 +227,7 @@ goalw CardinalArith.thy [cmult_def] "0 |*| i = 0"; by (asm_simp_tac (ZF_ss addsimps [prod_0_eqpoll RS cardinal_cong, - Card_0 RS Card_cardinal_eq]) 1); + Card_0 RS Card_cardinal_eq]) 1); qed "cmult_0"; (** 1 is the identity for multiplication **) @@ -239,7 +239,7 @@ goalw CardinalArith.thy [cmult_def, succ_def] "!!K. Card(K) ==> 1 |*| K = K"; by (asm_simp_tac (ZF_ss addsimps [prod_singleton_eqpoll RS cardinal_cong, - Card_cardinal_eq]) 1); + Card_cardinal_eq]) 1); qed "cmult_1"; (*** Some inequalities for multiplication ***) @@ -280,7 +280,7 @@ by (REPEAT (etac exE 1)); by (res_inst_tac [("x", "lam :A*B. ")] exI 1); by (res_inst_tac [("d", "%.")] - lam_injective 1); + lam_injective 1); by (typechk_tac (inj_is_fun::ZF_typechecks)); by (etac SigmaE 1); by (asm_simp_tac (ZF_ss addsimps [left_inverse]) 1); @@ -299,7 +299,7 @@ goalw CardinalArith.thy [eqpoll_def] "succ(A)*B eqpoll B + A*B"; by (rtac exI 1); by (res_inst_tac [("c", "%. if(x=A, Inl(y), Inr())"), - ("d", "case(%y. , %z.z)")] + ("d", "case(%y. , %z.z)")] lam_bijective 1); by (safe_tac (ZF_cs addSEs [sumE])); by (ALLGOALS @@ -321,13 +321,13 @@ by (nat_ind_tac "m" [mnat] 1); by (asm_simp_tac (arith_ss addsimps [cmult_0]) 1); by (asm_simp_tac (arith_ss addsimps [nat_into_Ord, cmult_succ_lemma, - nat_cadd_eq_add]) 1); + nat_cadd_eq_add]) 1); qed "nat_cmult_eq_mult"; goal CardinalArith.thy "!!m n. Card(n) ==> 2 |*| n = n |+| n"; by (asm_simp_tac (ZF_ss addsimps [Ord_0, Ord_succ, cmult_0, cmult_succ_lemma, Card_is_Ord, - read_instantiate [("j","0")] cadd_commute, cadd_0]) 1); + read_instantiate [("j","0")] cadd_commute, cadd_0]) 1); qed "cmult_2"; @@ -341,19 +341,19 @@ "!!i. nat lepoll A ==> cons(u,A) lepoll A"; by (etac exE 1); by (res_inst_tac [("x", - "lam z:cons(u,A). if(z=u, f`0, \ + "lam z:cons(u,A). if(z=u, f`0, \ \ if(z: range(f), f`succ(converse(f)`z), z))")] exI 1); -by (res_inst_tac [("d", "%y. if(y: range(f), \ +by (res_inst_tac [("d", "%y. if(y: range(f), \ \ nat_case(u, %z.f`z, converse(f)`y), y)")] lam_injective 1); by (fast_tac (ZF_cs addSIs [if_type, nat_0I, nat_succI, apply_type] addIs [inj_is_fun, inj_converse_fun]) 1); by (asm_simp_tac (ZF_ss addsimps [inj_is_fun RS apply_rangeI, - inj_converse_fun RS apply_rangeI, - inj_converse_fun RS apply_funtype, - left_inverse, right_inverse, nat_0I, nat_succI, - nat_case_0, nat_case_succ] + inj_converse_fun RS apply_rangeI, + inj_converse_fun RS apply_funtype, + left_inverse, right_inverse, nat_0I, nat_succI, + nat_case_0, nat_case_succ] setloop split_tac [expand_if]) 1); qed "nat_cons_lepoll"; @@ -378,7 +378,7 @@ goalw CardinalArith.thy [InfCard_def] "!!K L. [| InfCard(K); Card(L) |] ==> InfCard(K Un L)"; by (asm_simp_tac (ZF_ss addsimps [Card_Un, Un_upper1_le RSN (2,le_trans), - Card_is_Ord]) 1); + Card_is_Ord]) 1); qed "InfCard_Un"; (*Kunen's Lemma 10.11*) @@ -412,7 +412,7 @@ goalw CardinalArith.thy [inj_def] "!!K. Ord(K) ==> (lam :K*K. ) : inj(K*K, K*K*K)"; by (fast_tac (ZF_cs addss ZF_ss - addIs [lam_type, Un_least_lt RS ltD, ltI]) 1); + addIs [lam_type, Un_least_lt RS ltD, ltI]) 1); qed "csquare_lam_inj"; goalw CardinalArith.thy [csquare_rel_def] @@ -438,7 +438,7 @@ goalw CardinalArith.thy [pred_def] "!!K. z pred(K*K, , csquare_rel(K)) <= succ(z)*succ(z)"; -by (safe_tac (lemmas_cs addSEs [SigmaE])); (*avoids using succCI*) +by (safe_tac (lemmas_cs addSEs [SigmaE])); (*avoids using succCI*) by (rtac (csquareD RS conjE) 1); by (rewtac lt_def); by (assume_tac 4); @@ -466,14 +466,14 @@ by (REPEAT_FIRST (etac succE)); by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [subset_Un_iff RS iff_sym, - subset_Un_iff2 RS iff_sym, OrdmemD]))); + subset_Un_iff2 RS iff_sym, OrdmemD]))); qed "csquare_or_eqI"; (** The cardinality of initial segments **) goal CardinalArith.thy "!!K. [| Limit(K); x \ -\ ordermap(K*K, csquare_rel(K)) ` < \ +\ ordermap(K*K, csquare_rel(K)) ` < \ \ ordermap(K*K, csquare_rel(K)) ` "; by (subgoals_tac ["z \ \ (EX r X. r <= K*K & X <= K & well_ord(X,r) & i = ordertype(X,r))"; -by (fast_tac subset_cs 1); (*It's vital to avoid reasoning about <=*) +by (fast_tac subset_cs 1); (*It's vital to avoid reasoning about <=*) qed "jump_cardinal_iff"; (*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*) @@ -662,9 +662,9 @@ (*The proof by contradiction: the bijection f yields a wellordering of X whose ordertype is jump_cardinal(K). *) goal CardinalArith.thy - "!!K. [| well_ord(X,r); r <= K * K; X <= K; \ -\ f : bij(ordertype(X,r), jump_cardinal(K)) \ -\ |] ==> jump_cardinal(K) : jump_cardinal(K)"; + "!!K. [| well_ord(X,r); r <= K * K; X <= K; \ +\ f : bij(ordertype(X,r), jump_cardinal(K)) \ +\ |] ==> jump_cardinal(K) : jump_cardinal(K)"; by (subgoal_tac "f O ordermap(X,r): bij(X, jump_cardinal(K))" 1); by (REPEAT (ares_tac [comp_bij, ordermap_bij] 2)); by (resolve_tac [jump_cardinal_iff RS iffD2] 1); @@ -675,7 +675,7 @@ by (rtac (Ord_jump_cardinal RS well_ord_Memrel) 1); by (asm_simp_tac (ZF_ss addsimps [well_ord_Memrel RSN (2, bij_ordertype_vimage), - ordertype_Memrel, Ord_jump_cardinal]) 1); + ordertype_Memrel, Ord_jump_cardinal]) 1); qed "Card_jump_cardinal_lemma"; (*The hard part of Theorem 10.16: jump_cardinal(K) is itself a cardinal*) @@ -692,7 +692,7 @@ "!!K. Ord(K) ==> Card(csucc(K)) & K < csucc(K)"; by (rtac LeastI 1); by (REPEAT (ares_tac [conjI, Card_jump_cardinal, K_lt_jump_cardinal, - Ord_jump_cardinal] 1)); + Ord_jump_cardinal] 1)); qed "csucc_basic"; bind_thm ("Card_csucc", csucc_basic RS conjunct1); @@ -733,6 +733,6 @@ goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> InfCard(csucc(K))"; by (asm_simp_tac (ZF_ss addsimps [Card_csucc, Card_is_Ord, - lt_csucc RS leI RSN (2,le_trans)]) 1); + lt_csucc RS leI RSN (2,le_trans)]) 1); qed "InfCard_csucc"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Cardinal_AC.ML --- a/src/ZF/Cardinal_AC.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Cardinal_AC.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Cardinal_AC.ML +(* Title: ZF/Cardinal_AC.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Cardinal arithmetic WITH the Axiom of Choice @@ -14,7 +14,7 @@ goal Cardinal_AC.thy "|A| eqpoll A"; by (resolve_tac [AC_well_ord RS exE] 1); -by (eresolve_tac [well_ord_cardinal_eqpoll] 1); +by (etac well_ord_cardinal_eqpoll 1); qed "cardinal_eqpoll"; val cardinal_idem = cardinal_eqpoll RS cardinal_cong; @@ -22,13 +22,13 @@ goal Cardinal_AC.thy "!!X Y. |X| = |Y| ==> X eqpoll Y"; by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); -by (resolve_tac [well_ord_cardinal_eqE] 1); +by (rtac well_ord_cardinal_eqE 1); by (REPEAT_SOME assume_tac); qed "cardinal_eqE"; goal Cardinal_AC.thy "!!A B. A lepoll B ==> |A| le |B|"; by (resolve_tac [AC_well_ord RS exE] 1); -by (eresolve_tac [well_ord_lepoll_imp_Card_le] 1); +by (etac well_ord_lepoll_imp_Card_le 1); by (assume_tac 1); qed "lepoll_imp_Card_le"; @@ -36,7 +36,7 @@ by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); -by (resolve_tac [well_ord_cadd_assoc] 1); +by (rtac well_ord_cadd_assoc 1); by (REPEAT_SOME assume_tac); qed "cadd_assoc"; @@ -44,7 +44,7 @@ by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); -by (resolve_tac [well_ord_cmult_assoc] 1); +by (rtac well_ord_cmult_assoc 1); by (REPEAT_SOME assume_tac); qed "cmult_assoc"; @@ -52,13 +52,13 @@ by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); by (resolve_tac [AC_well_ord RS exE] 1); -by (resolve_tac [well_ord_cadd_cmult_distrib] 1); +by (rtac well_ord_cadd_cmult_distrib 1); by (REPEAT_SOME assume_tac); qed "cadd_cmult_distrib"; goal Cardinal_AC.thy "!!A. InfCard(|A|) ==> A*A eqpoll A"; by (resolve_tac [AC_well_ord RS exE] 1); -by (eresolve_tac [well_ord_InfCard_square_eq] 1); +by (etac well_ord_InfCard_square_eq 1); by (assume_tac 1); qed "InfCard_square_eq"; @@ -67,7 +67,7 @@ goal Cardinal_AC.thy "!!A B. |A| le |B| ==> A lepoll B"; by (resolve_tac [cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll RS - lepoll_trans] 1); + lepoll_trans] 1); by (eresolve_tac [le_imp_subset RS subset_imp_lepoll RS lepoll_trans] 1); by (resolve_tac [cardinal_eqpoll RS eqpoll_imp_lepoll] 1); qed "Card_le_imp_lepoll"; @@ -82,46 +82,46 @@ by (etac CollectE 1); by (res_inst_tac [("A1", "Y"), ("B1", "%y. f-``{y}")] (AC_Pi RS exE) 1); by (fast_tac (ZF_cs addSEs [apply_Pair]) 1); -by (resolve_tac [exI] 1); +by (rtac exI 1); by (rtac f_imp_injective 1); -by (resolve_tac [Pi_type] 1 THEN assume_tac 1); +by (rtac Pi_type 1 THEN assume_tac 1); by (fast_tac (ZF_cs addDs [apply_type] addDs [Pi_memberD]) 1); by (fast_tac (ZF_cs addDs [apply_type] addEs [apply_equality]) 1); qed "surj_implies_inj"; (*Kunen's Lemma 10.20*) goal Cardinal_AC.thy "!!f. f: surj(X,Y) ==> |Y| le |X|"; -by (resolve_tac [lepoll_imp_Card_le] 1); +by (rtac lepoll_imp_Card_le 1); by (eresolve_tac [surj_implies_inj RS exE] 1); by (rewtac lepoll_def); -by (eresolve_tac [exI] 1); +by (etac exI 1); qed "surj_implies_cardinal_le"; (*Kunen's Lemma 10.21*) goal Cardinal_AC.thy "!!K. [| InfCard(K); ALL i:K. |X(i)| le K |] ==> |UN i:K. X(i)| le K"; by (asm_full_simp_tac (ZF_ss addsimps [InfCard_is_Card, le_Card_iff]) 1); -by (resolve_tac [lepoll_trans] 1); +by (rtac lepoll_trans 1); by (resolve_tac [InfCard_square_eq RS eqpoll_imp_lepoll] 2); by (asm_simp_tac (ZF_ss addsimps [InfCard_is_Card, Card_cardinal_eq]) 2); -by (rewrite_goals_tac [lepoll_def]); +by (rewtac lepoll_def); by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1); by (etac (AC_ball_Pi RS exE) 1); -by (resolve_tac [exI] 1); +by (rtac exI 1); (*Lemma needed in both subgoals, for a fixed z*) by (subgoal_tac "ALL z: (UN i:K. X(i)). z: X(LEAST i. z:X(i)) & (LEAST i. z:X(i)) : K" 1); by (fast_tac (ZF_cs addSIs [Least_le RS lt_trans1 RS ltD, ltI] addSEs [LeastI, Ord_in_Ord]) 2); by (res_inst_tac [("c", "%z. "), - ("d", "%. converse(f`i) ` j")] - lam_injective 1); + ("d", "%. converse(f`i) ` j")] + lam_injective 1); (*Instantiate the lemma proved above*) by (ALLGOALS ball_tac); by (fast_tac (ZF_cs addEs [inj_is_fun RS apply_type] addDs [apply_type]) 1); -by (dresolve_tac [apply_type] 1); -by (eresolve_tac [conjunct2] 1); +by (dtac apply_type 1); +by (etac conjunct2 1); by (asm_simp_tac (ZF_ss addsimps [left_inverse]) 1); qed "cardinal_UN_le"; @@ -131,7 +131,7 @@ \ |UN i:K. X(i)| < csucc(K)"; by (asm_full_simp_tac (ZF_ss addsimps [Card_lt_csucc_iff, cardinal_UN_le, - InfCard_is_Card, Card_cardinal]) 1); + InfCard_is_Card, Card_cardinal]) 1); qed "cardinal_UN_lt_csucc"; (*The same again, for a union of ordinals. In use, j(i) is a bit like rank(i), @@ -159,9 +159,9 @@ (*Work backwards along the injection from W into K, given that W~=0.*) goal Perm.thy - "!!A. [| f: inj(A,B); a:A |] ==> \ + "!!A. [| f: inj(A,B); a:A |] ==> \ \ (UN x:A. C(x)) <= (UN y:B. C(if(y: range(f), converse(f)`y, a)))"; -by (resolve_tac [UN_least] 1); +by (rtac UN_least 1); by (res_inst_tac [("x1", "f`x")] (UN_upper RSN (2,subset_trans)) 1); by (eresolve_tac [inj_is_fun RS apply_type] 2 THEN assume_tac 2); by (asm_simp_tac @@ -174,17 +174,17 @@ "!!K. [| InfCard(K); |W| le K; ALL w:W. j(w) < csucc(K) |] ==> \ \ (UN w:W. j(w)) < csucc(K)"; by (excluded_middle_tac "W=0" 1); -by (asm_simp_tac (*solve the easy 0 case*) +by (asm_simp_tac (*solve the easy 0 case*) (ZF_ss addsimps [UN_0, InfCard_is_Card, Card_is_Ord RS Card_csucc, - Card_is_Ord, Ord_0_lt_csucc]) 2); + Card_is_Ord, Ord_0_lt_csucc]) 2); by (asm_full_simp_tac (ZF_ss addsimps [InfCard_is_Card, le_Card_iff, lepoll_def]) 1); by (safe_tac eq_cs); by (swap_res_tac [[inj_UN_subset, cardinal_UN_Ord_lt_csucc] - MRS lt_subset_trans] 1); + MRS lt_subset_trans] 1); by (REPEAT (assume_tac 1)); by (fast_tac (ZF_cs addSIs [Ord_UN] addEs [ltE]) 2); by (asm_simp_tac (ZF_ss addsimps [inj_converse_fun RS apply_type] - setloop split_tac [expand_if]) 1); + setloop split_tac [expand_if]) 1); qed "le_UN_Ord_lt_csucc"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Coind/ECR.ML --- a/src/ZF/Coind/ECR.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Coind/ECR.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Coind/ECR.ML +(* Title: ZF/Coind/ECR.ML ID: $Id$ - Author: Jacob Frost, Cambridge University Computer Laboratory + Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Coind/Language.ML --- a/src/ZF/Coind/Language.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Coind/Language.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Coind/Language.ML +(* Title: ZF/Coind/Language.ML ID: $Id$ - Author: Jacob Frost, Cambridge University Computer Laboratory + Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Coind/MT.ML --- a/src/ZF/Coind/MT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Coind/MT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Coind/MT.ML +(* Title: ZF/Coind/MT.ML ID: $Id$ - Author: Jacob Frost, Cambridge University Computer Laboratory + Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge *) @@ -12,7 +12,7 @@ (* ############################################################ *) val prems = goal MT.thy - "[| c:Const; hastyenv(ve,te);:ElabRel |] ==> \ + "[| c:Const; hastyenv(ve,te);:ElabRel |] ==> \ \ : HasTyRel"; by (cut_facts_tac prems 1); by (fast_tac htr_cs 1); @@ -20,7 +20,7 @@ val prems = goalw MT.thy [hastyenv_def] - "[| x:ve_dom(ve); hastyenv(ve,te); :ElabRel |] ==> \ + "[| x:ve_dom(ve); hastyenv(ve,te); :ElabRel |] ==> \ \ :HasTyRel"; by (cut_facts_tac prems 1); by (fast_tac static_cs 1); @@ -28,9 +28,9 @@ val prems = goalw MT.thy [hastyenv_def] - "[| ve:ValEnv; x:ExVar; e:Exp; hastyenv(ve,te); \ -\ :ElabRel \ -\ |] ==> \ + "[| ve:ValEnv; x:ExVar; e:Exp; hastyenv(ve,te); \ +\ :ElabRel \ +\ |] ==> \ \ : HasTyRel"; by (cut_facts_tac prems 1); by (best_tac htr_cs 1); @@ -51,9 +51,9 @@ (ematch_tac [te_owrE] i)); val prems = goalw MT.thy [hastyenv_def] - "[| ve:ValEnv; x:ExVar; e:Exp; f:ExVar; cl:Val; \ -\ v_clos(x,e,ve_owr(ve,f,cl)) = cl; \ -\ hastyenv(ve,te); :ElabRel |] ==> \ + "[| ve:ValEnv; x:ExVar; e:Exp; f:ExVar; cl:Val; \ +\ v_clos(x,e,ve_owr(ve,f,cl)) = cl; \ +\ hastyenv(ve,te); :ElabRel |] ==> \ \ :HasTyRel"; by (cut_facts_tac prems 1); by (etac elab_fixE 1); @@ -97,15 +97,15 @@ val prems = goal MT.thy - " [| ve:ValEnv; e1:Exp; e2:Exp; c1:Const; c2:Const; \ -\ :EvalRel; \ -\ ALL t te. \ + " [| ve:ValEnv; e1:Exp; e2:Exp; c1:Const; c2:Const; \ +\ :EvalRel; \ +\ ALL t te. \ \ hastyenv(ve,te) --> :ElabRel --> :HasTyRel; \ -\ : EvalRel; \ -\ ALL t te. \ +\ : EvalRel; \ +\ ALL t te. \ \ hastyenv(ve,te) --> :ElabRel --> :HasTyRel; \ -\ hastyenv(ve, te); \ -\ :ElabRel |] ==> \ +\ hastyenv(ve, te); \ +\ :ElabRel |] ==> \ \ :HasTyRel"; by (cut_facts_tac prems 1); by (etac elab_appE 1); @@ -113,23 +113,23 @@ qed "consistency_app1"; val prems = goal MT.thy - " [| ve:ValEnv; vem:ValEnv; e1:Exp; e2:Exp; em:Exp; xm:ExVar; v:Val; \ -\ :EvalRel; \ -\ ALL t te. \ -\ hastyenv(ve,te) --> \ -\ :ElabRel --> \ -\ :HasTyRel; \ -\ :EvalRel; \ -\ ALL t te. \ -\ hastyenv(ve,te) --> \ -\ :ElabRel --> \ -\ :HasTyRel; \ -\ :EvalRel; \ -\ ALL t te. \ -\ hastyenv(ve_owr(vem,xm,v2),te) --> \ -\ :ElabRel --> \ -\ :HasTyRel; \ -\ hastyenv(ve,te); :ElabRel |] ==> \ + " [| ve:ValEnv; vem:ValEnv; e1:Exp; e2:Exp; em:Exp; xm:ExVar; v:Val; \ +\ :EvalRel; \ +\ ALL t te. \ +\ hastyenv(ve,te) --> \ +\ :ElabRel --> \ +\ :HasTyRel; \ +\ :EvalRel; \ +\ ALL t te. \ +\ hastyenv(ve,te) --> \ +\ :ElabRel --> \ +\ :HasTyRel; \ +\ :EvalRel; \ +\ ALL t te. \ +\ hastyenv(ve_owr(vem,xm,v2),te) --> \ +\ :ElabRel --> \ +\ :HasTyRel; \ +\ hastyenv(ve,te); :ElabRel |] ==> \ \ :HasTyRel "; by (cut_facts_tac prems 1); by (etac elab_appE 1); @@ -157,7 +157,7 @@ fun mt_cases_tac t = ((rtac t 1) THEN (TRYALL assume_tac)); val prems = goal MT.thy - ":EvalRel ==> \ + ":EvalRel ==> \ \ (ALL t te. hastyenv(ve,te) --> :ElabRel --> :HasTyRel)"; by (rtac (EvalRel.mutual_induct RS spec RS spec RS spec RS impE) 1); by (resolve_tac prems 7 THEN assume_tac 7); @@ -172,11 +172,11 @@ val prems = goal MT.thy - "[| ve:ValEnv; te:TyEnv; \ -\ isofenv(ve,te); \ -\ :EvalRel; \ -\ :ElabRel \ -\ |] ==> \ + "[| ve:ValEnv; te:TyEnv; \ +\ isofenv(ve,te); \ +\ :EvalRel; \ +\ :ElabRel \ +\ |] ==> \ \ isof(c,t)"; by (cut_facts_tac prems 1); by (rtac (htr_constE) 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Coind/Map.ML --- a/src/ZF/Coind/Map.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Coind/Map.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Coind/Map.ML +(* Title: ZF/Coind/Map.ML ID: $Id$ - Author: Jacob Frost, Cambridge University Computer Laboratory + Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Coind/ROOT.ML --- a/src/ZF/Coind/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Coind/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Coind/MT.ML +(* Title: ZF/Coind/MT.ML ID: $Id$ - Author: Jacob Frost, Cambridge University Computer Laboratory + Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge Based upon the article @@ -13,7 +13,7 @@ Report, Computer Lab, University of Cambridge (1995). *) -ZF_build_completed; (*Make examples fail if ZF did*) +ZF_build_completed; (*Make examples fail if ZF did*) writeln"Root file for ZF/Coind"; proof_timing := true; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Coind/Static.ML --- a/src/ZF/Coind/Static.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Coind/Static.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Coind/Static.ML +(* Title: ZF/Coind/Static.ML ID: $Id$ - Author: Jacob Frost, Cambridge University Computer Laboratory + Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Coind/Types.ML --- a/src/ZF/Coind/Types.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Coind/Types.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Coind/Types.ML +(* Title: ZF/Coind/Types.ML ID: $Id$ - Author: Jacob Frost, Cambridge University Computer Laboratory + Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Coind/Values.ML --- a/src/ZF/Coind/Values.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Coind/Values.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Coind/Values.ML +(* Title: ZF/Coind/Values.ML ID: $Id$ - Author: Jacob Frost, Cambridge University Computer Laboratory + Author: Jacob Frost, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Datatype.ML --- a/src/ZF/Datatype.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Datatype.ML Tue Jan 30 13:42:57 1996 +0100 @@ -26,9 +26,9 @@ signature INDUCTIVE_THMS = sig - val monos : thm list (*monotonicity of each M operator*) - val type_intrs : thm list (*type-checking intro rules*) - val type_elims : thm list (*type-checking elim rules*) + val monos : thm list (*monotonicity of each M operator*) + val type_intrs : thm list (*type-checking intro rules*) + val type_elims : thm list (*type-checking elim rules*) end; functor Data_section_Fun @@ -37,7 +37,7 @@ struct structure Con = Constructor_Fun - (structure Const = Arg and Pr=Standard_Prod and Su=Standard_Sum); + (structure Const = Arg and Pr=Standard_Prod and Su=Standard_Sum); structure Inductive = Ind_section_Fun (open Arg; @@ -55,7 +55,7 @@ struct structure Con = Constructor_Fun - (structure Const = Arg and Pr=Quine_Prod and Su=Quine_Sum); + (structure Const = Arg and Pr=Quine_Prod and Su=Quine_Sum); structure CoInductive = CoInd_section_Fun (open Arg; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Epsilon.ML --- a/src/ZF/Epsilon.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Epsilon.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/epsilon.ML +(* Title: ZF/epsilon.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge For epsilon.thy. Epsilon induction and recursion @@ -53,7 +53,7 @@ (*Perform epsilon-induction on i. *) fun eps_ind_tac a = EVERY' [res_inst_tac [("a",a)] eps_induct, - rename_last_tac a ["1"]]; + rename_last_tac a ["1"]]; (*** Leastness of eclose ***) @@ -77,9 +77,9 @@ (*COMPLETELY DIFFERENT induction principle from eclose_induct!!*) val [major,base,step] = goal Epsilon.thy - "[| a: eclose(b); \ -\ !!y. [| y: b |] ==> P(y); \ -\ !!y z. [| y: eclose(b); P(y); z: y |] ==> P(z) \ + "[| a: eclose(b); \ +\ !!y. [| y: b |] ==> P(y); \ +\ !!y z. [| y: eclose(b); P(y); z: y |] ==> P(z) \ \ |] ==> P(a)"; by (rtac (major RSN (3, eclose_least RS subsetD RS CollectD2)) 1); by (rtac (CollectI RS subsetI) 2); @@ -127,7 +127,7 @@ by (rtac wfrec_ssubst 1); by (rtac wfrec_ssubst 1); by (asm_simp_tac (ZF_ss addsimps [under_Memrel_eclose, - jmemi RSN (2,mem_eclose_sing_trans)]) 1); + jmemi RSN (2,mem_eclose_sing_trans)]) 1); qed "wfrec_eclose_eq"; val [prem] = goal Epsilon.thy @@ -140,7 +140,7 @@ "transrec(a,H) = H(a, lam x:a. transrec(x,H))"; by (rtac wfrec_ssubst 1); by (simp_tac (ZF_ss addsimps [wfrec_eclose_eq2, arg_in_eclose_sing, - under_Memrel_eclose]) 1); + under_Memrel_eclose]) 1); qed "transrec"; (*Avoids explosions in proofs; resolve it with a meta-level definition.*) @@ -221,9 +221,9 @@ by (rtac (rank RS trans) 1); by (rtac le_anti_sym 1); by (DO_GOAL [rtac (Ord_rank RS Ord_succ RS UN_least_le), - etac (PowD RS rank_mono RS succ_leI)] 1); + etac (PowD RS rank_mono RS succ_leI)] 1); by (DO_GOAL [rtac ([Pow_top, le_refl] MRS UN_upper_le), - REPEAT o rtac (Ord_rank RS Ord_succ)] 1); + REPEAT o rtac (Ord_rank RS Ord_succ)] 1); qed "rank_Pow"; goal Epsilon.thy "rank(0) = 0"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/EquivClass.ML --- a/src/ZF/EquivClass.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/EquivClass.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/EquivClass.ML +(* Title: ZF/EquivClass.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge For EquivClass.thy. Equivalence relations in Zermelo-Fraenkel Set Theory @@ -27,7 +27,7 @@ goalw EquivClass.thy [equiv_def] "!!A r. equiv(A,r) ==> converse(r) O r = r"; by (fast_tac (subset_cs addSIs [equalityI, sym_trans_comp_subset, - refl_comp_subset]) 1); + refl_comp_subset]) 1); qed "equiv_comp_eq"; (*second half*) @@ -38,7 +38,7 @@ by (safe_tac ZF_cs); by (fast_tac (ZF_cs addSIs [converseI] addIs [compI]) 3); by (ALLGOALS (fast_tac - (ZF_cs addSIs [converseI] addIs [compI] addSEs [compE]))); + (ZF_cs addSIs [converseI] addIs [compI] addSEs [compE]))); qed "comp_equivI"; (** Equivalence classes **) @@ -52,7 +52,7 @@ goalw EquivClass.thy [equiv_def] "!!A r. [| equiv(A,r); : r |] ==> r``{a} = r``{b}"; by (safe_tac (subset_cs addSIs [equalityI, equiv_class_subset])); -by (rewrite_goals_tac [sym_def]); +by (rewtac sym_def); by (fast_tac ZF_cs 1); qed "equiv_class_eq"; @@ -85,13 +85,13 @@ goal EquivClass.thy "!!A r. equiv(A,r) ==> : r <-> r``{x} = r``{y} & x:A & y:A"; by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq] - addDs [equiv_type]) 1); + addDs [equiv_type]) 1); qed "equiv_class_eq_iff"; goal EquivClass.thy "!!A r. [| equiv(A,r); x: A; y: A |] ==> r``{x} = r``{y} <-> : r"; by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq] - addDs [equiv_type]) 1); + addDs [equiv_type]) 1); qed "eq_equiv_class_iff"; (*** Quotients ***) @@ -104,7 +104,7 @@ qed "quotientI"; val major::prems = goalw EquivClass.thy [quotient_def] - "[| X: A/r; !!x. [| X = r``{x}; x:A |] ==> P |] \ + "[| X: A/r; !!x. [| X = r``{x}; x:A |] ==> P |] \ \ ==> P"; by (rtac (major RS RepFunE) 1); by (eresolve_tac prems 1); @@ -143,8 +143,8 @@ (*type checking of UN x:r``{a}. b(x) *) val prems = goalw EquivClass.thy [quotient_def] - "[| equiv(A,r); congruent(r,b); X: A/r; \ -\ !!x. x : A ==> b(x) : B |] \ + "[| equiv(A,r); congruent(r,b); X: A/r; \ +\ !!x. x : A ==> b(x) : B |] \ \ ==> (UN x:X. b(x)) : B"; by (cut_facts_tac prems 1); by (safe_tac ZF_cs); @@ -157,7 +157,7 @@ val prems = goalw EquivClass.thy [quotient_def] "[| equiv(A,r); congruent(r,b); \ \ (UN x:X. b(x))=(UN y:Y. b(y)); X: A/r; Y: A/r; \ -\ !!x y. [| x:A; y:A; b(x)=b(y) |] ==> :r |] \ +\ !!x y. [| x:A; y:A; b(x)=b(y) |] ==> :r |] \ \ ==> X=Y"; by (cut_facts_tac prems 1); by (safe_tac ZF_cs); @@ -184,7 +184,7 @@ by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1); by (assume_tac 1); by (asm_simp_tac (ZF_ss addsimps [equivA RS UN_equiv_class, - congruent2_implies_congruent]) 1); + congruent2_implies_congruent]) 1); by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]); by (fast_tac ZF_cs 1); qed "congruent2_implies_congruent_UN"; @@ -194,28 +194,28 @@ \ ==> (UN x1:r``{a1}. UN x2:r``{a2}. b(x1,x2)) = b(a1,a2)"; by (cut_facts_tac prems 1); by (asm_simp_tac (ZF_ss addsimps [equivA RS UN_equiv_class, - congruent2_implies_congruent, - congruent2_implies_congruent_UN]) 1); + congruent2_implies_congruent, + congruent2_implies_congruent_UN]) 1); qed "UN_equiv_class2"; (*type checking*) val prems = goalw EquivClass.thy [quotient_def] - "[| equiv(A,r); congruent2(r,b); \ -\ X1: A/r; X2: A/r; \ -\ !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) : B \ + "[| equiv(A,r); congruent2(r,b); \ +\ X1: A/r; X2: A/r; \ +\ !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) : B \ \ |] ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B"; by (cut_facts_tac prems 1); by (safe_tac ZF_cs); by (REPEAT (ares_tac (prems@[UN_equiv_class_type, - congruent2_implies_congruent_UN, - congruent2_implies_congruent, quotientI]) 1)); + congruent2_implies_congruent_UN, + congruent2_implies_congruent, quotientI]) 1)); qed "UN_equiv_class_type2"; (*Suggested by John Harrison -- the two subproofs may be MUCH simpler than the direct proof*) val prems = goalw EquivClass.thy [congruent2_def,equiv_def,refl_def] - "[| equiv(A,r); \ + "[| equiv(A,r); \ \ !! y z w. [| w: A; : r |] ==> b(y,w) = b(z,w); \ \ !! y z w. [| w: A; : r |] ==> b(w,y) = b(w,z) \ \ |] ==> congruent2(r,b)"; @@ -227,9 +227,9 @@ qed "congruent2I"; val [equivA,commute,congt] = goal EquivClass.thy - "[| equiv(A,r); \ + "[| equiv(A,r); \ \ !! y z. [| y: A; z: A |] ==> b(y,z) = b(z,y); \ -\ !! y z w. [| w: A; : r |] ==> b(w,y) = b(w,z) \ +\ !! y z w. [| w: A; : r |] ==> b(w,y) = b(w,z) \ \ |] ==> congruent2(r,b)"; by (resolve_tac [equivA RS congruent2I] 1); by (rtac (commute RS trans) 1); @@ -242,8 +242,8 @@ (*Obsolete?*) val [equivA,ZinA,congt,commute] = goalw EquivClass.thy [quotient_def] "[| equiv(A,r); Z: A/r; \ -\ !!w. [| w: A |] ==> congruent(r, %z.b(w,z)); \ -\ !!x y. [| x: A; y: A |] ==> b(y,x) = b(x,y) \ +\ !!w. [| w: A |] ==> congruent(r, %z.b(w,z)); \ +\ !!x y. [| x: A; y: A |] ==> b(y,x) = b(x,y) \ \ |] ==> congruent(r, %w. UN z: Z. b(w,z))"; val congt' = rewrite_rule [congruent_def] congt; by (cut_facts_tac [ZinA] 1); @@ -252,6 +252,6 @@ by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1); by (assume_tac 1); by (asm_simp_tac (ZF_ss addsimps [commute, - [equivA, congt] MRS UN_equiv_class]) 1); + [equivA, congt] MRS UN_equiv_class]) 1); by (REPEAT (ares_tac [congt' RS spec RS spec RS mp] 1)); qed "congruent_commuteI"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Finite.ML --- a/src/ZF/Finite.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Finite.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Finite.ML +(* Title: ZF/Finite.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Finite powerset operator; finite function space @@ -33,7 +33,7 @@ \ |] ==> P(b)"; by (rtac (major RS Fin.induct) 1); by (excluded_middle_tac "a:b" 2); -by (etac (cons_absorb RS ssubst) 3 THEN assume_tac 3); (*backtracking!*) +by (etac (cons_absorb RS ssubst) 3 THEN assume_tac 3); (*backtracking!*) by (REPEAT (ares_tac prems 1)); qed "Fin_induct"; @@ -68,19 +68,19 @@ qed "Fin_subset"; val major::prems = goal Finite.thy - "[| c: Fin(A); b: Fin(A); \ -\ P(b); \ + "[| c: Fin(A); b: Fin(A); \ +\ P(b); \ \ !!x y. [| x: A; y: Fin(A); x:y; P(y) |] ==> P(y-{x}) \ \ |] ==> c<=b --> P(b-c)"; by (rtac (major RS Fin_induct) 1); by (rtac (Diff_cons RS ssubst) 2); by (ALLGOALS (asm_simp_tac (Fin_ss addsimps (prems@[Diff_0, cons_subset_iff, - Diff_subset RS Fin_subset])))); + Diff_subset RS Fin_subset])))); qed "Fin_0_induct_lemma"; val prems = goal Finite.thy - "[| b: Fin(A); \ -\ P(b); \ + "[| b: Fin(A); \ +\ P(b); \ \ !!x y. [| x: A; y: Fin(A); x:y; P(y) |] ==> P(y-{x}) \ \ |] ==> P(0)"; by (rtac (Diff_cancel RS subst) 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Fixedpt.ML --- a/src/ZF/Fixedpt.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Fixedpt.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/fixedpt.ML +(* Title: ZF/fixedpt.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge For fixedpt.thy. Least and greatest fixed points; the Knaster-Tarski Theorem @@ -129,8 +129,8 @@ (*This rule yields an induction hypothesis in which the components of a data structure may be assumed to be elements of lfp(D,h)*) val prems = goal Fixedpt.thy - "[| bnd_mono(D,h); a : lfp(D,h); \ -\ !!x. x : h(Collect(lfp(D,h),P)) ==> P(x) \ + "[| bnd_mono(D,h); a : lfp(D,h); \ +\ !!x. x : h(Collect(lfp(D,h),P)) ==> P(x) \ \ |] ==> P(a)"; by (rtac (Collect_is_pre_fixedpt RS lfp_lowerbound RS subsetD RS CollectD2) 1); by (rtac (lfp_subset RS (Collect_subset RS subset_trans)) 3); @@ -158,7 +158,7 @@ (*Monotonicity of lfp, where h precedes i under a domain-like partial order monotonicity of h is not strictly necessary; h must be bounded by D*) val [hmono,imono,subhi] = goal Fixedpt.thy - "[| bnd_mono(D,h); bnd_mono(E,i); \ + "[| bnd_mono(D,h); bnd_mono(E,i); \ \ !!X. X<=D ==> h(X) <= i(X) |] ==> lfp(D,h) <= lfp(E,i)"; by (rtac (bnd_monoD1 RS lfp_greatest) 1); by (rtac imono 1); @@ -294,7 +294,7 @@ (*Monotonicity of gfp!*) val [hmono,subde,subhi] = goal Fixedpt.thy - "[| bnd_mono(D,h); D <= E; \ + "[| bnd_mono(D,h); D <= E; \ \ !!X. X<=D ==> h(X) <= i(X) |] ==> gfp(D,h) <= gfp(E,i)"; by (rtac gfp_upperbound 1); by (rtac (hmono RS gfp_lemma2 RS subset_trans) 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/IMP/Com.ML --- a/src/ZF/IMP/Com.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/IMP/Com.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,13 +1,13 @@ -(* Title: ZF/IMP/Com.ML +(* Title: ZF/IMP/Com.ML ID: $Id$ - Author: Heiko Loetzbeyer & Robert Sandner, TUM + Author: Heiko Loetzbeyer & Robert Sandner, TUM Copyright 1994 TUM *) open Com; val assign_type = prove_goalw Com.thy [assign_def] - "!!n. [| sigma:loc -> nat; n:nat |] ==> sigma[n/x] : loc -> nat" + "!!n. [| sigma:loc -> nat; n:nat |] ==> sigma[n/x] : loc -> nat" (fn _ => [ fast_tac (ZF_cs addIs [apply_type,lam_type,if_type]) 1 ]); val type_cs = ZF_cs addSDs [evala.dom_subset RS subsetD, diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/IMP/Denotation.ML --- a/src/ZF/IMP/Denotation.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/IMP/Denotation.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/IMP/Denotation.ML +(* Title: ZF/IMP/Denotation.ML ID: $Id$ - Author: Heiko Loetzbeyer & Robert Sandner, TUM + Author: Heiko Loetzbeyer & Robert Sandner, TUM Copyright 1994 TUM *) @@ -20,7 +20,7 @@ (**** Type_intr for A ****) val A_type = prove_goal Denotation.thy - "!!a.[|a:aexp; sigma:loc->nat|] ==> A(a,sigma):nat" + "!!a.[|a:aexp; sigma:loc->nat|] ==> A(a,sigma):nat" (fn _ => [(etac aexp.induct 1), (rewrite_goals_tac A_rewrite_rules), (ALLGOALS (fast_tac (ZF_cs addSIs [apply_type])))]); @@ -28,7 +28,7 @@ (**** Type_intr for B ****) val B_type = prove_goal Denotation.thy - "!!b. [|b:bexp; sigma:loc->nat|] ==> B(b,sigma):bool" + "!!b. [|b:bexp; sigma:loc->nat|] ==> B(b,sigma):bool" (fn _ => [(etac bexp.induct 1), (rewrite_goals_tac B_rewrite_rules), (ALLGOALS (fast_tac (ZF_cs @@ -37,7 +37,7 @@ (**** C_subset ****) val C_subset = prove_goal Denotation.thy - "!!c. c:com ==> C(c) <= (loc->nat)*(loc->nat)" + "!!c. c:com ==> C(c) <= (loc->nat)*(loc->nat)" (fn _ => [(etac com.induct 1), (rewrite_tac C_rewrite_rules), (ALLGOALS (fast_tac (comp_cs addDs [lfp_subset RS subsetD])))]); @@ -45,17 +45,17 @@ (**** Type_elims for C ****) val C_type = prove_goal Denotation.thy - "[| :C(c); c:com; \ -\ !!c. [| x:loc->nat; y:loc->nat |] ==> R |] \ -\ ==> R" + "[| :C(c); c:com; \ +\ !!c. [| x:loc->nat; y:loc->nat |] ==> R |] \ +\ ==> R" (fn prems => [(cut_facts_tac prems 1), (fast_tac (ZF_cs addSIs prems addDs [(C_subset RS subsetD)]) 1)]); val C_type_fst = prove_goal Denotation.thy - "[| x:C(c); c:com; \ -\ !!c. [| fst(x):loc->nat |] ==> R |] \ -\ ==> R" + "[| x:C(c); c:com; \ +\ !!c. [| fst(x):loc->nat |] ==> R |] \ +\ ==> R" (fn prems => [(cut_facts_tac prems 1), (resolve_tac prems 1), (dtac (C_subset RS subsetD) 1), @@ -67,7 +67,7 @@ (**** bnd_mono (nat->nat*nat->nat,Gamma(b,c) ****) val Gamma_bnd_mono = prove_goalw Denotation.thy [bnd_mono_def,Gamma_def] - "!!c. c:com ==> bnd_mono ((loc->nat)*(loc->nat),Gamma(b,c))" + "!!c. c:com ==> bnd_mono ((loc->nat)*(loc->nat),Gamma(b,c))" (fn prems => [(best_tac (comp_cs addEs [C_type]) 1)]); (**** End ***) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/IMP/Equiv.ML --- a/src/ZF/IMP/Equiv.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/IMP/Equiv.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/IMP/Equiv.ML +(* Title: ZF/IMP/Equiv.ML ID: $Id$ - Author: Heiko Loetzbeyer & Robert Sandner, TUM + Author: Heiko Loetzbeyer & Robert Sandner, TUM Copyright 1994 TUM *) @@ -9,7 +9,7 @@ by (res_inst_tac [("x","n")] spec 1); (* quantify n *) by (res_inst_tac [("x","a")] aexp.induct 1); (* struct. ind. *) by (resolve_tac prems 1); (* type prem. *) -by (rewrite_goals_tac A_rewrite_rules); (* rewr. Den. *) +by (rewrite_goals_tac A_rewrite_rules); (* rewr. Den. *) by (TRYALL (fast_tac (ZF_cs addSIs (evala.intrs@prems) addSEs aexp_elim_cases))); qed "aexp_iff"; @@ -17,7 +17,7 @@ val aexp1 = prove_goal Equiv.thy "[| -a-> n; a: aexp; sigma: loc -> nat |] \ - \ ==> A(a,sigma) = n" (* destruction rule *) + \ ==> A(a,sigma) = n" (* destruction rule *) (fn prems => [(fast_tac (ZF_cs addSIs ((aexp_iff RS iffD1)::prems)) 1)]); val aexp2 = aexp_iff RS iffD2; @@ -35,10 +35,10 @@ val prems = goal Equiv.thy "[| b: bexp; sigma: loc -> nat |] ==> \ \ -b-> w <-> B(b,sigma) = w"; -by (res_inst_tac [("x","w")] spec 1); (* quantify w *) -by (res_inst_tac [("x","b")] bexp.induct 1); (* struct. ind. *) -by (resolve_tac prems 1); (* type prem. *) -by (rewrite_goals_tac B_rewrite_rules); (* rewr. Den. *) +by (res_inst_tac [("x","w")] spec 1); (* quantify w *) +by (res_inst_tac [("x","b")] bexp.induct 1); (* struct. ind. *) +by (resolve_tac prems 1); (* type prem. *) +by (rewrite_goals_tac B_rewrite_rules); (* rewr. Den. *) by (TRYALL (fast_tac (ZF_cs addSIs (evalb.intrs@prems@[aexp2]) addSDs [aexp1] addSEs bexp_elim_cases))); qed "bexp_iff"; @@ -70,12 +70,12 @@ (* while *) by (etac (rewrite_rule [Gamma_def] - (Gamma_bnd_mono RS lfp_Tarski RS ssubst)) 1); + (Gamma_bnd_mono RS lfp_Tarski RS ssubst)) 1); by (asm_simp_tac (ZF_ss addsimps [bexp1]) 1); by (fast_tac (comp_cs addSIs [bexp1,idI]@evalb_type_intrs) 1); by (etac (rewrite_rule [Gamma_def] - (Gamma_bnd_mono RS lfp_Tarski RS ssubst)) 1); + (Gamma_bnd_mono RS lfp_Tarski RS ssubst)) 1); by (asm_simp_tac (ZF_ss addsimps [bexp1]) 1); by (fast_tac (comp_cs addSIs [bexp1,compI]@evalb_type_intrs) 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/IMP/ROOT.ML --- a/src/ZF/IMP/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/IMP/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/IMP/ROOT.ML +(* Title: ZF/IMP/ROOT.ML ID: $Id$ - Author: Heiko Loetzbeyer & Robert Sandner, TUM + Author: Heiko Loetzbeyer & Robert Sandner, TUM Copyright 1994 TUM Executes the formalization of the denotational and operational semantics of a @@ -12,7 +12,7 @@ *) -ZF_build_completed; (*Make examples fail if ZF did*) +ZF_build_completed; (*Make examples fail if ZF did*) writeln"Root file for ZF/IMP"; proof_timing := true; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Inductive.ML --- a/src/ZF/Inductive.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Inductive.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Inductive.ML +(* Title: ZF/Inductive.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge (Co)Inductive Definitions for Zermelo-Fraenkel Set Theory @@ -17,38 +17,38 @@ structure Lfp = struct - val oper = Const("lfp", [iT,iT-->iT]--->iT) - val bnd_mono = Const("bnd_mono", [iT,iT-->iT]--->oT) - val bnd_monoI = bnd_monoI - val subs = def_lfp_subset - val Tarski = def_lfp_Tarski - val induct = def_induct + val oper = Const("lfp", [iT,iT-->iT]--->iT) + val bnd_mono = Const("bnd_mono", [iT,iT-->iT]--->oT) + val bnd_monoI = bnd_monoI + val subs = def_lfp_subset + val Tarski = def_lfp_Tarski + val induct = def_induct end; structure Standard_Prod = struct - val sigma = Const("Sigma", [iT, iT-->iT]--->iT) - val pair = Const("Pair", [iT,iT]--->iT) - val split_const = Const("split", [[iT,iT]--->iT, iT]--->iT) - val fsplit_const = Const("split", [[iT,iT]--->oT, iT]--->oT) - val pair_iff = Pair_iff - val split_eq = split - val fsplitI = splitI - val fsplitD = splitD - val fsplitE = splitE + val sigma = Const("Sigma", [iT, iT-->iT]--->iT) + val pair = Const("Pair", [iT,iT]--->iT) + val split_const = Const("split", [[iT,iT]--->iT, iT]--->iT) + val fsplit_const = Const("split", [[iT,iT]--->oT, iT]--->oT) + val pair_iff = Pair_iff + val split_eq = split + val fsplitI = splitI + val fsplitD = splitD + val fsplitE = splitE end; structure Standard_Sum = struct - val sum = Const("op +", [iT,iT]--->iT) - val inl = Const("Inl", iT-->iT) - val inr = Const("Inr", iT-->iT) - val elim = Const("case", [iT-->iT, iT-->iT, iT]--->iT) - val case_inl = case_Inl - val case_inr = case_Inr - val inl_iff = Inl_iff - val inr_iff = Inr_iff - val distinct = Inl_Inr_iff + val sum = Const("op +", [iT,iT]--->iT) + val inl = Const("Inl", iT-->iT) + val inr = Const("Inr", iT-->iT) + val elim = Const("case", [iT-->iT, iT-->iT, iT]--->iT) + val case_inl = case_Inl + val case_inr = case_Inr + val inl_iff = Inl_iff + val inr_iff = Inr_iff + val distinct = Inl_Inr_iff val distinct' = Inr_Inl_iff val free_SEs = Ind_Syntax.mk_free_SEs [distinct, distinct', inl_iff, inr_iff, Standard_Prod.pair_iff] @@ -60,21 +60,21 @@ let structure Intr_elim = Intr_elim_Fun(structure Inductive=Inductive and Fp=Lfp and - Pr=Standard_Prod and Su=Standard_Sum); + Pr=Standard_Prod and Su=Standard_Sum); structure Indrule = Indrule_Fun (structure Inductive=Inductive and - Pr=Standard_Prod and Su=Standard_Sum and - Intr_elim=Intr_elim) + Pr=Standard_Prod and Su=Standard_Sum and + Intr_elim=Intr_elim) in struct - val thy = Intr_elim.thy - val defs = Intr_elim.defs - val bnd_mono = Intr_elim.bnd_mono + val thy = Intr_elim.thy + val defs = Intr_elim.defs + val bnd_mono = Intr_elim.bnd_mono val dom_subset = Intr_elim.dom_subset - val intrs = Intr_elim.intrs - val elim = Intr_elim.elim - val mk_cases = Intr_elim.mk_cases + val intrs = Intr_elim.intrs + val elim = Intr_elim.elim + val mk_cases = Intr_elim.mk_cases open Indrule end end; @@ -86,38 +86,38 @@ structure Gfp = struct - val oper = Const("gfp", [iT,iT-->iT]--->iT) - val bnd_mono = Const("bnd_mono", [iT,iT-->iT]--->oT) - val bnd_monoI = bnd_monoI - val subs = def_gfp_subset - val Tarski = def_gfp_Tarski - val induct = def_Collect_coinduct + val oper = Const("gfp", [iT,iT-->iT]--->iT) + val bnd_mono = Const("bnd_mono", [iT,iT-->iT]--->oT) + val bnd_monoI = bnd_monoI + val subs = def_gfp_subset + val Tarski = def_gfp_Tarski + val induct = def_Collect_coinduct end; structure Quine_Prod = struct - val sigma = Const("QSigma", [iT, iT-->iT]--->iT) - val pair = Const("QPair", [iT,iT]--->iT) - val split_const = Const("qsplit", [[iT,iT]--->iT, iT]--->iT) - val fsplit_const = Const("qsplit", [[iT,iT]--->oT, iT]--->oT) - val pair_iff = QPair_iff - val split_eq = qsplit - val fsplitI = qsplitI - val fsplitD = qsplitD - val fsplitE = qsplitE + val sigma = Const("QSigma", [iT, iT-->iT]--->iT) + val pair = Const("QPair", [iT,iT]--->iT) + val split_const = Const("qsplit", [[iT,iT]--->iT, iT]--->iT) + val fsplit_const = Const("qsplit", [[iT,iT]--->oT, iT]--->oT) + val pair_iff = QPair_iff + val split_eq = qsplit + val fsplitI = qsplitI + val fsplitD = qsplitD + val fsplitE = qsplitE end; structure Quine_Sum = struct - val sum = Const("op <+>", [iT,iT]--->iT) - val inl = Const("QInl", iT-->iT) - val inr = Const("QInr", iT-->iT) - val elim = Const("qcase", [iT-->iT, iT-->iT, iT]--->iT) - val case_inl = qcase_QInl - val case_inr = qcase_QInr - val inl_iff = QInl_iff - val inr_iff = QInr_iff - val distinct = QInl_QInr_iff + val sum = Const("op <+>", [iT,iT]--->iT) + val inl = Const("QInl", iT-->iT) + val inr = Const("QInr", iT-->iT) + val elim = Const("qcase", [iT-->iT, iT-->iT, iT]--->iT) + val case_inl = qcase_QInl + val case_inr = qcase_QInr + val inl_iff = QInl_iff + val inr_iff = QInr_iff + val distinct = QInl_QInr_iff val distinct' = QInr_QInl_iff val free_SEs = Ind_Syntax.mk_free_SEs [distinct, distinct', inl_iff, inr_iff, Quine_Prod.pair_iff] @@ -136,16 +136,16 @@ let structure Intr_elim = Intr_elim_Fun(structure Inductive=Inductive and Fp=Gfp and - Pr=Quine_Prod and Su=Quine_Sum); + Pr=Quine_Prod and Su=Quine_Sum); in struct - val thy = Intr_elim.thy - val defs = Intr_elim.defs - val bnd_mono = Intr_elim.bnd_mono + val thy = Intr_elim.thy + val defs = Intr_elim.defs + val bnd_mono = Intr_elim.bnd_mono val dom_subset = Intr_elim.dom_subset - val intrs = Intr_elim.intrs - val elim = Intr_elim.elim - val mk_cases = Intr_elim.mk_cases + val intrs = Intr_elim.intrs + val elim = Intr_elim.elim + val mk_cases = Intr_elim.mk_cases val coinduct = Intr_elim.raw_induct end end; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/InfDatatype.ML --- a/src/ZF/InfDatatype.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/InfDatatype.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/InfDatatype.ML +(* Title: ZF/InfDatatype.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Infinite-branching datatype definitions @@ -11,17 +11,17 @@ |> standard; goal InfDatatype.thy - "!!K. [| f: W -> Vfrom(A,csucc(K)); |W| le K; InfCard(K) \ + "!!K. [| f: W -> Vfrom(A,csucc(K)); |W| le K; InfCard(K) \ \ |] ==> EX j. f: W -> Vfrom(A,j) & j < csucc(K)"; by (res_inst_tac [("x", "UN w:W. LEAST i. f`w : Vfrom(A,i)")] exI 1); -by (resolve_tac [conjI] 1); -by (resolve_tac [le_UN_Ord_lt_csucc] 2); +by (rtac conjI 1); +by (rtac le_UN_Ord_lt_csucc 2); by (rtac ballI 4 THEN - eresolve_tac [fun_Limit_VfromE] 4 THEN REPEAT_SOME assume_tac); + etac fun_Limit_VfromE 4 THEN REPEAT_SOME assume_tac); by (fast_tac (ZF_cs addEs [Least_le RS lt_trans1, ltE]) 2); -by (resolve_tac [Pi_type] 1); +by (rtac Pi_type 1); by (rename_tac "w" 2); -by (eresolve_tac [fun_Limit_VfromE] 2 THEN REPEAT_SOME assume_tac); +by (etac fun_Limit_VfromE 2 THEN REPEAT_SOME assume_tac); by (subgoal_tac "f`w : Vfrom(A, LEAST i. f`w : Vfrom(A,i))" 1); by (fast_tac (ZF_cs addEs [LeastI, ltE]) 2); by (eresolve_tac [[subset_refl, UN_upper] MRS Vfrom_mono RS subsetD] 1); @@ -29,7 +29,7 @@ qed "fun_Vcsucc_lemma"; goal InfDatatype.thy - "!!K. [| W <= Vfrom(A,csucc(K)); |W| le K; InfCard(K) \ + "!!K. [| W <= Vfrom(A,csucc(K)); |W| le K; InfCard(K) \ \ |] ==> EX j. W <= Vfrom(A,j) & j < csucc(K)"; by (asm_full_simp_tac (ZF_ss addsimps [subset_iff_id, fun_Vcsucc_lemma]) 1); qed "subset_Vcsucc"; @@ -40,18 +40,18 @@ \ W -> Vfrom(A,csucc(K)) <= Vfrom(A,csucc(K))"; by (safe_tac (ZF_cs addSDs [fun_Vcsucc_lemma, subset_Vcsucc])); by (resolve_tac [Vfrom RS ssubst] 1); -by (dresolve_tac [fun_is_rel] 1); +by (dtac fun_is_rel 1); (*This level includes the function, and is below csucc(K)*) by (res_inst_tac [("a1", "succ(succ(j Un ja))")] (UN_I RS UnI2) 1); by (eresolve_tac [subset_trans RS PowI] 2); by (fast_tac (ZF_cs addIs [Pair_in_Vfrom, Vfrom_UnI1, Vfrom_UnI2]) 2); by (REPEAT (ares_tac [ltD, InfCard_csucc, InfCard_is_Limit, - Limit_has_succ, Un_least_lt] 1)); + Limit_has_succ, Un_least_lt] 1)); qed "fun_Vcsucc"; goal InfDatatype.thy - "!!K. [| f: W -> Vfrom(A, csucc(K)); |W| le K; InfCard(K); \ -\ W <= Vfrom(A,csucc(K)) \ + "!!K. [| f: W -> Vfrom(A, csucc(K)); |W| le K; InfCard(K); \ +\ W <= Vfrom(A,csucc(K)) \ \ |] ==> f: Vfrom(A,csucc(K))"; by (REPEAT (ares_tac [fun_Vcsucc RS subsetD] 1)); qed "fun_in_Vcsucc"; @@ -65,8 +65,8 @@ "!!K. InfCard(K) ==> K -> Vfrom(A,csucc(K)) <= Vfrom(A,csucc(K))"; by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1); by (REPEAT (ares_tac [fun_Vcsucc, Ord_cardinal_le, - i_subset_Vfrom, - lt_csucc RS leI RS le_imp_subset RS subset_trans] 1)); + i_subset_Vfrom, + lt_csucc RS leI RS le_imp_subset RS subset_trans] 1)); qed "Card_fun_Vcsucc"; goal InfDatatype.thy @@ -89,13 +89,13 @@ (*For handling Cardinals of the form (nat Un |X|) *) bind_thm ("InfCard_nat_Un_cardinal", - [InfCard_nat, Card_cardinal] MRS InfCard_Un); + [InfCard_nat, Card_cardinal] MRS InfCard_Un); bind_thm ("le_nat_Un_cardinal", - [Ord_nat, Card_cardinal RS Card_is_Ord] MRS Un_upper2_le); + [Ord_nat, Card_cardinal RS Card_is_Ord] MRS Un_upper2_le); bind_thm ("UN_upper_cardinal", - UN_upper RS subset_imp_lepoll RS lepoll_imp_Card_le); + UN_upper RS subset_imp_lepoll RS lepoll_imp_Card_le); (*For most K-branching datatypes with domain Vfrom(A, csucc(K)) *) val inf_datatype_intrs = diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Let.ML --- a/src/ZF/Let.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Let.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Let +(* Title: ZF/Let ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge Let expressions -- borrowed from HOL @@ -10,5 +10,5 @@ val [prem] = goalw thy [Let_def] "(!!x. x=t ==> P(u(x))) ==> P(let x=t in u(x))"; -br (refl RS prem) 1; +by (rtac (refl RS prem) 1); qed "LetI"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/List.ML --- a/src/ZF/List.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/List.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/List.ML +(* Title: ZF/List.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Datatype definition of Lists @@ -20,8 +20,8 @@ (*Perform induction on l, then prove the major premise using prems. *) fun list_ind_tac a prems i = EVERY [res_inst_tac [("x",a)] list.induct i, - rename_last_tac a ["1"] (i+2), - ares_tac prems i]; + rename_last_tac a ["1"] (i+2), + ares_tac prems i]; goal List.thy "list(A) = {0} + (A * list(A))"; let open list; val rew = rewrite_rule con_defs in @@ -43,7 +43,7 @@ by (rtac lfp_lowerbound 1); by (rtac (A_subset_univ RS univ_mono) 2); by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, - Pair_in_univ]) 1); + Pair_in_univ]) 1); qed "list_univ"; (*These two theorems justify datatypes involving list(nat), list(A), ...*) @@ -138,7 +138,7 @@ \ |] ==> list_rec(l,c,h) : C(l)"; by (list_ind_tac "l" prems 1); by (ALLGOALS (asm_simp_tac - (ZF_ss addsimps (prems@[list_rec_Nil,list_rec_Cons])))); + (ZF_ss addsimps (prems@[list_rec_Nil,list_rec_Cons])))); qed "list_rec_type"; (** Versions for use with definitions **) @@ -156,7 +156,7 @@ qed "def_list_rec_Cons"; fun list_recs def = map standard - ([def] RL [def_list_rec_Nil, def_list_rec_Cons]); + ([def] RL [def_list_rec_Nil, def_list_rec_Cons]); (** map **) @@ -229,9 +229,9 @@ val list_ss = arith_ss addsimps list.case_eqns addsimps [list_rec_Nil, list_rec_Cons, - map_Nil, map_Cons, app_Nil, app_Cons, - length_Nil, length_Cons, rev_Nil, rev_Cons, - flat_Nil, flat_Cons, list_add_Nil, list_add_Cons] + map_Nil, map_Cons, app_Nil, app_Cons, + length_Nil, length_Cons, rev_Nil, rev_Cons, + flat_Nil, flat_Cons, list_add_Nil, list_add_Cons] setsolver (type_auto_tac list_typechecks); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Nat.ML --- a/src/ZF/Nat.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Nat.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/nat.ML +(* Title: ZF/nat.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge For nat.thy. Natural numbers in Zermelo-Fraenkel Set Theory @@ -41,7 +41,7 @@ goal Nat.thy "bool <= nat"; by (REPEAT (ares_tac [subsetI,nat_0I,nat_1I] 1 - ORELSE eresolve_tac [boolE,ssubst] 1)); + ORELSE eresolve_tac [boolE,ssubst] 1)); qed "bool_subset_nat"; val bool_into_nat = bool_subset_nat RS subsetD; @@ -59,8 +59,8 @@ (*Perform induction on n, then prove the n:nat subgoal using prems. *) fun nat_ind_tac a prems i = EVERY [res_inst_tac [("n",a)] nat_induct i, - rename_last_tac a ["1"] (i+2), - ares_tac prems i]; + rename_last_tac a ["1"] (i+2), + ares_tac prems i]; val major::prems = goal Nat.thy "[| n: nat; n=0 ==> P; !!x. [| x: nat; n=succ(x) |] ==> P |] ==> P"; @@ -94,12 +94,12 @@ qed "Limit_nat"; goal Nat.thy "!!i. Limit(i) ==> nat le i"; -by (resolve_tac [subset_imp_le] 1); +by (rtac subset_imp_le 1); by (rtac subsetI 1); -by (eresolve_tac [nat_induct] 1); +by (etac nat_induct 1); by (fast_tac (ZF_cs addIs [Limit_has_succ RS ltD, ltI, Limit_is_Ord]) 2); by (REPEAT (ares_tac [Limit_has_0 RS ltD, - Ord_nat, Limit_is_Ord] 1)); + Ord_nat, Limit_is_Ord] 1)); qed "nat_le_Limit"; (* succ(i): nat ==> i: nat *) @@ -163,13 +163,13 @@ by (etac nat_induct 1); by (ALLGOALS (EVERY' [rtac (impI RS impI), rtac (nat_induct RS ballI), assume_tac, - fast_tac lt_cs, fast_tac lt_cs])); + fast_tac lt_cs, fast_tac lt_cs])); qed "succ_lt_induct_lemma"; val prems = goal Nat.thy - "[| m P(m,succ(x)) \ + "[| m P(m,succ(x)) \ \ |] ==> P(m,n)"; by (res_inst_tac [("P4","P")] (succ_lt_induct_lemma RS mp RS mp RS bspec RS mp) 1); @@ -209,7 +209,7 @@ "m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))"; by (rtac nat_rec_trans 1); by (simp_tac (ZF_ss addsimps [prem, nat_case_succ, nat_succI, Memrel_iff, - vimage_singleton_iff]) 1); + vimage_singleton_iff]) 1); qed "nat_rec_succ"; (** The union of two natural numbers is a natural number -- their maximum **) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Order.ML --- a/src/ZF/Order.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Order.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,12 +1,12 @@ -(* Title: ZF/Order.ML +(* Title: ZF/Order.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Orders in Zermelo-Fraenkel Set Theory Results from the book "Set Theory: an Introduction to Independence Proofs" - by Ken Kunen. Chapter 1, section 6. + by Ken Kunen. Chapter 1, section 6. *) open Order; @@ -34,12 +34,12 @@ (*Does the case analysis, deleting linear(A,r) and proving trivial subgoals*) val linear_case_tac = SELECT_GOAL (EVERY [etac linearE 1, assume_tac 1, assume_tac 1, - REPEAT_SOME (assume_tac ORELSE' contr_tac)]); + REPEAT_SOME (assume_tac ORELSE' contr_tac)]); (** General properties of well_ord **) goalw Order.thy [irrefl_def, part_ord_def, tot_ord_def, - trans_on_def, well_ord_def] + trans_on_def, well_ord_def] "!!r. [| wf[A](r); linear(A,r) |] ==> well_ord(A,r)"; by (asm_simp_tac (ZF_ss addsimps [wf_on_not_refl]) 1); by (fast_tac (ZF_cs addEs [linearE, wf_on_asym, wf_on_chain3]) 1); @@ -195,17 +195,17 @@ by (linear_case_tac 1); by (REPEAT (EVERY [eresolve_tac [wf_on_not_refl RS notE] 1, - eresolve_tac [ssubst] 2, - fast_tac ZF_cs 2, - REPEAT (ares_tac [apply_type] 1)])); + etac ssubst 2, + fast_tac ZF_cs 2, + REPEAT (ares_tac [apply_type] 1)])); qed "mono_map_is_inj"; (** Order-isomorphisms -- or similarities, as Suppes calls them **) val prems = goalw Order.thy [ord_iso_def] - "[| f: bij(A, B); \ -\ !!x y. [| x:A; y:A |] ==> : r <-> : s \ + "[| f: bij(A, B); \ +\ !!x y. [| x:A; y:A |] ==> : r <-> : s \ \ |] ==> f: ord_iso(A,r,B,s)"; by (fast_tac (ZF_cs addSIs prems) 1); qed "ord_isoI"; @@ -233,7 +233,7 @@ by (etac CollectE 1); by (etac (bspec RS bspec RS iffD2) 1); by (REPEAT (eresolve_tac [asm_rl, - bij_converse_bij RS bij_is_fun RS apply_type] 1)); + bij_converse_bij RS bij_is_fun RS apply_type] 1)); by (asm_simp_tac (ZF_ss addsimps [right_inverse_bij]) 1); qed "ord_iso_converse"; @@ -241,7 +241,7 @@ (*Rewriting with bijections and converse (function inverse)*) val bij_inverse_ss = ZF_ss setsolver (type_auto_tac [ord_iso_is_bij, bij_is_fun, apply_type, - bij_converse_bij, comp_fun, comp_bij]) + bij_converse_bij, comp_fun, comp_bij]) addsimps [right_inverse_bij, left_inverse_bij, comp_fun_apply]; @@ -276,7 +276,7 @@ (** Two monotone maps can make an order-isomorphism **) goalw Order.thy [ord_iso_def, mono_map_def] - "!!f g. [| f: mono_map(A,r,B,s); g: mono_map(B,s,A,r); \ + "!!f g. [| f: mono_map(A,r,B,s); g: mono_map(B,s,A,r); \ \ f O g = id(B); g O f = id(A) |] ==> f: ord_iso(A,r,B,s)"; by (safe_tac ZF_cs); by (REPEAT_FIRST (ares_tac [fg_imp_bijective])); @@ -288,8 +288,8 @@ qed "mono_ord_isoI"; goal Order.thy - "!!B. [| well_ord(A,r); well_ord(B,s); \ -\ f: mono_map(A,r,B,s); converse(f): mono_map(B,s,A,r) \ + "!!B. [| well_ord(A,r); well_ord(B,s); \ +\ f: mono_map(A,r,B,s); converse(f): mono_map(B,s,A,r) \ \ |] ==> f: ord_iso(A,r,B,s)"; by (REPEAT (ares_tac [mono_ord_isoI] 1)); by (forward_tac [mono_map_is_fun RS fun_is_rel] 1); @@ -327,7 +327,7 @@ by (dtac equalityD1 1); by (fast_tac (ZF_cs addSIs [bexI]) 1); by (fast_tac (ZF_cs addSIs [bexI] - addIs [bij_is_fun RS apply_type]) 1); + addIs [bij_is_fun RS apply_type]) 1); qed "wf_on_ord_iso"; goalw Order.thy [well_ord_def, tot_ord_def] @@ -359,14 +359,14 @@ by (REPEAT_FIRST (ares_tac [pred_subset])); (*Now we know f`x < x *) by (EVERY1 [dtac (ord_iso_is_bij RS bij_is_fun RS apply_type), - assume_tac]); + assume_tac]); (*Now we also know f`x : pred(A,x,r); contradiction! *) by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, pred_def]) 1); qed "well_ord_iso_predE"; (*Simple consequence of Lemma 6.1*) goal Order.thy - "!!r. [| well_ord(A,r); f : ord_iso(pred(A,a,r), r, pred(A,c,r), r); \ + "!!r. [| well_ord(A,r); f : ord_iso(pred(A,a,r), r, pred(A,c,r), r); \ \ a:A; c:A |] ==> a=c"; by (forward_tac [well_ord_is_trans_on] 1); by (forward_tac [well_ord_is_linear] 1); @@ -374,20 +374,20 @@ by (dtac ord_iso_sym 1); by (REPEAT (*because there are two symmetric cases*) (EVERY [eresolve_tac [pred_subset RSN (2, well_ord_subset) RS - well_ord_iso_predE] 1, - fast_tac (ZF_cs addSIs [predI]) 2, - asm_simp_tac (ZF_ss addsimps [trans_pred_pred_eq]) 1])); + well_ord_iso_predE] 1, + fast_tac (ZF_cs addSIs [predI]) 2, + asm_simp_tac (ZF_ss addsimps [trans_pred_pred_eq]) 1])); qed "well_ord_iso_pred_eq"; (*Does not assume r is a wellordering!*) goalw Order.thy [ord_iso_def, pred_def] - "!!r. [| f : ord_iso(A,r,B,s); a:A |] ==> \ + "!!r. [| f : ord_iso(A,r,B,s); a:A |] ==> \ \ f `` pred(A,a,r) = pred(B, f`a, s)"; by (etac CollectE 1); by (asm_simp_tac (ZF_ss addsimps [[bij_is_fun, Collect_subset] MRS image_fun]) 1); by (safe_tac (eq_cs addSEs [bij_is_fun RS apply_type])); -by (resolve_tac [RepFun_eqI] 1); +by (rtac RepFun_eqI 1); by (fast_tac (ZF_cs addSIs [right_inverse_bij RS sym]) 1); by (asm_simp_tac bij_inverse_ss 1); qed "ord_iso_image_pred"; @@ -395,10 +395,10 @@ (*But in use, A and B may themselves be initial segments. Then use trans_pred_pred_eq to simplify the pred(pred...) terms. See just below.*) goal Order.thy - "!!r. [| f : ord_iso(A,r,B,s); a:A |] ==> \ + "!!r. [| f : ord_iso(A,r,B,s); a:A |] ==> \ \ restrict(f, pred(A,a,r)) : ord_iso(pred(A,a,r), r, pred(B, f`a, s), s)"; by (asm_simp_tac (ZF_ss addsimps [ord_iso_image_pred RS sym]) 1); -by (rewrite_goals_tac [ord_iso_def]); +by (rewtac ord_iso_def); by (etac CollectE 1); by (rtac CollectI 1); by (asm_full_simp_tac (ZF_ss addsimps [pred_def]) 2); @@ -407,15 +407,15 @@ (*Tricky; a lot of forward proof!*) goal Order.thy - "!!r. [| well_ord(A,r); well_ord(B,s); : r; \ -\ f : ord_iso(pred(A,a,r), r, pred(B,b,s), s); \ -\ g : ord_iso(pred(A,c,r), r, pred(B,d,s), s); \ + "!!r. [| well_ord(A,r); well_ord(B,s); : r; \ +\ f : ord_iso(pred(A,a,r), r, pred(B,b,s), s); \ +\ g : ord_iso(pred(A,c,r), r, pred(B,d,s), s); \ \ a:A; c:A; b:B; d:B |] ==> : s"; by (forward_tac [ord_iso_is_bij RS bij_is_fun RS apply_type] 1 THEN REPEAT1 (eresolve_tac [asm_rl, predI, predE] 1)); by (subgoal_tac "b = g`a" 1); by (asm_simp_tac ZF_ss 1); -by (resolve_tac [well_ord_iso_pred_eq] 1); +by (rtac well_ord_iso_pred_eq 1); by (REPEAT_SOME assume_tac); by (forward_tac [ord_iso_restrict_pred] 1 THEN REPEAT1 (eresolve_tac [asm_rl, predI] 1)); @@ -479,44 +479,44 @@ goalw Order.thy [ord_iso_map_def, function_def] "!!B. well_ord(B,s) ==> function(ord_iso_map(A,r,B,s))"; by (safe_tac ZF_cs); -by (resolve_tac [well_ord_iso_pred_eq] 1); +by (rtac well_ord_iso_pred_eq 1); by (REPEAT_SOME assume_tac); by (eresolve_tac [ord_iso_sym RS ord_iso_trans] 1); by (assume_tac 1); qed "function_ord_iso_map"; goal Order.thy - "!!B. well_ord(B,s) ==> ord_iso_map(A,r,B,s) \ + "!!B. well_ord(B,s) ==> ord_iso_map(A,r,B,s) \ \ : domain(ord_iso_map(A,r,B,s)) -> range(ord_iso_map(A,r,B,s))"; by (asm_simp_tac (ZF_ss addsimps [Pi_iff, function_ord_iso_map, - ord_iso_map_subset RS domain_times_range]) 1); + ord_iso_map_subset RS domain_times_range]) 1); qed "ord_iso_map_fun"; goalw Order.thy [mono_map_def] - "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s) \ -\ : mono_map(domain(ord_iso_map(A,r,B,s)), r, \ -\ range(ord_iso_map(A,r,B,s)), s)"; + "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s) \ +\ : mono_map(domain(ord_iso_map(A,r,B,s)), r, \ +\ range(ord_iso_map(A,r,B,s)), s)"; by (asm_simp_tac (ZF_ss addsimps [ord_iso_map_fun]) 1); by (safe_tac ZF_cs); by (subgoals_tac ["x:A", "xa:A", "y:B", "ya:B"] 1); by (REPEAT (fast_tac (ZF_cs addSEs [ord_iso_map_subset RS subsetD RS SigmaE]) 2)); by (asm_simp_tac (ZF_ss addsimps [ord_iso_map_fun RSN (2,apply_equality)]) 1); -by (rewrite_goals_tac [ord_iso_map_def]); +by (rewtac ord_iso_map_def); by (safe_tac (ZF_cs addSEs [UN_E])); -by (resolve_tac [well_ord_iso_preserving] 1 THEN REPEAT_FIRST assume_tac); +by (rtac well_ord_iso_preserving 1 THEN REPEAT_FIRST assume_tac); qed "ord_iso_map_mono_map"; goalw Order.thy [mono_map_def] - "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s) \ -\ : ord_iso(domain(ord_iso_map(A,r,B,s)), r, \ -\ range(ord_iso_map(A,r,B,s)), s)"; -by (resolve_tac [well_ord_mono_ord_isoI] 1); + "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s) \ +\ : ord_iso(domain(ord_iso_map(A,r,B,s)), r, \ +\ range(ord_iso_map(A,r,B,s)), s)"; +by (rtac well_ord_mono_ord_isoI 1); by (resolve_tac [converse_ord_iso_map RS subst] 4); by (asm_simp_tac (ZF_ss addsimps [ord_iso_map_subset RS converse_converse, - domain_converse, range_converse]) 4); + domain_converse, range_converse]) 4); by (REPEAT (ares_tac [ord_iso_map_mono_map] 3)); by (ALLGOALS (etac well_ord_subset)); by (ALLGOALS (resolve_tac [domain_ord_iso_map, range_ord_iso_map])); @@ -524,8 +524,8 @@ (*One way of saying that domain(ord_iso_map(A,r,B,s)) is downwards-closed*) goalw Order.thy [ord_iso_map_def] - "!!B. [| well_ord(A,r); well_ord(B,s); \ -\ a: A; a ~: domain(ord_iso_map(A,r,B,s)) \ + "!!B. [| well_ord(A,r); well_ord(B,s); \ +\ a: A; a ~: domain(ord_iso_map(A,r,B,s)) \ \ |] ==> domain(ord_iso_map(A,r,B,s)) <= pred(A, a, r)"; by (safe_tac (ZF_cs addSIs [predI])); (*Case analysis on xaa vs a in r *) @@ -546,7 +546,7 @@ (*For the 4-way case analysis in the main result*) goal Order.thy "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ -\ domain(ord_iso_map(A,r,B,s)) = A | \ +\ domain(ord_iso_map(A,r,B,s)) = A | \ \ (EX x:A. domain(ord_iso_map(A,r,B,s)) = pred(A,x,r))"; by (forward_tac [well_ord_is_wf] 1); by (rewrite_goals_tac [wf_on_def, wf_def]); @@ -566,8 +566,8 @@ (*As above, by duality*) goal Order.thy - "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ -\ range(ord_iso_map(A,r,B,s)) = B | \ + "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ +\ range(ord_iso_map(A,r,B,s)) = B | \ \ (EX y:B. range(ord_iso_map(A,r,B,s))= pred(B,y,s))"; by (resolve_tac [converse_ord_iso_map RS subst] 1); by (asm_simp_tac @@ -576,22 +576,22 @@ (*Kunen's Theorem 6.3: Fundamental Theorem for Well-Ordered Sets*) goal Order.thy - "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ -\ ord_iso_map(A,r,B,s) : ord_iso(A, r, B, s) | \ + "!!B. [| well_ord(A,r); well_ord(B,s) |] ==> \ +\ ord_iso_map(A,r,B,s) : ord_iso(A, r, B, s) | \ \ (EX x:A. ord_iso_map(A,r,B,s) : ord_iso(pred(A,x,r), r, B, s)) | \ \ (EX y:B. ord_iso_map(A,r,B,s) : ord_iso(A, r, pred(B,y,s), s))"; by (forw_inst_tac [("B","B")] domain_ord_iso_map_cases 1); by (forw_inst_tac [("B","B")] range_ord_iso_map_cases 2); by (REPEAT_FIRST (eresolve_tac [asm_rl, disjE, bexE])); -by (ALLGOALS (dresolve_tac [ord_iso_map_ord_iso] THEN' assume_tac THEN' - asm_full_simp_tac (ZF_ss addsimps [bexI]))); +by (ALLGOALS (dtac ord_iso_map_ord_iso THEN' assume_tac THEN' + asm_full_simp_tac (ZF_ss addsimps [bexI]))); by (resolve_tac [wf_on_not_refl RS notE] 1); -by (eresolve_tac [well_ord_is_wf] 1); +by (etac well_ord_is_wf 1); by (assume_tac 1); by (subgoal_tac ": ord_iso_map(A,r,B,s)" 1); -by (dresolve_tac [rangeI] 1); +by (dtac rangeI 1); by (asm_full_simp_tac (ZF_ss addsimps [pred_def]) 1); -by (rewrite_goals_tac [ord_iso_map_def]); +by (rewtac ord_iso_map_def); by (fast_tac ZF_cs 1); qed "well_ord_trichotomy"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/OrderArith.ML --- a/src/ZF/OrderArith.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/OrderArith.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/OrderArith.ML +(* Title: ZF/OrderArith.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Towards ordinal arithmetic @@ -36,10 +36,10 @@ (** Elimination Rule **) val major::prems = goalw OrderArith.thy [radd_def] - "[| : radd(A,r,B,s); \ -\ !!x y. [| p'=Inl(x); x:A; p=Inr(y); y:B |] ==> Q; \ -\ !!x' x. [| p'=Inl(x'); p=Inl(x); : r; x':A; x:A |] ==> Q; \ -\ !!y' y. [| p'=Inr(y'); p=Inr(y); : s; y':B; y:B |] ==> Q \ + "[| : radd(A,r,B,s); \ +\ !!x y. [| p'=Inl(x); x:A; p=Inr(y); y:B |] ==> Q; \ +\ !!x' x. [| p'=Inl(x'); p=Inl(x); : r; x':A; x:A |] ==> Q; \ +\ !!y' y. [| p'=Inr(y'); p=Inr(y); : s; y':B; y:B |] ==> Q \ \ |] ==> Q"; by (cut_facts_tac [major] 1); (*Split into the three cases*) @@ -48,8 +48,8 @@ (*Apply each premise to correct subgoal; can't just use fast_tac because hyp_subst_tac would delete equalities too quickly*) by (EVERY (map (fn prem => - EVERY1 [rtac prem, assume_tac, REPEAT o fast_tac sum_cs]) - prems)); + EVERY1 [rtac prem, assume_tac, REPEAT o fast_tac sum_cs]) + prems)); qed "raddE"; (** Type checking **) @@ -63,7 +63,7 @@ (** Linearity **) val radd_ss = sum_ss addsimps [radd_Inl_iff, radd_Inr_iff, - radd_Inl_Inr_iff, radd_Inr_Inl_iff]; + radd_Inl_Inr_iff, radd_Inr_Inl_iff]; goalw OrderArith.thy [linear_def] "!!r s. [| linear(A,r); linear(B,s) |] ==> linear(A+B,radd(A,r,B,s))"; @@ -114,7 +114,7 @@ val case_ss = bij_inverse_ss addsimps [Inl_iff, Inl_Inr_iff, Inr_iff, Inr_Inl_iff, - case_Inl, case_Inr, InlI, InrI]; + case_Inl, case_Inr, InlI, InrI]; goal OrderArith.thy "!!f g. [| f: bij(A,C); g: bij(B,D) |] ==> \ @@ -127,8 +127,8 @@ qed "sum_bij"; goalw OrderArith.thy [ord_iso_def] - "!!r s. [| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> \ -\ (lam z:A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \ + "!!r s. [| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> \ +\ (lam z:A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \ \ : ord_iso(A+B, radd(A,r,B,s), A'+B', radd(A',r',B',s'))"; by (safe_tac (ZF_cs addSIs [sum_bij])); (*Do the beta-reductions now*) @@ -143,7 +143,7 @@ (*Could we prove an ord_iso result? Perhaps ord_iso(A+B, radd(A,r,B,s), A Un B, r Un s) *) goal OrderArith.thy - "!!A B. A Int B = 0 ==> \ + "!!A B. A Int B = 0 ==> \ \ (lam z:A+B. case(%x.x, %y.y, z)) : bij(A+B, A Un B)"; by (res_inst_tac [("d", "%z. if(z:A, Inl(z), Inr(z))")] lam_bijective 1); @@ -166,7 +166,7 @@ goal OrderArith.thy "(lam z:(A+B)+C. case(case(Inl, %y.Inr(Inl(y))), %y. Inr(Inr(y)), z)) \ -\ : ord_iso((A+B)+C, radd(A+B, radd(A,r,B,s), C, t), \ +\ : ord_iso((A+B)+C, radd(A+B, radd(A,r,B,s), C, t), \ \ A+(B+C), radd(A, r, B+C, radd(B,s,C,t)))"; by (resolve_tac [sum_assoc_bij RS ord_isoI] 1); by (REPEAT_FIRST (etac sumE)); @@ -179,16 +179,16 @@ (** Rewrite rule. Can be used to obtain introduction rules **) goalw OrderArith.thy [rmult_def] - "!!r s. <, > : rmult(A,r,B,s) <-> \ -\ (: r & a':A & a:A & b': B & b: B) | \ + "!!r s. <, > : rmult(A,r,B,s) <-> \ +\ (: r & a':A & a:A & b': B & b: B) | \ \ (: s & a'=a & a:A & b': B & b: B)"; by (fast_tac ZF_cs 1); qed "rmult_iff"; val major::prems = goal OrderArith.thy - "[| <, > : rmult(A,r,B,s); \ -\ [| : r; a':A; a:A; b':B; b:B |] ==> Q; \ -\ [| : s; a:A; a'=a; b':B; b:B |] ==> Q \ + "[| <, > : rmult(A,r,B,s); \ +\ [| : r; a':A; a:A; b':B; b:B |] ==> Q; \ +\ [| : s; a:A; a'=a; b':B; b:B |] ==> Q \ \ |] ==> Q"; by (rtac (major RS (rmult_iff RS iffD1) RS disjE) 1); by (DEPTH_SOLVE (eresolve_tac ([asm_rl, conjE] @ prems) 1)); @@ -262,8 +262,8 @@ qed "prod_bij"; goalw OrderArith.thy [ord_iso_def] - "!!r s. [| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> \ -\ (lam :A*B. ) \ + "!!r s. [| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> \ +\ (lam :A*B. ) \ \ : ord_iso(A*B, rmult(A,r,B,s), A'*B', rmult(A',r',B',s'))"; by (safe_tac (ZF_cs addSIs [prod_bij])); by (ALLGOALS @@ -281,7 +281,7 @@ (*Used??*) goal OrderArith.thy - "!!x xr. well_ord({x},xr) ==> \ + "!!x xr. well_ord({x},xr) ==> \ \ (lam z:A. ) : ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))"; by (resolve_tac [singleton_prod_bij RS ord_isoI] 1); by (asm_simp_tac (ZF_ss addsimps [rmult_iff]) 1); @@ -294,10 +294,10 @@ "!!a. a~:C ==> \ \ (lam x:C*B + D. case(%x.x, %y., x)) \ \ : bij(C*B + D, C*B Un {a}*D)"; -by (resolve_tac [subst_elem] 1); +by (rtac subst_elem 1); by (resolve_tac [id_bij RS sum_bij RS comp_bij] 1); -by (resolve_tac [singleton_prod_bij] 1); -by (resolve_tac [sum_disjoint_bij] 1); +by (rtac singleton_prod_bij 1); +by (rtac sum_disjoint_bij 1); by (fast_tac eq_cs 1); by (asm_simp_tac (ZF_ss addcongs [case_cong] addsimps [id_conv]) 1); by (resolve_tac [comp_lam RS trans RS sym] 1); @@ -308,8 +308,8 @@ goal OrderArith.thy "!!A. [| a:A; well_ord(A,r) |] ==> \ \ (lam x:pred(A,a,r)*B + pred(B,b,s). case(%x.x, %y., x)) \ -\ : ord_iso(pred(A,a,r)*B + pred(B,b,s), \ -\ radd(A*B, rmult(A,r,B,s), B, s), \ +\ : ord_iso(pred(A,a,r)*B + pred(B,b,s), \ +\ radd(A*B, rmult(A,r,B,s), B, s), \ \ pred(A,a,r)*B Un {a}*pred(B,b,s), rmult(A,r,B,s))"; by (resolve_tac [prod_sum_singleton_bij RS ord_isoI] 1); by (asm_simp_tac @@ -349,8 +349,8 @@ qed "prod_assoc_bij"; goal OrderArith.thy - "(lam <, z>:(A*B)*C. >) \ -\ : ord_iso((A*B)*C, rmult(A*B, rmult(A,r,B,s), C, t), \ + "(lam <, z>:(A*B)*C. >) \ +\ : ord_iso((A*B)*C, rmult(A*B, rmult(A,r,B,s), C, t), \ \ A*(B*C), rmult(A, r, B*C, rmult(B,s,C,t)))"; by (resolve_tac [prod_assoc_bij RS ord_isoI] 1); by (REPEAT_FIRST (eresolve_tac [SigmaE, ssubst])); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/OrderType.ML --- a/src/ZF/OrderType.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/OrderType.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/OrderType.ML +(* Title: ZF/OrderType.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Order types and ordinal arithmetic in Zermelo-Fraenkel Set Theory @@ -19,7 +19,7 @@ by (rtac (wf_Memrel RS wf_imp_wf_on) 1); by (resolve_tac [prem RS ltE] 1); by (asm_simp_tac (ZF_ss addsimps [linear_def, Memrel_iff, - [ltI, prem] MRS lt_trans2 RS ltD]) 1); + [ltI, prem] MRS lt_trans2 RS ltD]) 1); by (REPEAT (resolve_tac [ballI, Ord_linear] 1)); by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1)); qed "le_well_ord_Memrel"; @@ -54,7 +54,7 @@ (*Kunen's Theorem 7.3 (ii), page 16. Isomorphic ordinals are equal*) goal OrderType.thy - "!!i. [| Ord(i); Ord(j); f: ord_iso(i,Memrel(i),j,Memrel(j)) \ + "!!i. [| Ord(i); Ord(j); f: ord_iso(i,Memrel(i),j,Memrel(j)) \ \ |] ==> i=j"; by (res_inst_tac [("i","i"),("j","j")] Ord_linear_lt 1); by (REPEAT (eresolve_tac [asm_rl, ord_iso_sym, Ord_iso_implies_eq_lemma] 1)); @@ -88,7 +88,7 @@ "!!r. [| wf[A](r); x:A |] ==> \ \ ordermap(A,r) ` x = {ordermap(A,r)`y . y : pred(A,x,r)}"; by (asm_simp_tac (ZF_ss addsimps [ordermap_eq_image, pred_subset, - ordermap_type RS image_fun]) 1); + ordermap_type RS image_fun]) 1); qed "ordermap_pred_unfold"; (*pred-unfolded version. NOT suitable for rewriting -- loops!*) @@ -98,8 +98,8 @@ fun ordermap_elim_tac i = EVERY [etac (ordermap_unfold RS equalityD1 RS subsetD RS RepFunE) i, - assume_tac (i+1), - assume_tac i]; + assume_tac (i+1), + assume_tac i]; goalw OrderType.thy [well_ord_def, tot_ord_def, part_ord_def] "!!r. [| well_ord(A,r); x:A |] ==> Ord(ordermap(A,r) ` x)"; @@ -128,7 +128,7 @@ (*** ordermap preserves the orderings in both directions ***) goal OrderType.thy - "!!r. [| : r; wf[A](r); w: A; x: A |] ==> \ + "!!r. [| : r; wf[A](r); w: A; x: A |] ==> \ \ ordermap(A,r)`w : ordermap(A,r)`x"; by (eres_inst_tac [("x1", "x")] (ordermap_unfold RS ssubst) 1); by (assume_tac 1); @@ -149,14 +149,14 @@ qed "converse_ordermap_mono"; bind_thm ("ordermap_surj", - rewrite_rule [symmetric ordertype_def] - (ordermap_type RS surj_image)); + rewrite_rule [symmetric ordertype_def] + (ordermap_type RS surj_image)); goalw OrderType.thy [well_ord_def, tot_ord_def, bij_def, inj_def] "!!r. well_ord(A,r) ==> ordermap(A,r) : bij(A, ordertype(A,r))"; by (fast_tac (ZF_cs addSIs [ordermap_type, ordermap_surj] - addEs [linearE] - addDs [ordermap_mono] + addEs [linearE] + addDs [ordermap_mono] addss (ZF_ss addsimps [mem_not_refl])) 1); qed "ordermap_bij"; @@ -171,27 +171,27 @@ by (fast_tac (ZF_cs addSEs [MemrelE, converse_ordermap_mono]) 2); by (rewtac well_ord_def); by (fast_tac (ZF_cs addSIs [MemrelI, ordermap_mono, - ordermap_type RS apply_type]) 1); + ordermap_type RS apply_type]) 1); qed "ordertype_ord_iso"; goal OrderType.thy - "!!f. [| f: ord_iso(A,r,B,s); well_ord(B,s) |] ==> \ + "!!f. [| f: ord_iso(A,r,B,s); well_ord(B,s) |] ==> \ \ ordertype(A,r) = ordertype(B,s)"; by (forward_tac [well_ord_ord_iso] 1 THEN assume_tac 1); -by (resolve_tac [Ord_iso_implies_eq] 1 - THEN REPEAT (eresolve_tac [Ord_ordertype] 1)); +by (rtac Ord_iso_implies_eq 1 + THEN REPEAT (etac Ord_ordertype 1)); by (deepen_tac (ZF_cs addIs [ord_iso_trans, ord_iso_sym] addSEs [ordertype_ord_iso]) 0 1); qed "ordertype_eq"; goal OrderType.thy - "!!A B. [| ordertype(A,r) = ordertype(B,s); \ + "!!A B. [| ordertype(A,r) = ordertype(B,s); \ \ well_ord(A,r); well_ord(B,s) \ \ |] ==> EX f. f: ord_iso(A,r,B,s)"; -by (resolve_tac [exI] 1); +by (rtac exI 1); by (resolve_tac [ordertype_ord_iso RS ord_iso_trans] 1); by (assume_tac 1); -by (eresolve_tac [ssubst] 1); +by (etac ssubst 1); by (eresolve_tac [ordertype_ord_iso RS ord_iso_sym] 1); qed "ordertype_eq_imp_ord_iso"; @@ -200,9 +200,9 @@ (*Ordertype of Memrel*) goal OrderType.thy "!!i. j le i ==> ordertype(j,Memrel(i)) = j"; by (resolve_tac [Ord_iso_implies_eq RS sym] 1); -by (eresolve_tac [ltE] 1); +by (etac ltE 1); by (REPEAT (ares_tac [le_well_ord_Memrel, Ord_ordertype] 1)); -by (resolve_tac [ord_iso_trans] 1); +by (rtac ord_iso_trans 1); by (eresolve_tac [le_well_ord_Memrel RS ordertype_ord_iso] 2); by (resolve_tac [id_bij RS ord_isoI] 1); by (asm_simp_tac (ZF_ss addsimps [id_conv, Memrel_iff]) 1); @@ -215,7 +215,7 @@ goal OrderType.thy "ordertype(0,r) = 0"; by (resolve_tac [id_bij RS ord_isoI RS ordertype_eq RS trans] 1); by (etac emptyE 1); -by (resolve_tac [well_ord_0] 1); +by (rtac well_ord_0 1); by (resolve_tac [Ord_0 RS ordertype_Memrel] 1); qed "ordertype_0"; @@ -227,8 +227,8 @@ (*Ordermap returns the same result if applied to an initial segment*) goal OrderType.thy - "!!r. [| well_ord(A,r); y:A; z: pred(A,y,r) |] ==> \ -\ ordermap(pred(A,y,r), r) ` z = ordermap(A, r) ` z"; + "!!r. [| well_ord(A,r); y:A; z: pred(A,y,r) |] ==> \ +\ ordermap(pred(A,y,r), r) ` z = ordermap(A, r) ` z"; by (forward_tac [[well_ord_is_wf, pred_subset] MRS wf_on_subset_A] 1); by (wf_on_ind_tac "z" [] 1); by (safe_tac (ZF_cs addSEs [predE])); @@ -255,7 +255,7 @@ by (asm_simp_tac (ZF_ss addsimps [ordertype_unfold, pred_subset RSN (2, well_ord_subset)]) 1); by (fast_tac (ZF_cs addIs [ordermap_pred_eq_ordermap, RepFun_eqI] - addEs [predE]) 1); + addEs [predE]) 1); qed "ordertype_pred_subset"; goal OrderType.thy @@ -264,12 +264,12 @@ by (resolve_tac [ordertype_pred_subset RS subset_imp_le RS leE] 1); by (REPEAT (ares_tac [Ord_ordertype, well_ord_subset, pred_subset] 1)); by (eresolve_tac [sym RS ordertype_eq_imp_ord_iso RS exE] 1); -by (eresolve_tac [well_ord_iso_predE] 3); +by (etac well_ord_iso_predE 3); by (REPEAT (ares_tac [pred_subset, well_ord_subset] 1)); qed "ordertype_pred_lt"; (*May rewrite with this -- provided no rules are supplied for proving that - well_ord(pred(A,x,r), r) *) + well_ord(pred(A,x,r), r) *) goal OrderType.thy "!!A r. well_ord(A,r) ==> \ \ ordertype(A,r) = {ordertype(pred(A,x,r),r). x:A}"; @@ -277,10 +277,10 @@ by (fast_tac (ZF_cs addss (ZF_ss addsimps [ordertype_def, - well_ord_is_wf RS ordermap_eq_image, - ordermap_type RS image_fun, - ordermap_pred_eq_ordermap, - pred_subset])) + well_ord_is_wf RS ordermap_eq_image, + ordermap_type RS image_fun, + ordermap_pred_eq_ordermap, + pred_subset])) 1); qed "ordertype_pred_unfold"; @@ -289,15 +289,15 @@ (*proof by Krzysztof Grabczewski*) goalw OrderType.thy [Ord_alt_def] "!!i. Ord(i) ==> Ord_alt(i)"; -by (resolve_tac [conjI] 1); -by (eresolve_tac [well_ord_Memrel] 1); +by (rtac conjI 1); +by (etac well_ord_Memrel 1); by (rewrite_goals_tac [Ord_def, Transset_def, pred_def, Memrel_def]); by (fast_tac eq_cs 1); qed "Ord_is_Ord_alt"; (*proof by lcp*) goalw OrderType.thy [Ord_alt_def, Ord_def, Transset_def, well_ord_def, - tot_ord_def, part_ord_def, trans_on_def] + tot_ord_def, part_ord_def, trans_on_def] "!!i. Ord_alt(i) ==> Ord(i)"; by (asm_full_simp_tac (ZF_ss addsimps [Memrel_iff, pred_Memrel]) 1); by (safe_tac ZF_cs); @@ -346,7 +346,7 @@ (*In fact, pred(A+B, Inl(a), radd(A,r,B,s)) = pred(A,a,r)+0 *) goalw OrderType.thy [pred_def] "!!A B. a:A ==> \ -\ (lam x:pred(A,a,r). Inl(x)) \ +\ (lam x:pred(A,a,r). Inl(x)) \ \ : bij(pred(A,a,r), pred(A+B, Inl(a), radd(A,r,B,s)))"; by (res_inst_tac [("d", "case(%x.x, %y.y)")] lam_bijective 1); by (safe_tac sum_cs); @@ -366,7 +366,7 @@ goalw OrderType.thy [pred_def, id_def] "!!A B. b:B ==> \ -\ id(A+pred(B,b,s)) \ +\ id(A+pred(B,b,s)) \ \ : bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))"; by (res_inst_tac [("d", "%z.z")] lam_bijective 1); by (safe_tac sum_cs); @@ -393,12 +393,12 @@ goalw OrderType.thy [oadd_def] "!!i. Ord(i) ==> i++0 = i"; by (asm_simp_tac (ZF_ss addsimps [Memrel_0, ordertype_sum_0_eq, - ordertype_Memrel, well_ord_Memrel]) 1); + ordertype_Memrel, well_ord_Memrel]) 1); qed "oadd_0"; goalw OrderType.thy [oadd_def] "!!i. Ord(i) ==> 0++i = i"; by (asm_simp_tac (ZF_ss addsimps [Memrel_0, ordertype_0_sum_eq, - ordertype_Memrel, well_ord_Memrel]) 1); + ordertype_Memrel, well_ord_Memrel]) 1); qed "oadd_0_left"; @@ -406,20 +406,20 @@ proofs by lcp. ***) goalw OrderType.thy [oadd_def] "!!i j k. [| k k < i++j"; -by (resolve_tac [ltE] 1 THEN assume_tac 1); -by (resolve_tac [ltI] 1); +by (rtac ltE 1 THEN assume_tac 1); +by (rtac ltI 1); by (REPEAT (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel] 2)); by (asm_simp_tac (ZF_ss addsimps [ordertype_pred_unfold, - well_ord_radd, well_ord_Memrel, - ordertype_pred_Inl_eq, - lt_pred_Memrel, leI RS le_ordertype_Memrel] - setloop rtac (InlI RSN (2,RepFun_eqI))) 1); + well_ord_radd, well_ord_Memrel, + ordertype_pred_Inl_eq, + lt_pred_Memrel, leI RS le_ordertype_Memrel] + setloop rtac (InlI RSN (2,RepFun_eqI))) 1); qed "lt_oadd1"; (*Thus also we obtain the rule i++j = k ==> i le k *) goal OrderType.thy "!!i j. [| Ord(i); Ord(j) |] ==> i le i++j"; -by (resolve_tac [all_lt_imp_le] 1); +by (rtac all_lt_imp_le 1); by (REPEAT (ares_tac [Ord_oadd, lt_oadd1] 1)); qed "oadd_le_self"; @@ -433,25 +433,25 @@ qed "id_ord_iso_Memrel"; goal OrderType.thy - "!!k. [| well_ord(A,r); k \ -\ ordertype(A+k, radd(A, r, k, Memrel(j))) = \ + "!!k. [| well_ord(A,r); k \ +\ ordertype(A+k, radd(A, r, k, Memrel(j))) = \ \ ordertype(A+k, radd(A, r, k, Memrel(k)))"; -by (eresolve_tac [ltE] 1); +by (etac ltE 1); by (resolve_tac [ord_iso_refl RS sum_ord_iso_cong RS ordertype_eq] 1); by (eresolve_tac [OrdmemD RS id_ord_iso_Memrel RS ord_iso_sym] 1); by (REPEAT_FIRST (ares_tac [well_ord_radd, well_ord_Memrel])); qed "ordertype_sum_Memrel"; goalw OrderType.thy [oadd_def] "!!i j k. [| k i++k < i++j"; -by (resolve_tac [ltE] 1 THEN assume_tac 1); +by (rtac ltE 1 THEN assume_tac 1); by (resolve_tac [ordertype_pred_unfold RS equalityD2 RS subsetD RS ltI] 1); by (REPEAT_FIRST (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel])); -by (resolve_tac [RepFun_eqI] 1); -by (eresolve_tac [InrI] 2); +by (rtac RepFun_eqI 1); +by (etac InrI 2); by (asm_simp_tac (ZF_ss addsimps [ordertype_pred_Inr_eq, well_ord_Memrel, - lt_pred_Memrel, leI RS le_ordertype_Memrel, - ordertype_sum_Memrel]) 1); + lt_pred_Memrel, leI RS le_ordertype_Memrel, + ordertype_sum_Memrel]) 1); qed "oadd_lt_mono2"; goal OrderType.thy @@ -482,13 +482,13 @@ by (etac revcut_rl 1); by (asm_full_simp_tac (ZF_ss addsimps [ordertype_pred_unfold, well_ord_radd, - well_ord_Memrel]) 1); + well_ord_Memrel]) 1); by (eresolve_tac [ltD RS RepFunE] 1); by (fast_tac (sum_cs addss - (ZF_ss addsimps [ordertype_pred_Inl_eq, well_ord_Memrel, - ltI, lt_pred_Memrel, le_ordertype_Memrel, leI, - ordertype_pred_Inr_eq, - ordertype_sum_Memrel])) 1); + (ZF_ss addsimps [ordertype_pred_Inl_eq, well_ord_Memrel, + ltI, lt_pred_Memrel, le_ordertype_Memrel, leI, + ordertype_pred_Inr_eq, + ordertype_sum_Memrel])) 1); qed "lt_oadd_disj"; @@ -498,11 +498,11 @@ "!!i j k. [| Ord(i); Ord(j); Ord(k) |] ==> (i++j)++k = i++(j++k)"; by (resolve_tac [ordertype_eq RS trans] 1); by (rtac ([ordertype_ord_iso RS ord_iso_sym, ord_iso_refl] MRS - sum_ord_iso_cong) 1); + sum_ord_iso_cong) 1); by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel, Ord_ordertype] 1)); by (resolve_tac [sum_assoc_ord_iso RS ordertype_eq RS trans] 1); by (rtac ([ord_iso_refl, ordertype_ord_iso] MRS sum_ord_iso_cong RS - ordertype_eq) 2); + ordertype_eq) 2); by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel, Ord_ordertype] 1)); qed "oadd_assoc"; @@ -512,7 +512,7 @@ by (eresolve_tac [ltI RS lt_oadd_disj RS disjE] 1); by (REPEAT (ares_tac [Ord_oadd] 1)); by (fast_tac (ZF_cs addIs [lt_oadd1, oadd_lt_mono2] - addss (ZF_ss addsimps [Ord_mem_iff_lt, Ord_oadd])) 3); + addss (ZF_ss addsimps [Ord_mem_iff_lt, Ord_oadd])) 3); by (fast_tac ZF_cs 2); by (fast_tac (ZF_cs addSEs [ltE]) 1); qed "oadd_unfold"; @@ -535,7 +535,7 @@ "[| Ord(i); !!x. x:A ==> Ord(j(x)); a:A |] ==> \ \ i ++ (UN x:A. j(x)) = (UN x:A. i++j(x))"; by (fast_tac (eq_cs addIs (prems @ [ltI, Ord_UN, Ord_oadd, - lt_oadd1 RS ltD, oadd_lt_mono2 RS ltD]) + lt_oadd1 RS ltD, oadd_lt_mono2 RS ltD]) addSEs [ltE, ltI RS lt_oadd_disj RS disjE]) 1); qed "oadd_UN"; @@ -543,8 +543,8 @@ "!!i j. [| Ord(i); Limit(j) |] ==> i++j = (UN k:j. i++k)"; by (forward_tac [Limit_has_0 RS ltD] 1); by (asm_simp_tac (ZF_ss addsimps [Limit_is_Ord RS Ord_in_Ord, - oadd_UN RS sym, Union_eq_UN RS sym, - Limit_Union_eq]) 1); + oadd_UN RS sym, Union_eq_UN RS sym, + Limit_Union_eq]) 1); qed "oadd_Limit"; (** Order/monotonicity properties of ordinal addition **) @@ -554,28 +554,28 @@ by (asm_simp_tac (ZF_ss addsimps [oadd_0, Ord_0_le]) 1); by (asm_simp_tac (ZF_ss addsimps [oadd_succ, succ_leI]) 1); by (asm_simp_tac (ZF_ss addsimps [oadd_Limit]) 1); -by (resolve_tac [le_trans] 1); -by (resolve_tac [le_implies_UN_le_UN] 2); +by (rtac le_trans 1); +by (rtac le_implies_UN_le_UN 2); by (fast_tac ZF_cs 2); by (asm_simp_tac (ZF_ss addsimps [Union_eq_UN RS sym, Limit_Union_eq, - le_refl, Limit_is_Ord]) 1); + le_refl, Limit_is_Ord]) 1); qed "oadd_le_self2"; goal OrderType.thy "!!i j k. [| k le j; Ord(i) |] ==> k++i le j++i"; by (forward_tac [lt_Ord] 1); by (forward_tac [le_Ord2] 1); -by (eresolve_tac [trans_induct3] 1); +by (etac trans_induct3 1); by (asm_simp_tac (ZF_ss addsimps [oadd_0]) 1); by (asm_simp_tac (ZF_ss addsimps [oadd_succ, succ_le_iff]) 1); by (asm_simp_tac (ZF_ss addsimps [oadd_Limit]) 1); -by (resolve_tac [le_implies_UN_le_UN] 1); +by (rtac le_implies_UN_le_UN 1); by (fast_tac ZF_cs 1); qed "oadd_le_mono1"; goal OrderType.thy "!!i j. [| i' le i; j' i'++j' < i++j"; -by (resolve_tac [lt_trans1] 1); +by (rtac lt_trans1 1); by (REPEAT (eresolve_tac [asm_rl, oadd_le_mono1, oadd_lt_mono2, ltE, - Ord_succD] 1)); + Ord_succD] 1)); qed "oadd_lt_mono"; goal OrderType.thy "!!i j. [| i' le i; j' le j |] ==> i'++j' le i++j"; @@ -585,7 +585,7 @@ goal OrderType.thy "!!i j k. [| Ord(i); Ord(j); Ord(k) |] ==> i++j le i++k <-> j le k"; by (asm_simp_tac (ZF_ss addsimps [oadd_lt_iff2, oadd_succ RS sym, - Ord_succ]) 1); + Ord_succ]) 1); qed "oadd_le_iff2"; @@ -598,17 +598,17 @@ by (res_inst_tac [("d", "case(%x.x, %y.y)")] lam_bijective 1); by (fast_tac (sum_cs addSIs [if_type]) 1); by (fast_tac (ZF_cs addSIs [case_type]) 1); -by (eresolve_tac [sumE] 2); +by (etac sumE 2); by (ALLGOALS (asm_simp_tac (sum_ss setloop split_tac [expand_if]))); qed "bij_sum_Diff"; goal OrderType.thy - "!!i j. i le j ==> \ -\ ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j))) = \ + "!!i j. i le j ==> \ +\ ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j))) = \ \ ordertype(j, Memrel(j))"; by (safe_tac (ZF_cs addSDs [le_subset_iff RS iffD1])); by (resolve_tac [bij_sum_Diff RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1); -by (eresolve_tac [well_ord_Memrel] 3); +by (etac well_ord_Memrel 3); by (assume_tac 1); by (asm_simp_tac (radd_ss setloop split_tac [expand_if] addsimps [Memrel_iff]) 1); @@ -619,32 +619,32 @@ qed "ordertype_sum_Diff"; goalw OrderType.thy [oadd_def, odiff_def] - "!!i j. i le j ==> \ + "!!i j. i le j ==> \ \ i ++ (j--i) = ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j)))"; by (safe_tac (ZF_cs addSDs [le_subset_iff RS iffD1])); by (resolve_tac [sum_ord_iso_cong RS ordertype_eq] 1); -by (eresolve_tac [id_ord_iso_Memrel] 1); +by (etac id_ord_iso_Memrel 1); by (resolve_tac [ordertype_ord_iso RS ord_iso_sym] 1); by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel RS well_ord_subset, - Diff_subset] 1)); + Diff_subset] 1)); qed "oadd_ordertype_Diff"; goal OrderType.thy "!!i j. i le j ==> i ++ (j--i) = j"; by (asm_simp_tac (ZF_ss addsimps [oadd_ordertype_Diff, ordertype_sum_Diff, - ordertype_Memrel, lt_Ord2 RS Ord_succD]) 1); + ordertype_Memrel, lt_Ord2 RS Ord_succD]) 1); qed "oadd_odiff_inverse"; goalw OrderType.thy [odiff_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i--j)"; by (REPEAT (ares_tac [Ord_ordertype, well_ord_Memrel RS well_ord_subset, - Diff_subset] 1)); + Diff_subset] 1)); qed "Ord_odiff"; (*By oadd_inject, the difference between i and j is unique. Note that we get i++j = k ==> j = k--i. *) goal OrderType.thy "!!i j. [| Ord(i); Ord(j) |] ==> (i++j) -- i = j"; -br oadd_inject 1; +by (rtac oadd_inject 1); by (REPEAT (ares_tac [Ord_ordertype, Ord_oadd, Ord_odiff] 2)); by (asm_simp_tac (ZF_ss addsimps [oadd_odiff_inverse, oadd_le_self]) 1); qed "odiff_oadd_inverse"; @@ -654,9 +654,9 @@ by (rtac (k_le_i RS lt_Ord RSN (2,oadd_lt_cancel2)) 1); by (simp_tac (ZF_ss addsimps [i_lt_j, k_le_i, [k_le_i, leI] MRS le_trans, - oadd_odiff_inverse]) 1); + oadd_odiff_inverse]) 1); by (REPEAT (resolve_tac (Ord_odiff :: - ([i_lt_j, k_le_i] RL [lt_Ord, lt_Ord2])) 1)); + ([i_lt_j, k_le_i] RL [lt_Ord, lt_Ord2])) 1)); qed "odiff_lt_mono2"; @@ -671,7 +671,7 @@ goalw OrderType.thy [pred_def] "!!A B. [| a:A; b:B |] ==> \ -\ pred(A*B, , rmult(A,r,B,s)) = \ +\ pred(A*B, , rmult(A,r,B,s)) = \ \ pred(A,a,r)*B Un ({a} * pred(B,b,s))"; by (safe_tac eq_cs); by (ALLGOALS (asm_full_simp_tac (ZF_ss addsimps [rmult_iff]))); @@ -681,11 +681,11 @@ goal OrderType.thy "!!A B. [| a:A; b:B; well_ord(A,r); well_ord(B,s) |] ==> \ \ ordertype(pred(A*B, , rmult(A,r,B,s)), rmult(A,r,B,s)) = \ -\ ordertype(pred(A,a,r)*B + pred(B,b,s), \ +\ ordertype(pred(A,a,r)*B + pred(B,b,s), \ \ radd(A*B, rmult(A,r,B,s), B, s))"; by (asm_simp_tac (ZF_ss addsimps [pred_Pair_eq]) 1); by (resolve_tac [ordertype_eq RS sym] 1); -by (resolve_tac [prod_sum_singleton_ord_iso] 1); +by (rtac prod_sum_singleton_ord_iso 1); by (REPEAT_FIRST (ares_tac [pred_subset, well_ord_rmult RS well_ord_subset])); by (fast_tac (ZF_cs addSEs [predE]) 1); qed "ordertype_pred_Pair_eq"; @@ -696,14 +696,14 @@ \ rmult(i,Memrel(i),j,Memrel(j))) = \ \ j**i' ++ j'"; by (asm_simp_tac (ZF_ss addsimps [ordertype_pred_Pair_eq, lt_pred_Memrel, - ltD, lt_Ord2, well_ord_Memrel]) 1); -by (resolve_tac [trans] 1); + ltD, lt_Ord2, well_ord_Memrel]) 1); +by (rtac trans 1); by (resolve_tac [ordertype_ord_iso RS sum_ord_iso_cong RS ordertype_eq] 2); -by (resolve_tac [ord_iso_refl] 3); +by (rtac ord_iso_refl 3); by (resolve_tac [id_bij RS ord_isoI RS ordertype_eq] 1); by (REPEAT_FIRST (eresolve_tac [SigmaE, sumE, ltE, ssubst])); by (REPEAT_FIRST (ares_tac [well_ord_rmult, well_ord_radd, well_ord_Memrel, - Ord_ordertype])); + Ord_ordertype])); by (ALLGOALS (asm_simp_tac (radd_ss addsimps [rmult_iff, id_conv, Memrel_iff]))); by (safe_tac ZF_cs); @@ -714,23 +714,23 @@ "!!i j. [| Ord(i); Ord(j); k \ \ EX j' i'. k = j**i' ++ j' & j' j**i' ++ j' < j**i"; -by (resolve_tac [ltI] 1); +by (rtac ltI 1); by (asm_simp_tac (ZF_ss addsimps [Ord_ordertype, well_ord_rmult, well_ord_Memrel, - lt_Ord2]) 2); + lt_Ord2]) 2); by (asm_simp_tac (ZF_ss addsimps [ordertype_pred_unfold, - well_ord_rmult, well_ord_Memrel, lt_Ord2]) 1); -by (resolve_tac [RepFun_eqI] 1); + well_ord_rmult, well_ord_Memrel, lt_Ord2]) 1); +by (rtac RepFun_eqI 1); by (fast_tac (ZF_cs addSEs [ltE]) 2); by (asm_simp_tac (ZF_ss addsimps [ordertype_pred_Pair_lemma, ltI, symmetric omult_def]) 1); @@ -740,7 +740,7 @@ "!!i j. [| Ord(i); Ord(j) |] ==> j**i = (UN j':j. UN i':i. {j**i' ++ j'})"; by (rtac (subsetI RS equalityI) 1); by (resolve_tac [lt_omult RS exE] 1); -by (eresolve_tac [ltI] 3); +by (etac ltI 3); by (REPEAT (ares_tac [Ord_omult] 1)); by (fast_tac (ZF_cs addSEs [ltE]) 1); by (fast_tac (ZF_cs addIs [omult_oadd_lt RS ltD, ltI]) 1); @@ -764,7 +764,7 @@ by (resolve_tac [ord_isoI RS ordertype_eq RS trans] 1); by (res_inst_tac [("c", "snd"), ("d", "%z.<0,z>")] lam_bijective 1); by (REPEAT_FIRST (eresolve_tac [snd_type, SigmaE, succE, emptyE, - well_ord_Memrel, ordertype_Memrel])); + well_ord_Memrel, ordertype_Memrel])); by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [rmult_iff, Memrel_iff]))); qed "omult_1"; @@ -772,7 +772,7 @@ by (resolve_tac [ord_isoI RS ordertype_eq RS trans] 1); by (res_inst_tac [("c", "fst"), ("d", "%z.")] lam_bijective 1); by (REPEAT_FIRST (eresolve_tac [fst_type, SigmaE, succE, emptyE, - well_ord_Memrel, ordertype_Memrel])); + well_ord_Memrel, ordertype_Memrel])); by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [rmult_iff, Memrel_iff]))); qed "omult_1_left"; @@ -782,14 +782,14 @@ "!!i. [| Ord(i); Ord(j); Ord(k) |] ==> i**(j++k) = (i**j)++(i**k)"; by (resolve_tac [ordertype_eq RS trans] 1); by (rtac ([ordertype_ord_iso RS ord_iso_sym, ord_iso_refl] MRS - prod_ord_iso_cong) 1); + prod_ord_iso_cong) 1); by (REPEAT (ares_tac [well_ord_rmult, well_ord_radd, well_ord_Memrel, - Ord_ordertype] 1)); + Ord_ordertype] 1)); by (rtac (sum_prod_distrib_ord_iso RS ordertype_eq RS trans) 1); by (rtac ordertype_eq 2); by (rtac ([ordertype_ord_iso, ordertype_ord_iso] MRS sum_ord_iso_cong) 2); by (REPEAT (ares_tac [well_ord_rmult, well_ord_radd, well_ord_Memrel, - Ord_ordertype] 1)); + Ord_ordertype] 1)); qed "oadd_omult_distrib"; goal OrderType.thy "!!i. [| Ord(i); Ord(j) |] ==> i**succ(j) = (i**j)++i"; @@ -803,12 +803,12 @@ "!!i j k. [| Ord(i); Ord(j); Ord(k) |] ==> (i**j)**k = i**(j**k)"; by (resolve_tac [ordertype_eq RS trans] 1); by (rtac ([ord_iso_refl, ordertype_ord_iso RS ord_iso_sym] MRS - prod_ord_iso_cong) 1); + prod_ord_iso_cong) 1); by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1)); by (resolve_tac [prod_assoc_ord_iso RS ord_iso_sym RS - ordertype_eq RS trans] 1); + ordertype_eq RS trans] 1); by (rtac ([ordertype_ord_iso, ord_iso_refl] MRS prod_ord_iso_cong RS - ordertype_eq) 2); + ordertype_eq) 2); by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Ord_ordertype] 1)); qed "omult_assoc"; @@ -826,7 +826,7 @@ "!!i j. [| Ord(i); Limit(j) |] ==> i**j = (UN k:j. i**k)"; by (asm_simp_tac (ZF_ss addsimps [Limit_is_Ord RS Ord_in_Ord, omult_UN RS sym, - Union_eq_UN RS sym, Limit_Union_eq]) 1); + Union_eq_UN RS sym, Limit_Union_eq]) 1); qed "omult_Limit"; @@ -836,52 +836,52 @@ goal OrderType.thy "!!i j. [| k k < i**j"; by (safe_tac (ZF_cs addSEs [ltE] addSIs [ltI, Ord_omult])); by (asm_simp_tac (ZF_ss addsimps [omult_unfold]) 1); -by (REPEAT (eresolve_tac [UN_I] 1)); +by (REPEAT (etac UN_I 1)); by (asm_simp_tac (ZF_ss addsimps [omult_0, oadd_0_left]) 1); qed "lt_omult1"; goal OrderType.thy "!!i j. [| Ord(i); 0 i le i**j"; -by (resolve_tac [all_lt_imp_le] 1); +by (rtac all_lt_imp_le 1); by (REPEAT (ares_tac [Ord_omult, lt_omult1, lt_Ord2] 1)); qed "omult_le_self"; goal OrderType.thy "!!i j k. [| k le j; Ord(i) |] ==> k**i le j**i"; by (forward_tac [lt_Ord] 1); by (forward_tac [le_Ord2] 1); -by (eresolve_tac [trans_induct3] 1); +by (etac trans_induct3 1); by (asm_simp_tac (ZF_ss addsimps [omult_0, le_refl, Ord_0]) 1); by (asm_simp_tac (ZF_ss addsimps [omult_succ, oadd_le_mono]) 1); by (asm_simp_tac (ZF_ss addsimps [omult_Limit]) 1); -by (resolve_tac [le_implies_UN_le_UN] 1); +by (rtac le_implies_UN_le_UN 1); by (fast_tac ZF_cs 1); qed "omult_le_mono1"; goal OrderType.thy "!!i j k. [| k i**k < i**j"; -by (resolve_tac [ltI] 1); +by (rtac ltI 1); by (asm_simp_tac (ZF_ss addsimps [omult_unfold, lt_Ord2]) 1); by (safe_tac (ZF_cs addSEs [ltE] addSIs [Ord_omult])); -by (REPEAT (eresolve_tac [UN_I] 1)); +by (REPEAT (etac UN_I 1)); by (asm_simp_tac (ZF_ss addsimps [omult_0, oadd_0, Ord_omult]) 1); qed "omult_lt_mono2"; goal OrderType.thy "!!i j k. [| k le j; Ord(i) |] ==> i**k le i**j"; -by (resolve_tac [subset_imp_le] 1); +by (rtac subset_imp_le 1); by (safe_tac (ZF_cs addSEs [ltE, make_elim Ord_succD] addSIs [Ord_omult])); by (asm_full_simp_tac (ZF_ss addsimps [omult_unfold]) 1); by (deepen_tac (ZF_cs addEs [Ord_trans, UN_I]) 0 1); qed "omult_le_mono2"; goal OrderType.thy "!!i j. [| i' le i; j' le j |] ==> i'**j' le i**j"; -by (resolve_tac [le_trans] 1); +by (rtac le_trans 1); by (REPEAT (eresolve_tac [asm_rl, omult_le_mono1, omult_le_mono2, ltE, - Ord_succD] 1)); + Ord_succD] 1)); qed "omult_le_mono"; goal OrderType.thy "!!i j. [| i' le i; j' i'**j' < i**j"; -by (resolve_tac [lt_trans1] 1); +by (rtac lt_trans1 1); by (REPEAT (eresolve_tac [asm_rl, omult_le_mono1, omult_lt_mono2, ltE, - Ord_succD] 1)); + Ord_succD] 1)); qed "omult_lt_mono"; goal OrderType.thy "!!i j. [| Ord(i); 0 i le j**i"; @@ -889,16 +889,16 @@ by (eres_inst_tac [("i","i")] trans_induct3 1); by (asm_simp_tac (ZF_ss addsimps [omult_0, Ord_0 RS le_refl]) 1); by (asm_simp_tac (ZF_ss addsimps [omult_succ, succ_le_iff]) 1); -by (eresolve_tac [lt_trans1] 1); +by (etac lt_trans1 1); by (res_inst_tac [("b", "j**x")] (oadd_0 RS subst) 1 THEN rtac oadd_lt_mono2 2); by (REPEAT (ares_tac [Ord_omult] 1)); by (asm_simp_tac (ZF_ss addsimps [omult_Limit]) 1); -by (resolve_tac [le_trans] 1); -by (resolve_tac [le_implies_UN_le_UN] 2); +by (rtac le_trans 1); +by (rtac le_implies_UN_le_UN 2); by (fast_tac ZF_cs 2); by (asm_simp_tac (ZF_ss addsimps [Union_eq_UN RS sym, Limit_Union_eq, - Limit_is_Ord RS le_refl]) 1); + Limit_is_Ord RS le_refl]) 1); qed "omult_le_self2"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Ordinal.ML --- a/src/ZF/Ordinal.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Ordinal.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Ordinal.thy +(* Title: ZF/Ordinal.thy ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge For Ordinal.thy. Ordinals in Zermelo-Fraenkel Set Theory @@ -29,7 +29,7 @@ val [prem1,prem2] = goalw Ordinal.thy [Pair_def] "[| Transset(C); : C |] ==> a:C & b: C"; -by (cut_facts_tac [prem2] 1); +by (cut_facts_tac [prem2] 1); by (fast_tac (ZF_cs addSDs [prem1 RS Transset_doubleton_D]) 1); qed "Transset_Pair_D"; @@ -137,7 +137,7 @@ goal Ordinal.thy "!!i. Ord(i) ==> Ord(succ(i))"; by (REPEAT (ares_tac [OrdI,Transset_succ] 1 ORELSE eresolve_tac [succE,ssubst,Ord_is_Transset, - Ord_contains_Transset] 1)); + Ord_contains_Transset] 1)); qed "Ord_succ"; bind_thm ("Ord_1", Ord_0 RS Ord_succ); @@ -205,11 +205,11 @@ qed "not_lt0"; goal Ordinal.thy "!!i j. j Ord(j)"; -by (eresolve_tac [ltE] 1 THEN assume_tac 1); +by (etac ltE 1 THEN assume_tac 1); qed "lt_Ord"; goal Ordinal.thy "!!i j. j Ord(i)"; -by (eresolve_tac [ltE] 1 THEN assume_tac 1); +by (etac ltE 1 THEN assume_tac 1); qed "lt_Ord2"; (* "ja le j ==> Ord(j)" *) @@ -317,12 +317,12 @@ Proof idea: show A<=B by applying the foundation axiom to A-B *) goalw Ordinal.thy [wf_def] "wf(Memrel(A))"; by (EVERY1 [rtac (foundation RS disjE RS allI), - etac disjI1, - etac bexE, - rtac (impI RS allI RS bexI RS disjI2), - etac MemrelE, - etac bspec, - REPEAT o assume_tac]); + etac disjI1, + etac bexE, + rtac (impI RS allI RS bexI RS disjI2), + etac MemrelE, + etac bspec, + REPEAT o assume_tac]); qed "wf_Memrel"; (*Transset(i) does not suffice, though ALL j:i.Transset(j) does*) @@ -370,8 +370,8 @@ (*Perform induction on i, then prove the Ord(i) subgoal using prems. *) fun trans_ind_tac a prems i = EVERY [res_inst_tac [("i",a)] trans_induct i, - rename_last_tac a ["1"] (i+1), - ares_tac prems i]; + rename_last_tac a ["1"] (i+1), + ares_tac prems i]; (*** Fundamental properties of the epsilon ordering (< on ordinals) ***) @@ -466,7 +466,7 @@ goal Ordinal.thy "j le i <-> j<=i & Ord(i) & Ord(j)"; by (fast_tac (ZF_cs addSEs [subset_imp_le, le_imp_subset] - addEs [ltE, make_elim Ord_succD]) 1); + addEs [ltE, make_elim Ord_succD]) 1); qed "le_subset_iff"; goal Ordinal.thy "i le succ(j) <-> i le j | i=succ(j) & Ord(i)"; @@ -547,7 +547,7 @@ val [ordi,ordj,ordk] = goal Ordinal.thy "[| Ord(i); Ord(j); Ord(k) |] ==> i Un j : k <-> i:k & j:k"; by (cut_facts_tac [[ordi,ordj] MRS - read_instantiate [("k","k")] Un_least_lt_iff] 1); + read_instantiate [("k","k")] Un_least_lt_iff] 1); by (asm_full_simp_tac (ZF_ss addsimps [lt_def,ordi,ordj,ordk]) 1); qed "Un_least_mem_iff"; @@ -601,10 +601,10 @@ val [leprem] = goal Ordinal.thy "[| !!x. x:A ==> c(x) le d(x) |] ==> (UN x:A. c(x)) le (UN x:A. d(x))"; -by (resolve_tac [UN_least_le] 1); -by (resolve_tac [UN_upper_le] 2); +by (rtac UN_least_le 1); +by (rtac UN_upper_le 2); by (REPEAT (ares_tac [leprem] 2)); -by (resolve_tac [Ord_UN] 1); +by (rtac Ord_UN 1); by (REPEAT (eresolve_tac [asm_rl, leprem RS ltE] 1 ORELSE dtac Ord_succD 1)); qed "le_implies_UN_le_UN"; @@ -664,22 +664,22 @@ qed "Ord_cases_disj"; val major::prems = goal Ordinal.thy - "[| Ord(i); \ -\ i=0 ==> P; \ -\ !!j. [| Ord(j); i=succ(j) |] ==> P; \ -\ Limit(i) ==> P \ + "[| Ord(i); \ +\ i=0 ==> P; \ +\ !!j. [| Ord(j); i=succ(j) |] ==> P; \ +\ Limit(i) ==> P \ \ |] ==> P"; by (cut_facts_tac [major RS Ord_cases_disj] 1); by (REPEAT (eresolve_tac (prems@[asm_rl, disjE, exE, conjE]) 1)); qed "Ord_cases"; val major::prems = goal Ordinal.thy - "[| Ord(i); \ -\ P(0); \ -\ !!x. [| Ord(x); P(x) |] ==> P(succ(x)); \ -\ !!x. [| Limit(x); ALL y:x. P(y) |] ==> P(x) \ + "[| Ord(i); \ +\ P(0); \ +\ !!x. [| Ord(x); P(x) |] ==> P(succ(x)); \ +\ !!x. [| Limit(x); ALL y:x. P(y) |] ==> P(x) \ \ |] ==> P(i)"; by (resolve_tac [major RS trans_induct] 1); -by (eresolve_tac [Ord_cases] 1); +by (etac Ord_cases 1); by (ALLGOALS (fast_tac (ZF_cs addIs prems))); qed "trans_induct3"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Perm.ML --- a/src/ZF/Perm.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Perm.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Perm.ML +(* Title: ZF/Perm.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1991 University of Cambridge The theory underlying permutation groups @@ -19,7 +19,7 @@ goalw Perm.thy [surj_def] "!!f A B. f : Pi(A,B) ==> f: surj(A,range(f))"; by (fast_tac (ZF_cs addIs [apply_equality] - addEs [range_of_fun,domain_type]) 1); + addEs [range_of_fun,domain_type]) 1); qed "fun_is_surj"; goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> range(f)=B"; @@ -35,9 +35,9 @@ qed "f_imp_surjective"; val prems = goal Perm.thy - "[| !!x. x:A ==> c(x): B; \ -\ !!y. y:B ==> d(y): A; \ -\ !!y. y:B ==> c(d(y)) = y \ + "[| !!x. x:A ==> c(x): B; \ +\ !!y. y:B ==> d(y): A; \ +\ !!y. y:B ==> c(d(y)) = y \ \ |] ==> (lam x:A.c(x)) : surj(A,B)"; by (res_inst_tac [("d", "d")] f_imp_surjective 1); by (ALLGOALS (asm_simp_tac (ZF_ss addsimps ([lam_type]@prems)) )); @@ -78,8 +78,8 @@ qed "f_imp_injective"; val prems = goal Perm.thy - "[| !!x. x:A ==> c(x): B; \ -\ !!x. x:A ==> d(c(x)) = x \ + "[| !!x. x:A ==> c(x): B; \ +\ !!x. x:A ==> d(c(x)) = x \ \ |] ==> (lam x:A.c(x)) : inj(A,B)"; by (res_inst_tac [("d", "d")] f_imp_injective 1); by (ALLGOALS (asm_simp_tac (ZF_ss addsimps ([lam_type]@prems)) )); @@ -99,10 +99,10 @@ bind_thm ("bij_is_fun", (bij_is_inj RS inj_is_fun)); val prems = goalw Perm.thy [bij_def] - "[| !!x. x:A ==> c(x): B; \ -\ !!y. y:B ==> d(y): A; \ -\ !!x. x:A ==> d(c(x)) = x; \ -\ !!y. y:B ==> c(d(y)) = y \ + "[| !!x. x:A ==> c(x): B; \ +\ !!y. y:B ==> d(y): A; \ +\ !!x. x:A ==> d(c(x)) = x; \ +\ !!y. y:B ==> c(d(y)) = y \ \ |] ==> (lam x:A.c(x)) : bij(A,B)"; by (REPEAT (ares_tac (prems @ [IntI, lam_injective, lam_surjective]) 1)); qed "lam_bijective"; @@ -164,7 +164,7 @@ val [prem] = goal Perm.thy "f: inj(A,B) ==> converse(f) : range(f)->A"; by (cut_facts_tac [prem] 1); by (asm_full_simp_tac (ZF_ss addsimps [inj_def, Pi_iff, domain_converse]) 1); -by (resolve_tac [conjI] 1); +by (rtac conjI 1); by (deepen_tac ZF_cs 0 2); by (simp_tac (ZF_ss addsimps [function_def, converse_iff]) 1); by (fast_tac (ZF_cs addEs [prem RSN (3,inj_equality)]) 1); @@ -201,16 +201,16 @@ goalw Perm.thy [bij_def] "!!f. [| f: bij(A,B); b: B |] ==> f`(converse(f)`b) = b"; by (EVERY1 [etac IntE, etac right_inverse, - etac (surj_range RS ssubst), - assume_tac]); + etac (surj_range RS ssubst), + assume_tac]); qed "right_inverse_bij"; (** Converses of injections, surjections, bijections **) goal Perm.thy "!!f A B. f: inj(A,B) ==> converse(f): inj(range(f), A)"; -by (resolve_tac [f_imp_injective] 1); -by (eresolve_tac [inj_converse_fun] 1); -by (resolve_tac [right_inverse] 1); +by (rtac f_imp_injective 1); +by (etac inj_converse_fun 1); +by (rtac right_inverse 1); by (REPEAT (assume_tac 1)); qed "inj_converse_inj"; @@ -247,8 +247,8 @@ val compEpair = rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac) - THEN prune_params_tac) - (read_instantiate [("xz","")] compE); + THEN prune_params_tac) + (read_instantiate [("xz","")] compE); val comp_cs = ZF_cs addSIs [idI] addIs [compI] addSEs [compE,idE]; @@ -327,16 +327,16 @@ goal Perm.thy "!!f g. [| g: A->B; f: B->C; a:A |] ==> (f O g)`a = f`(g`a)"; by (REPEAT (ares_tac [comp_fun,apply_equality,compI, - apply_Pair,apply_type] 1)); + apply_Pair,apply_type] 1)); qed "comp_fun_apply"; (*Simplifies compositions of lambda-abstractions*) val [prem] = goal Perm.thy - "[| !!x. x:A ==> b(x): B \ + "[| !!x. x:A ==> b(x): B \ \ |] ==> (lam y:B.c(y)) O (lam x:A. b(x)) = (lam x:A. c(b(x)))"; -by (resolve_tac [fun_extension] 1); -by (resolve_tac [comp_fun] 1); -by (resolve_tac [lam_funtype] 2); +by (rtac fun_extension 1); +by (rtac comp_fun 1); +by (rtac lam_funtype 2); by (typechk_tac (prem::ZF_typechecks)); by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply] setsolver type_auto_tac [lam_type, lam_funtype, prem]) 1); @@ -412,11 +412,11 @@ val injfD = prem RSN (3,inj_equality); by (cut_facts_tac [prem RS inj_is_fun] 1); by (fast_tac (comp_cs addIs [equalityI,apply_Pair] - addEs [domain_type, make_elim injfD]) 1); + addEs [domain_type, make_elim injfD]) 1); qed "left_comp_inverse"; (*right inverse of composition; one inclusion is - f: A->B ==> f O converse(f) <= id(B) + f: A->B ==> f O converse(f) <= id(B) *) val [prem] = goalw Perm.thy [surj_def] "f: surj(A,B) ==> f O converse(f) = id(B)"; @@ -454,7 +454,7 @@ goal Perm.thy "!!f A B. [| converse(f): B->A; f: A->B |] ==> f : bij(A,B)"; by (asm_simp_tac (ZF_ss addsimps [fg_imp_bijective, comp_eq_id_iff, - left_inverse_lemma, right_inverse_lemma]) 1); + left_inverse_lemma, right_inverse_lemma]) 1); qed "invertible_imp_bijective"; (** Unions of functions -- cf similar theorems on func.ML **) @@ -463,9 +463,9 @@ "!!f g. [| f: surj(A,B); g: surj(C,D); A Int C = 0 |] ==> \ \ (f Un g) : surj(A Un C, B Un D)"; by (DEPTH_SOLVE_1 (eresolve_tac [fun_disjoint_apply1, fun_disjoint_apply2] 1 - ORELSE ball_tac 1 - ORELSE (rtac trans 1 THEN atac 2) - ORELSE step_tac (ZF_cs addIs [fun_disjoint_Un]) 1)); + ORELSE ball_tac 1 + ORELSE (rtac trans 1 THEN atac 2) + ORELSE step_tac (ZF_cs addIs [fun_disjoint_Un]) 1)); qed "surj_disjoint_Un"; (*A simple, high-level proof; the version for injections follows from it, @@ -551,6 +551,6 @@ (*Assumption ALL w:A. ALL x:A. f`w = f`x --> w=x would make asm_simp_tac loop using ZF_ss! But FOL_ss ignores the assumption...*) by (ALLGOALS (asm_simp_tac - (FOL_ss addsimps [fun_extend_apply2,fun_extend_apply1]))); + (FOL_ss addsimps [fun_extend_apply2,fun_extend_apply1]))); by (ALLGOALS (fast_tac (ZF_cs addIs [apply_type]))); qed "inj_extend"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/QPair.ML --- a/src/ZF/QPair.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/QPair.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/QPair.ML +(* Title: ZF/QPair.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge For QPair.thy. @@ -59,7 +59,7 @@ val QSigmaE2 = rule_by_tactic (REPEAT_FIRST (etac QPair_inject ORELSE' bound_hyp_subst_tac) - THEN prune_params_tac) + THEN prune_params_tac) (read_instantiate [("c","")] QSigmaE); qed_goal "QSigmaD1" thy " : QSigma(A,B) ==> a : A" @@ -143,8 +143,8 @@ qed "qsplitI"; val major::sigma::prems = goalw thy [qsplit_def] - "[| qsplit(R,z); z:QSigma(A,B); \ -\ !!x y. [| z = ; R(x,y) |] ==> P \ + "[| qsplit(R,z); z:QSigma(A,B); \ +\ !!x y. [| z = ; R(x,y) |] ==> P \ \ |] ==> P"; by (rtac (sigma RS QSigmaE) 1); by (cut_facts_tac [major] 1); @@ -177,7 +177,7 @@ (assume_tac 1) ]); val qconverse_cs = qpair_cs addSIs [qconverseI] - addSEs [qconverseD,qconverseE]; + addSEs [qconverseD,qconverseE]; qed_goal "qconverse_qconverse" thy "!!A B r. r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r" @@ -292,7 +292,7 @@ by (rtac (major RS qsumE) 1); by (ALLGOALS (etac ssubst)); by (ALLGOALS (asm_simp_tac (ZF_ss addsimps - (prems@[qcase_QInl,qcase_QInr])))); + (prems@[qcase_QInl,qcase_QInr])))); qed "qcase_type"; (** Rules for the Part primitive **) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/QUniv.ML --- a/src/ZF/QUniv.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/QUniv.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/quniv +(* Title: ZF/quniv ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge For quniv.thy. A small universe for lazy recursive types @@ -59,14 +59,14 @@ qed "Pow_univ_subset_quniv"; bind_thm ("univ_subset_into_quniv", - PowI RS (Pow_univ_subset_quniv RS subsetD)); + PowI RS (Pow_univ_subset_quniv RS subsetD)); bind_thm ("zero_in_quniv", zero_in_univ RS univ_into_quniv); bind_thm ("one_in_quniv", one_in_univ RS univ_into_quniv); bind_thm ("two_in_quniv", two_in_univ RS univ_into_quniv); bind_thm ("A_subset_quniv", - [A_subset_univ, univ_subset_quniv] MRS subset_trans); + [A_subset_univ, univ_subset_quniv] MRS subset_trans); val A_into_quniv = A_subset_quniv RS subsetD; @@ -110,13 +110,13 @@ qed "QSigma_quniv"; bind_thm ("QSigma_subset_quniv", - [QSigma_mono, QSigma_quniv] MRS subset_trans); + [QSigma_mono, QSigma_quniv] MRS subset_trans); (*The opposite inclusion*) goalw QUniv.thy [quniv_def,QPair_def] "!!A a b. : quniv(A) ==> a: quniv(A) & b: quniv(A)"; by (etac ([Transset_eclose RS Transset_univ, PowD] MRS - Transset_includes_summands RS conjE) 1); + Transset_includes_summands RS conjE) 1); by (REPEAT (ares_tac [conjI,PowI] 1)); qed "quniv_QPair_D"; @@ -149,13 +149,13 @@ (*** The natural numbers ***) bind_thm ("nat_subset_quniv", - [nat_subset_univ, univ_subset_quniv] MRS subset_trans); + [nat_subset_univ, univ_subset_quniv] MRS subset_trans); (* n:nat ==> n:quniv(A) *) bind_thm ("nat_into_quniv", (nat_subset_quniv RS subsetD)); bind_thm ("bool_subset_quniv", - [bool_subset_univ, univ_subset_quniv] MRS subset_trans); + [bool_subset_univ, univ_subset_quniv] MRS subset_trans); bind_thm ("bool_into_quniv", bool_subset_quniv RS subsetD); @@ -191,7 +191,7 @@ qed "Pair_in_Vfrom_D"; goal Univ.thy - "!!X. Transset(X) ==> \ + "!!X. Transset(X) ==> \ \ (a*b) Int Vfrom(X, succ(i)) <= (a Int Vfrom(X,i)) * (b Int Vfrom(X,i))"; by (fast_tac (ZF_cs addSDs [Pair_in_Vfrom_D]) 1); qed "product_Int_Vfrom_subset"; @@ -199,12 +199,12 @@ (*** Intersecting with Vfrom... ***) goalw QUniv.thy [QPair_def,sum_def] - "!!X. Transset(X) ==> \ + "!!X. Transset(X) ==> \ \ Int Vfrom(X, succ(i)) <= "; by (rtac (Int_Un_distrib RS ssubst) 1); by (rtac Un_mono 1); by (REPEAT (ares_tac [product_Int_Vfrom_subset RS subset_trans, - [Int_lower1, subset_refl] MRS Sigma_mono] 1)); + [Int_lower1, subset_refl] MRS Sigma_mono] 1)); qed "QPair_Int_Vfrom_succ_subset"; (**** "Take-lemma" rules for proving a=b by coinduction and c: quniv(A) ****) @@ -212,14 +212,14 @@ (*Rule for level i -- preserving the level, not decreasing it*) goalw QUniv.thy [QPair_def] - "!!X. Transset(X) ==> \ + "!!X. Transset(X) ==> \ \ Int Vfrom(X,i) <= "; by (etac (Transset_Vfrom RS Transset_sum_Int_subset) 1); qed "QPair_Int_Vfrom_subset"; (*[| a Int Vset(i) <= c; b Int Vset(i) <= d |] ==> Int Vset(i) <= *) bind_thm ("QPair_Int_Vset_subset_trans", - [Transset_0 RS QPair_Int_Vfrom_subset, QPair_mono] MRS subset_trans); + [Transset_0 RS QPair_Int_Vfrom_subset, QPair_mono] MRS subset_trans); goal QUniv.thy "!!i. [| Ord(i) \ @@ -233,5 +233,5 @@ by (rtac (succI1 RS UN_upper) 1); (*Limit(i) case*) by (asm_simp_tac (ZF_ss addsimps [Limit_Vfrom_eq, Int_UN_distrib, subset_refl, - UN_mono, QPair_Int_Vset_subset_trans]) 1); + UN_mono, QPair_Int_Vset_subset_trans]) 1); qed "QPair_Int_Vset_subset_UN"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ROOT.ML --- a/src/ZF/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ROOT +(* Title: ZF/ROOT ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Adds Zermelo-Fraenkel Set Theory to a database containing First-Order Logic. @@ -15,14 +15,14 @@ fun CHECK_SOLVED (Tactic tf) = Tactic (fn state => case Sequence.pull (tf state) of - None => error"DO_GOAL: tactic list failed" + None => error"DO_GOAL: tactic list failed" | Some(x,_) => - if has_fewer_prems 1 x then - Sequence.cons(x, Sequence.null) - else (writeln"DO_GOAL: unsolved goals!!"; - writeln"Final proof state was ..."; - print_goals (!goals_limit) x; - raise ERROR)); + if has_fewer_prems 1 x then + Sequence.cons(x, Sequence.null) + else (writeln"DO_GOAL: unsolved goals!!"; + writeln"Final proof state was ..."; + print_goals (!goals_limit) x; + raise ERROR)); fun DO_GOAL tfs = SELECT_GOAL (CHECK_SOLVED (EVERY1 tfs)); @@ -40,4 +40,4 @@ (*printing functions are inherited from FOL*) print_depth 8; -val ZF_build_completed = (); (*indicate successful build*) +val ZF_build_completed = (); (*indicate successful build*) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Rel.ML --- a/src/ZF/Rel.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Rel.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Rel.ML +(* Title: ZF/Rel.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge For Rel.thy. Relations in Zermelo-Fraenkel Set Theory diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Resid/Confluence.ML --- a/src/ZF/Resid/Confluence.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Resid/Confluence.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: Confluence.ML +(* Title: Confluence.ML ID: $Id$ - Author: Ole Rasmussen + Author: Ole Rasmussen Copyright 1995 University of Cambridge Logic Image: ZF *) @@ -27,7 +27,7 @@ goalw Confluence.thy [confluence_def,strip_def] "!!u.[|confluence(Spar_red1)|]==> strip"; by (resolve_tac [impI RS allI RS allI] 1); -by (eresolve_tac [par_red_induct] 1); +by (etac par_red_induct 1); by (fast_tac reduc_cs 1); by (fast_tac (ZF_cs addIs [Spar_red.trans]) 1); val strip_lemma_r = result(); @@ -36,7 +36,7 @@ goalw Confluence.thy [confluence_def,strip_def] "!!u.strip==> confluence(Spar_red)"; by (resolve_tac [impI RS allI RS allI] 1); -by (eresolve_tac [par_red_induct] 1); +by (etac par_red_induct 1); by (fast_tac reduc_cs 1); by (step_tac ZF_cs 1); by (dres_inst_tac [("x1","z")] (spec RS mp) 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Resid/Conversion.ML --- a/src/ZF/Resid/Conversion.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Resid/Conversion.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: Conversion.ML +(* Title: Conversion.ML ID: $Id$ - Author: Ole Rasmussen + Author: Ole Rasmussen Copyright 1995 University of Cambridge Logic Image: ZF *) @@ -14,8 +14,8 @@ goal Conversion.thy "!!u.m<--->n ==> n<--->m"; -by (eresolve_tac [conv_induct] 1); -by (eresolve_tac [conv1_induct] 1); +by (etac conv_induct 1); +by (etac conv1_induct 1); by (resolve_tac [Sconv.trans] 4); by (ALLGOALS(asm_simp_tac (conv_ss) )); val conv_sym = result(); @@ -26,13 +26,13 @@ goal Conversion.thy "!!u.m<--->n ==> EX p.(m --->p) & (n ---> p)"; -by (eresolve_tac [conv_induct] 1); -by (eresolve_tac [conv1_induct] 1); +by (etac conv_induct 1); +by (etac conv1_induct 1); by (fast_tac (reduc_cs addIs [red1D1,redD2]) 1); by (fast_tac (reduc_cs addIs [red1D1,redD2]) 1); by (fast_tac (reduc_cs addIs [red1D1,redD2]) 1); by (cut_facts_tac [confluence_beta_reduction] 1); -by (rewrite_goals_tac [confluence_def]); +by (rewtac confluence_def); by (step_tac ZF_cs 1); by (dres_inst_tac [("x4","n"),("x3","pa"),("x1","pb")] (spec RS spec RS mp RS spec RS mp) 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Resid/Cube.ML --- a/src/ZF/Resid/Cube.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Resid/Cube.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: Cube.ML +(* Title: Cube.ML ID: $Id$ - Author: Ole Rasmussen + Author: Ole Rasmussen Copyright 1995 University of Cambridge Logic Image: ZF *) @@ -8,8 +8,8 @@ open Cube; val prism_ss = (res1L_ss addsimps - [commutation,residuals_preserve_comp,sub_preserve_reg, - residuals_preserve_reg,sub_comp,sub_comp RS comp_sym]); + [commutation,residuals_preserve_comp,sub_preserve_reg, + residuals_preserve_reg,sub_comp,sub_comp RS comp_sym]); (* ------------------------------------------------------------------------- *) (* Prism theorem *) @@ -21,7 +21,7 @@ "!!u.v<==u ==> \ \ regular(u)-->(ALL w.w~v-->w~u--> \ \ w |> u = (w|>v) |> (u|>v))"; -by (eresolve_tac [sub_induct] 1); +by (etac sub_induct 1); by (asm_simp_tac (prism_ss) 1); by (asm_simp_tac (prism_ss ) 1); by (asm_simp_tac (prism_ss ) 1); @@ -38,7 +38,7 @@ "!!u.[|v <== u; regular(u); w~v|]==> \ \ w |> u = (w|>v) |> (u|>v)"; by (resolve_tac [prism_l RS mp RS spec RS mp RS mp ] 1); -by (resolve_tac [comp_trans] 4); +by (rtac comp_trans 4); by (assume_tac 4); by (ALLGOALS(asm_simp_tac (prism_ss))); val prism = result(); @@ -71,7 +71,7 @@ \ regular(vu) & (w|>v)~uv & regular(uv) "; by (forw_inst_tac [("x1","v")] (comp_sym RS comp_trans) 1); by (REPEAT(step_tac ZF_cs 1)); -by (resolve_tac [cube] 1); +by (rtac cube 1); by (ALLGOALS(asm_simp_tac (prism_ss addsimps [comp_sym_iff]))); val paving = result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Resid/ROOT.ML --- a/src/ZF/Resid/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Resid/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Resid/ROOT +(* Title: ZF/Resid/ROOT ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1995 University of Cambridge Executes the Residuals example. @@ -12,7 +12,7 @@ J. Functional Programming 4(3) 1994, 371-394. *) -ZF_build_completed; (*Make examples fail if ZF did*) +ZF_build_completed; (*Make examples fail if ZF did*) writeln"Root file for ZF/Resid"; proof_timing := true; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Resid/Redex.ML --- a/src/ZF/Resid/Redex.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Resid/Redex.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: Redex.ML +(* Title: Redex.ML ID: $Id$ - Author: Ole Rasmussen + Author: Ole Rasmussen Copyright 1995 University of Cambridge Logic Image: ZF *) @@ -28,7 +28,7 @@ val redex_rec_App = result(); val redexes_ss = (arith_ss addsimps - ([redex_rec_App,redex_rec_Fun,redex_rec_Var]@redexes.intrs)); + ([redex_rec_App,redex_rec_Fun,redex_rec_Var]@redexes.intrs)); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Resid/Reduction.ML --- a/src/ZF/Resid/Reduction.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Resid/Reduction.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: Reduction.ML +(* Title: Reduction.ML ID: $Id$ - Author: Ole Rasmussen + Author: Ole Rasmussen Copyright 1995 University of Cambridge Logic Image: ZF *) @@ -24,14 +24,14 @@ val reduc_cs = res_cs addIs (Sred1.intrs@[Sred.one_step,Sred.refl]@Spar_red1.intrs@ - [Spar_red.one_step,lambda.dom_subset RS subsetD, - unmark_type]@lambda.intrs@bool_typechecks) + [Spar_red.one_step,lambda.dom_subset RS subsetD, + unmark_type]@lambda.intrs@bool_typechecks) addSEs [Spar_red1.mk_cases redexes.con_defs "Fun(t) =1=> Fun(u)"]; val reduc_ss = term_ss addsimps (Sred1.intrs@[Sred.one_step,Sred.refl]@Spar_red1.intrs@ - [Spar_red.one_step,substL_type,redD1,redD2,par_redD1, - par_redD2,par_red1D2,unmark_type]); + [Spar_red.one_step,substL_type,redD1,redD2,par_redD1, + par_redD2,par_red1D2,unmark_type]); val reducL_ss = reduc_ss setloop (SELECT_GOAL (safe_tac reduc_cs)); @@ -45,21 +45,21 @@ (* ------------------------------------------------------------------------- *) goal Reduction.thy "!!u. m--->n ==> Fun(m) ---> Fun(n)"; -by (eresolve_tac [red_induct] 1); +by (etac red_induct 1); by (resolve_tac [Sred.trans] 3); by (ALLGOALS (asm_simp_tac reduc_ss )); val red_Fun = result(); goal Reduction.thy "!!u.[|n:lambda; m ---> m'|]==> Apl(m,n)--->Apl(m',n)"; -by (eresolve_tac [red_induct] 1); +by (etac red_induct 1); by (resolve_tac [Sred.trans] 3); by (ALLGOALS (asm_simp_tac reduc_ss )); val red_Apll = result(); goal Reduction.thy "!!u.[|n:lambda; m ---> m'|]==> Apl(n,m)--->Apl(n,m')"; -by (eresolve_tac [red_induct] 1); +by (etac red_induct 1); by (resolve_tac [Sred.trans] 3); by (ALLGOALS (asm_simp_tac reduc_ss )); val red_Aplr = result(); @@ -89,19 +89,19 @@ val refl_par_red1 = result(); goal Reduction.thy "!!u.m-1->n ==> m=1=>n"; -by (eresolve_tac [red1_induct] 1); +by (etac red1_induct 1); by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [refl_par_red1]) )); val red1_par_red1 = result(); goal Reduction.thy "!!u.m--->n ==> m===>n"; -by (eresolve_tac [red_induct] 1); +by (etac red_induct 1); by (resolve_tac [Spar_red.trans] 3); by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [refl_par_red1,red1_par_red1]) )); val red_par_red = result(); goal Reduction.thy "!!u.m===>n ==> m--->n"; -by (eresolve_tac [par_red_induct] 1); -by (eresolve_tac [par_red1_induct] 1); +by (etac par_red_induct 1); +by (etac par_red1_induct 1); by (resolve_tac [Sred.trans] 5); by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [red_Fun,red_beta,red_Apl]) )); val par_red_red = result(); @@ -113,7 +113,7 @@ goal Reduction.thy "!!u.m=1=>n ==> EX v. m|>v = n & m~v & regular(v)"; -by (eresolve_tac [par_red1_induct] 1); +by (etac par_red1_induct 1); by (step_tac ZF_cs 1); by (ALLGOALS (REPEAT o (resolve_tac [exI,conjI]))); by (TRYALL(resolve_tac [res_redex,res_App,res_Fun,res_Var])); @@ -137,7 +137,7 @@ \ unmark(subst_rec(u,v,k)) = subst_rec(unmark(u),unmark(v),k)"; by (eresolve_tac [redexes.induct] 1); by (ALLGOALS (asm_full_simp_tac - ((addsplit reduc_ss) addsimps [unmmark_lift_rec]))); + ((addsplit reduc_ss) addsimps [unmmark_lift_rec]))); val unmmark_subst_rec = result(); @@ -147,7 +147,7 @@ goal Reduction.thy "!!u.u~v ==> regular(v) --> unmark(u) =1=> unmark(u|>v)"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS (asm_full_simp_tac (reducL_ss addsimps [unmmark_subst_rec]) )); by (dres_inst_tac [("psi", "Fun(?u) =1=> ?w")] asm_rl 1); by (asm_full_simp_tac reducL_ss 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Resid/Residuals.ML --- a/src/ZF/Resid/Residuals.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Resid/Residuals.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: Residuals.ML +(* Title: Residuals.ML ID: $Id$ - Author: Ole Rasmussen + Author: Ole Rasmussen Copyright 1995 University of Cambridge Logic Image: ZF *) @@ -16,26 +16,26 @@ val res_cs = union_cs addIs (Sres.intrs@redexes.intrs@Sreg.intrs@ - [subst_type]@nat_typechecks) + [subst_type]@nat_typechecks) addSEs (redexes.free_SEs@ [Sres.mk_cases redexes.con_defs "residuals(Var(n),Var(n),v)", - Sres.mk_cases redexes.con_defs "residuals(Fun(t),Fun(u),v)", - Sres.mk_cases redexes.con_defs - "residuals(App(b, u1, u2), App(0, v1, v2),v)", - Sres.mk_cases redexes.con_defs + Sres.mk_cases redexes.con_defs "residuals(Fun(t),Fun(u),v)", + Sres.mk_cases redexes.con_defs + "residuals(App(b, u1, u2), App(0, v1, v2),v)", + Sres.mk_cases redexes.con_defs "residuals(App(b, u1, u2), App(1, Fun(v1), v2),v)", - Sres.mk_cases redexes.con_defs "residuals(Var(n),u,v)", - Sres.mk_cases redexes.con_defs "residuals(Fun(t),u,v)", - Sres.mk_cases redexes.con_defs "residuals(App(b, u1, u2), w,v)", - Sres.mk_cases redexes.con_defs "residuals(u,Var(n),v)", - Sres.mk_cases redexes.con_defs "residuals(u,Fun(t),v)", - Sres.mk_cases redexes.con_defs "residuals(w,App(b, u1, u2),v)", - Ssub.mk_cases redexes.con_defs "Var(n) <== u", - Ssub.mk_cases redexes.con_defs "Fun(n) <== u", - Ssub.mk_cases redexes.con_defs "u <== Fun(n)", - Ssub.mk_cases redexes.con_defs "App(1,Fun(t),a) <== u", - Ssub.mk_cases redexes.con_defs "App(0,t,a) <== u", - redexes.mk_cases redexes.con_defs "Fun(t):redexes"]); + Sres.mk_cases redexes.con_defs "residuals(Var(n),u,v)", + Sres.mk_cases redexes.con_defs "residuals(Fun(t),u,v)", + Sres.mk_cases redexes.con_defs "residuals(App(b, u1, u2), w,v)", + Sres.mk_cases redexes.con_defs "residuals(u,Var(n),v)", + Sres.mk_cases redexes.con_defs "residuals(u,Fun(t),v)", + Sres.mk_cases redexes.con_defs "residuals(w,App(b, u1, u2),v)", + Ssub.mk_cases redexes.con_defs "Var(n) <== u", + Ssub.mk_cases redexes.con_defs "Fun(n) <== u", + Ssub.mk_cases redexes.con_defs "u <== Fun(n)", + Ssub.mk_cases redexes.con_defs "App(1,Fun(t),a) <== u", + Ssub.mk_cases redexes.con_defs "App(0,t,a) <== u", + redexes.mk_cases redexes.con_defs "Fun(t):redexes"]); val res_ss = subst_ss addsimps (Sres.intrs); val resD1 =Sres.dom_subset RS subsetD RS SigmaD1; @@ -49,14 +49,14 @@ goal Residuals.thy "!!u.residuals(u,v,w) ==> ALL w1.residuals(u,v,w1) --> w1 = w"; -by (eresolve_tac [residuals_induct] 1); +by (etac residuals_induct 1); by (ALLGOALS (fast_tac res_cs )); val residuals_function = result(); val res_is_func = residuals_function RS spec RS mp; goal Residuals.thy "!!u.u~v ==> regular(v) --> (EX w.residuals(u,v,w))"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS (fast_tac res_cs)); val residuals_intro = result(); @@ -102,20 +102,20 @@ goal Residuals.thy "!!n.[|s~t; regular(t)|]==> regular(t) --> s |> t:redexes"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS (asm_full_simp_tac - (res_ss addsimps [res_Var,res_Fun,res_App,res_redex] - setloop (SELECT_GOAL (safe_tac res_cs))))); + (res_ss addsimps [res_Var,res_Fun,res_App,res_redex] + setloop (SELECT_GOAL (safe_tac res_cs))))); by (dres_inst_tac [("psi", "Fun(?u) |> ?v : redexes")] asm_rl 1); by (asm_full_simp_tac - (res_ss addsimps [res_Fun] - setloop (SELECT_GOAL (safe_tac res_cs))) 1); + (res_ss addsimps [res_Fun] + setloop (SELECT_GOAL (safe_tac res_cs))) 1); val resfunc_type = result(); val res1_ss = (res_ss addsimps - [res_Var,res_Fun,res_App,res_redex,lift_rec_preserve_comp, - lift_rec_preserve_reg,subst_rec_preserve_comp,resfunc_type, - subst_rec_preserve_reg]@redexes.free_iffs); + [res_Var,res_Fun,res_App,res_redex,lift_rec_preserve_comp, + lift_rec_preserve_reg,subst_rec_preserve_comp,resfunc_type, + subst_rec_preserve_reg]@redexes.free_iffs); val res1L_ss = res1_ss setloop (SELECT_GOAL (safe_tac res_cs)); @@ -125,20 +125,20 @@ goal Residuals.thy "!!u.u<==v ==> u~v"; -by (eresolve_tac [sub_induct] 1); +by (etac sub_induct 1); by (ALLGOALS (asm_simp_tac res1_ss)); val sub_comp = result(); goal Residuals.thy "!!u.u<==v ==> regular(v) --> regular(u)"; -by (eresolve_tac [sub_induct] 1); +by (etac sub_induct 1); by (ALLGOALS (asm_simp_tac res1L_ss)); val sub_preserve_reg = result(); goal Residuals.thy "!!u.[|u~v; k:nat|]==> regular(v)--> (ALL n:nat. \ \ lift_rec(u,n) |> lift_rec(v,n) = lift_rec(u |> v,n))"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (step_tac res_cs 1); by (ALLGOALS (asm_full_simp_tac ((addsplit res1_ss) addsimps [lift_subst]))); by (dres_inst_tac [("psi", "ALL x:nat.Fun(?u(x)) |> ?v(x) = ?w(x)")] asm_rl 1); @@ -149,7 +149,7 @@ "!!u.u1~u2 ==> ALL v1 v2.v1~v2 --> regular(v2) --> regular(u2) -->\ \ (ALL n:nat.subst_rec(v1,u1,n) |> subst_rec(v2,u2,n) = \ \ subst_rec(v1 |> v2, u1 |> u2,n))"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (step_tac (res_cs) 1); by (ALLGOALS (asm_full_simp_tac ((addsplit res1_ss) addsimps ([residuals_lift_rec])))); @@ -170,7 +170,7 @@ goal Residuals.thy "!!u.u~v ==> ALL w. u~w --> v~w --> regular(w) --> (u|>w) ~ (v|>w)"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS (asm_simp_tac (res1L_ss ))); by (dresolve_tac [spec RS mp RS mp RS mp] 1 THEN resolve_tac Scomp.intrs 1 THEN resolve_tac Scomp.intrs 2 @@ -182,7 +182,7 @@ goal Residuals.thy "!!u.u~v ==> regular(u) --> regular(v) --> regular(u|>v)"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (safe_tac res_cs); by (TRYALL (dres_inst_tac [("psi", "regular(Fun(?u) |> ?v)")] asm_rl)); by (ALLGOALS (asm_full_simp_tac res1L_ss)); @@ -194,20 +194,20 @@ goal Residuals.thy "!!u.u~v ==> v ~ (u un v)"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS (asm_full_simp_tac res_ss)); val union_preserve_comp = result(); goal Residuals.thy "!!u.u ~ v ==> regular(v) --> u|>v = (u un v)|>v"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (safe_tac res_cs); by (dres_inst_tac [("psi", "Fun(?u) |> ?v = ?w")] asm_rl 3); by (ALLGOALS (asm_full_simp_tac (res1L_ss addsimps - [union_preserve_comp,comp_sym_iff]))); + [union_preserve_comp,comp_sym_iff]))); by (asm_full_simp_tac (res1_ss addsimps - [union_preserve_comp RS comp_sym, - comp_sym RS union_preserve_comp RS comp_sym]) 1); + [union_preserve_comp RS comp_sym, + comp_sym RS union_preserve_comp RS comp_sym]) 1); val preservation1 = result(); val preservation = preservation1 RS mp; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Resid/SubUnion.ML --- a/src/ZF/Resid/SubUnion.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Resid/SubUnion.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: SubUnion.ML +(* Title: SubUnion.ML ID: $Id$ - Author: Ole Rasmussen + Author: Ole Rasmussen Copyright 1995 University of Cambridge Logic Image: ZF *) @@ -45,23 +45,23 @@ val union_ss = redexes_ss addsimps (Ssub.intrs@bool_simps@bool_typechecks@ - Sreg.intrs@Scomp.intrs@ - [or_1 RSN (3,or_commute RS trans), - or_0 RSN (3,or_commute RS trans), - union_App,union_Fun,union_Var,compD2,compD1,regD]); + Sreg.intrs@Scomp.intrs@ + [or_1 RSN (3,or_commute RS trans), + or_0 RSN (3,or_commute RS trans), + union_App,union_Fun,union_Var,compD2,compD1,regD]); val union_cs = (ZF_cs addIs Scomp.intrs addSEs - [Sreg.mk_cases redexes.con_defs "regular(App(b,f,a))", - Sreg.mk_cases redexes.con_defs "regular(Fun(b))", - Sreg.mk_cases redexes.con_defs "regular(Var(b))", - Scomp.mk_cases redexes.con_defs "Fun(u) ~ Fun(t)", - Scomp.mk_cases redexes.con_defs "u ~ Fun(t)", - Scomp.mk_cases redexes.con_defs "u ~ Var(n)", - Scomp.mk_cases redexes.con_defs "u ~ App(b,t,a)", - Scomp.mk_cases redexes.con_defs "Fun(t) ~ v", - Scomp.mk_cases redexes.con_defs "App(b,f,a) ~ v", - Scomp.mk_cases redexes.con_defs "Var(n) ~ u" - ]); + [Sreg.mk_cases redexes.con_defs "regular(App(b,f,a))", + Sreg.mk_cases redexes.con_defs "regular(Fun(b))", + Sreg.mk_cases redexes.con_defs "regular(Var(b))", + Scomp.mk_cases redexes.con_defs "Fun(u) ~ Fun(t)", + Scomp.mk_cases redexes.con_defs "u ~ Fun(t)", + Scomp.mk_cases redexes.con_defs "u ~ Var(n)", + Scomp.mk_cases redexes.con_defs "u ~ App(b,t,a)", + Scomp.mk_cases redexes.con_defs "Fun(t) ~ v", + Scomp.mk_cases redexes.con_defs "App(b,f,a) ~ v", + Scomp.mk_cases redexes.con_defs "Var(n) ~ u" + ]); @@ -76,7 +76,7 @@ goal SubUnion.thy "!!u.u ~ v ==> v ~ u"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS(fast_tac union_cs)); val comp_sym = result(); @@ -88,7 +88,7 @@ goal SubUnion.thy "!!u.u ~ v ==> ALL w.v ~ w-->u ~ w"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS(fast_tac union_cs)); val comp_trans1 = result(); @@ -103,21 +103,21 @@ goal SubUnion.thy "!!u.u ~ v ==> u <== (u un v)"; -by (eresolve_tac [comp_induct] 1); -by (eresolve_tac [boolE] 3); +by (etac comp_induct 1); +by (etac boolE 3); by (ALLGOALS(asm_full_simp_tac union_ss)); val union_l = result(); goal SubUnion.thy "!!u.u ~ v ==> v <== (u un v)"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (eres_inst_tac [("c","b2")] boolE 3); by (ALLGOALS(asm_full_simp_tac union_ss)); val union_r = result(); goal SubUnion.thy "!!u.u ~ v ==> u un v = v un u"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS(asm_simp_tac (union_ss addsimps [or_commute]))); val union_sym = result(); @@ -127,9 +127,9 @@ goal SubUnion.thy "!!u.u ~ v ==> regular(u)-->regular(v)-->regular(u un v)"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS(asm_full_simp_tac - (union_ss setloop(SELECT_GOAL (safe_tac union_cs))))); + (union_ss setloop(SELECT_GOAL (safe_tac union_cs))))); by (dres_inst_tac [("psi", "regular(Fun(?u) un ?v)")] asm_rl 1); by (asm_full_simp_tac union_ss 1); val union_preserve_regular = result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Resid/Substitution.ML --- a/src/ZF/Resid/Substitution.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Resid/Substitution.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: Substitution.ML +(* Title: Substitution.ML ID: $Id$ - Author: Ole Rasmussen + Author: Ole Rasmussen Copyright 1995 University of Cambridge Logic Image: ZF *) @@ -19,26 +19,26 @@ val succ_pred = prove_goal Arith.thy "!!i.[|j:nat; i:nat|]==> i < j --> succ(j #- 1) = j" (fn prems =>[(etac nat_induct 1), - (fast_tac lt_cs 1), - (asm_simp_tac arith_ss 1)]); + (fast_tac lt_cs 1), + (asm_simp_tac arith_ss 1)]); goal Arith.thy "!!i.[|succ(x) x < n#-1 "; -by (resolve_tac [succ_leE] 1); +by (rtac succ_leE 1); by (forward_tac [nat_into_Ord RS le_refl RS lt_trans] 1 THEN assume_tac 1); by (asm_simp_tac (arith_ss addsimps [succ_pred]) 1); val lt_pred = result(); goal Arith.thy "!!i.[|n < succ(x); p n#-1 < x "; -by (resolve_tac [succ_leE] 1); +by (rtac succ_leE 1); by (asm_simp_tac (arith_ss addsimps [succ_pred]) 1); val gt_pred = result(); val lift_ss = (union_ss addsimps - [add_0_right,add_succ_right,nat_into_Ord, - not_lt_iff_le,if_P,if_not_P]); + [add_0_right,add_succ_right,nat_into_Ord, + not_lt_iff_le,if_P,if_not_P]); (* ------------------------------------------------------------------------- *) @@ -111,7 +111,7 @@ val subst_App = result(); fun addsplit ss = (ss setloop (split_tac [expand_if]) - addsimps [lift_rec_Var,subst_Var]); + addsimps [lift_rec_Var,subst_Var]); goal Substitution.thy @@ -127,15 +127,15 @@ by (eresolve_tac [redexes.induct] 1); by (ALLGOALS(asm_full_simp_tac ((addsplit lift_ss) addsimps [subst_Fun,subst_App, - lift_rec_type]))); + lift_rec_type]))); val subst_type_a = result(); val subst_type = subst_type_a RS bspec RS bspec; val subst_ss = (lift_ss addsimps - [subst_eq,subst_gt,subst_lt,subst_Fun,subst_App,subst_type, - lift_rec_le,lift_rec_gt,lift_rec_Fun,lift_rec_App, - lift_rec_type]); + [subst_eq,subst_gt,subst_lt,subst_Fun,subst_App,subst_type, + lift_rec_le,lift_rec_gt,lift_rec_Fun,lift_rec_App, + lift_rec_type]); (* ------------------------------------------------------------------------- *) @@ -201,7 +201,7 @@ by (eres_inst_tac [("i","x")] leE 1); by (forward_tac [lt_trans1] 1 THEN assume_tac 1); by (ALLGOALS(asm_full_simp_tac - (subst_ss addsimps [succ_pred,leI,gt_pred]))); + (subst_ss addsimps [succ_pred,leI,gt_pred]))); by (hyp_subst_tac 1); by (asm_full_simp_tac (subst_ss addsimps [leI]) 1); by (excluded_middle_tac "na < xa" 1); @@ -234,7 +234,7 @@ by (excluded_middle_tac "na le succ(xa)" 1); by (asm_full_simp_tac (subst_ss) 1); by (forward_tac [nat_into_Ord RS le_refl RS lt_trans] 1 THEN assume_tac 1); -by (eresolve_tac [leE] 1); +by (etac leE 1); by (asm_simp_tac (subst_ss addsimps [succ_pred,lt_pred]) 2); by (forward_tac [succ_leI RS lt_trans] 1 THEN assume_tac 1); by (forw_inst_tac [("i","x")] @@ -267,16 +267,16 @@ goal Substitution.thy "!!n.[|n:nat; u ~ v|]==> ALL m:nat.lift_rec(u,m) ~ lift_rec(v,m)"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS(asm_simp_tac (subst_ss addsimps [comp_refl]))); val lift_rec_preserve_comp = result(); goal Substitution.thy "!!n.u2 ~ v2 ==> ALL m:nat.ALL u1:redexes.ALL v1:redexes.\ \ u1 ~ v1--> subst_rec(u1,u2,m) ~ subst_rec(v1,v2,m)"; -by (eresolve_tac [comp_induct] 1); +by (etac comp_induct 1); by (ALLGOALS(asm_full_simp_tac ((addsplit subst_ss) addsimps - ([lift_rec_preserve_comp,comp_refl])))); + ([lift_rec_preserve_comp,comp_refl])))); val subst_rec_preserve_comp = result(); goal Substitution.thy @@ -290,5 +290,5 @@ \ ALL m:nat.ALL u:redexes.regular(u)-->regular(subst_rec(u,v,m))"; by (eresolve_tac [Sreg.induct] 1); by (ALLGOALS(asm_full_simp_tac ((addsplit subst_ss) addsimps - [lift_rec_preserve_reg]))); + [lift_rec_preserve_reg]))); val subst_rec_preserve_reg = result(); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Resid/Terms.ML --- a/src/ZF/Resid/Terms.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Resid/Terms.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: Terms.ML +(* Title: Terms.ML ID: $Id$ - Author: Ole Rasmussen + Author: Ole Rasmussen Copyright 1995 University of Cambridge Logic Image: ZF *) @@ -33,7 +33,7 @@ val term_ss = res1_ss addsimps ([unmark_App,unmark_Fun,unmark_Var, - lambda.dom_subset RS subsetD]@lambda.intrs); + lambda.dom_subset RS subsetD]@lambda.intrs); (* ------------------------------------------------------------------------- *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Sum.ML --- a/src/ZF/Sum.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Sum.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Sum +(* Title: ZF/Sum ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge Disjoint sums in Zermelo-Fraenkel Set Theory @@ -134,7 +134,7 @@ by (rtac (major RS sumE) 1); by (ALLGOALS (etac ssubst)); by (ALLGOALS (asm_simp_tac (ZF_ss addsimps - (prems@[case_Inl,case_Inr])))); + (prems@[case_Inl,case_Inr])))); qed "case_type"; goal Sum.thy @@ -150,16 +150,16 @@ qed "expand_case"; val major::prems = goal Sum.thy - "[| z: A+B; \ -\ !!x. x:A ==> c(x)=c'(x); \ -\ !!y. y:B ==> d(y)=d'(y) \ + "[| z: A+B; \ +\ !!x. x:A ==> c(x)=c'(x); \ +\ !!y. y:B ==> d(y)=d'(y) \ \ |] ==> case(c,d,z) = case(c',d',z)"; by (resolve_tac [major RS sumE] 1); by (ALLGOALS (asm_simp_tac (ZF_ss addsimps ([case_Inl, case_Inr] @ prems)))); qed "case_cong"; val [major] = goal Sum.thy - "z: A+B ==> \ + "z: A+B ==> \ \ case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) = \ \ case(%x. c(c'(x)), %y. d(d'(y)), z)"; by (resolve_tac [major RS sumE] 1); @@ -167,8 +167,8 @@ qed "case_case"; val sum_ss = ZF_ss addsimps [InlI, InrI, Inl_iff, Inr_iff, - Inl_Inr_iff, Inr_Inl_iff, - case_Inl, case_Inr]; + Inl_Inr_iff, Inr_Inl_iff, + case_Inl, case_Inr]; (*** More rules for Part(A,h) ***) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Trancl.ML --- a/src/ZF/Trancl.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Trancl.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/trancl.ML +(* Title: ZF/trancl.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge For trancl.thy. Transitive closure of a relation @@ -17,7 +17,7 @@ val [prem] = goalw Trancl.thy [rtrancl_def] "r<=s ==> r^* <= s^*"; by (rtac lfp_mono 1); by (REPEAT (resolve_tac [rtrancl_bnd_mono, prem, subset_refl, id_mono, - comp_mono, Un_mono, field_mono, Sigma_mono] 1)); + comp_mono, Un_mono, field_mono, Sigma_mono] 1)); qed "rtrancl_mono"; (* r^* = id(field(r)) Un ( r O r^* ) *) @@ -56,7 +56,7 @@ goal Trancl.thy "field(r^*) = field(r)"; by (fast_tac (eq_cs addIs [r_into_rtrancl] - addSDs [rtrancl_type RS subsetD]) 1); + addSDs [rtrancl_type RS subsetD]) 1); qed "rtrancl_field"; @@ -75,9 +75,9 @@ Tried adding the typing hypotheses y,z:field(r), but these caused expensive case splits!*) val major::prems = goal Trancl.thy - "[| : r^*; \ -\ P(a); \ -\ !!y z.[| : r^*; : r; P(y) |] ==> P(z) \ + "[| : r^*; \ +\ P(a); \ +\ !!y z.[| : r^*; : r; P(y) |] ==> P(z) \ \ |] ==> P(b)"; (*by induction on this formula*) by (subgoal_tac "ALL y. = --> P(y)" 1); @@ -97,8 +97,8 @@ (*elimination of rtrancl -- by induction on a special formula*) val major::prems = goal Trancl.thy - "[| : r^*; (a=b) ==> P; \ -\ !!y.[| : r^*; : r |] ==> P |] \ + "[| : r^*; (a=b) ==> P; \ +\ !!y.[| : r^*; : r |] ==> P |] \ \ ==> P"; by (subgoal_tac "a = b | (EX y. : r^* & : r)" 1); (*see HOL/trancl*) @@ -151,9 +151,9 @@ (*Nice induction rule for trancl*) val major::prems = goal Trancl.thy - "[| : r^+; \ -\ !!y. [| : r |] ==> P(y); \ -\ !!y z.[| : r^+; : r; P(y) |] ==> P(z) \ + "[| : r^+; \ +\ !!y. [| : r |] ==> P(y); \ +\ !!y z.[| : r^+; : r; P(y) |] ==> P(z) \ \ |] ==> P(b)"; by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1); (*by induction on this formula*) @@ -168,7 +168,7 @@ val major::prems = goal Trancl.thy "[| : r^+; \ \ : r ==> P; \ -\ !!y.[| : r^+; : r |] ==> P \ +\ !!y.[| : r^+; : r |] ==> P \ \ |] ==> P"; by (subgoal_tac " : r | (EX y. : r^+ & : r)" 1); by (fast_tac (ZF_cs addIs prems) 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Univ.ML --- a/src/ZF/Univ.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Univ.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Univ +(* Title: ZF/Univ ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge The cumulative hierarchy and a small universe for recursive types @@ -128,7 +128,7 @@ goal Univ.thy "!!i. Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"; by (rtac (Vfrom RS trans) 1); by (rtac (succI1 RS RepFunI RS Union_upper RSN - (2, equalityI RS subst_context)) 1); + (2, equalityI RS subst_context)) 1); by (rtac UN_least 1); by (rtac (subset_refl RS Vfrom_mono RS Pow_mono) 1); by (etac (ltI RS le_imp_subset) 1); @@ -180,8 +180,8 @@ qed "Limit_VfromI"; val prems = goal Univ.thy - "[| a: Vfrom(A,i); Limit(i); \ -\ !!x. [| x R \ + "[| a: Vfrom(A,i); Limit(i); \ +\ !!x. [| x R \ \ |] ==> R"; by (rtac (Limit_Vfrom_eq RS equalityD1 RS subsetD RS UN_E) 1); by (REPEAT (ares_tac (prems @ [ltI, Limit_is_Ord]) 1)); @@ -230,10 +230,10 @@ qed "product_VLimit"; bind_thm ("Sigma_subset_VLimit", - [Sigma_mono, product_VLimit] MRS subset_trans); + [Sigma_mono, product_VLimit] MRS subset_trans); bind_thm ("nat_subset_VLimit", - [nat_le_Limit RS le_imp_subset, i_subset_Vfrom] MRS subset_trans); + [nat_le_Limit RS le_imp_subset, i_subset_Vfrom] MRS subset_trans); goal Univ.thy "!!i. [| n: nat; Limit(i) |] ==> n : Vfrom(A,i)"; by (REPEAT (ares_tac [nat_subset_VLimit RS subsetD] 1)); @@ -271,7 +271,7 @@ by (eps_ind_tac "i" 1); by (rtac (Vfrom RS ssubst) 1); by (fast_tac (ZF_cs addSIs [Transset_Union_family, Transset_Un, - Transset_Pow]) 1); + Transset_Pow]) 1); qed "Transset_Vfrom"; goal Univ.thy "!!A i. Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))"; @@ -303,9 +303,9 @@ (*General theorem for membership in Vfrom(A,i) when i is a limit ordinal*) val [aprem,bprem,limiti,step] = goal Univ.thy - "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); \ + "[| a: Vfrom(A,i); b: Vfrom(A,i); Limit(i); \ \ !!x y j. [| j EX k. h(x,y): Vfrom(A,k) & k \ +\ |] ==> EX k. h(x,y): Vfrom(A,k) & k \ \ h(a,b) : Vfrom(A,i)"; (*Infer that a, b occur at ordinals x,xa < i.*) by (rtac ([aprem,limiti] MRS Limit_VfromE) 1); @@ -334,7 +334,7 @@ \ a*b : Vfrom(A,i)"; by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1); by (REPEAT (ares_tac [exI, conjI, prod_in_Vfrom, transset, - limiti RS Limit_has_succ] 1)); + limiti RS Limit_has_succ] 1)); qed "prod_in_VLimit"; (** Disjoint sums, aka Quine ordered pairs **) @@ -346,7 +346,7 @@ by (rtac subset_mem_Vfrom 1); by (rewtac Transset_def); by (fast_tac (ZF_cs addIs [zero_in_Vfrom, Pair_in_Vfrom, - i_subset_Vfrom RS subsetD]) 1); + i_subset_Vfrom RS subsetD]) 1); qed "sum_in_Vfrom"; val [aprem,bprem,limiti,transset] = goal Univ.thy @@ -354,7 +354,7 @@ \ a+b : Vfrom(A,i)"; by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1); by (REPEAT (ares_tac [exI, conjI, sum_in_Vfrom, transset, - limiti RS Limit_has_succ] 1)); + limiti RS Limit_has_succ] 1)); qed "sum_in_VLimit"; (** function space! **) @@ -379,7 +379,7 @@ \ a->b : Vfrom(A,i)"; by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1); by (REPEAT (ares_tac [exI, conjI, fun_in_Vfrom, transset, - limiti RS Limit_has_succ] 1)); + limiti RS Limit_has_succ] 1)); qed "fun_in_VLimit"; @@ -437,10 +437,10 @@ by (rtac equalityI 1); by (safe_tac ZF_cs); by (EVERY' [rtac UN_I, - etac (i_subset_Vfrom RS subsetD), - etac (Ord_in_Ord RS rank_of_Ord RS ssubst), - assume_tac, - rtac succI1] 3); + etac (i_subset_Vfrom RS subsetD), + etac (Ord_in_Ord RS rank_of_Ord RS ssubst), + assume_tac, + rtac succI1] 3); by (REPEAT (eresolve_tac [asm_rl, VsetD RS ltD, Ord_trans] 1)); qed "rank_Vset"; @@ -454,7 +454,7 @@ val [iprem] = goal Univ.thy "[| !!i. Ord(i) ==> a Int Vset(i) <= b |] ==> a <= b"; by (rtac ([subset_refl, arg_subset_Vset_rank] MRS - Int_greatest RS subset_trans) 1); + Int_greatest RS subset_trans) 1); by (rtac (Ord_rank RS iprem) 1); qed "Int_Vset_subset"; @@ -473,7 +473,7 @@ goalw Univ.thy [Vrec_def] "Vrec(a,H) = H(a, lam x:Vset(rank(a)). Vrec(x,H))"; by (rtac (transrec RS ssubst) 1); by (simp_tac (ZF_ss addsimps [Ord_rank, Ord_succ, VsetD RS ltD RS beta, - VsetI RS beta, le_refl]) 1); + VsetI RS beta, le_refl]) 1); qed "Vrec"; (*This form avoids giant explosions in proofs. NOTE USE OF == *) @@ -508,8 +508,8 @@ qed "subset_univ_eq_Int"; val [aprem, iprem] = goal Univ.thy - "[| a <= univ(X); \ -\ !!i. i:nat ==> a Int Vfrom(X,i) <= b \ + "[| a <= univ(X); \ +\ !!i. i:nat ==> a Int Vfrom(X,i) <= b \ \ |] ==> a <= b"; by (rtac (aprem RS subset_univ_eq_Int RS ssubst) 1); by (rtac UN_least 1); @@ -615,19 +615,19 @@ goal Univ.thy "!!i. [| b: Fin(Vfrom(A,i)); Limit(i) |] ==> EX j. b <= Vfrom(A,j) & j Fin(Vfrom(A,i)) <= Vfrom(A,i)"; by (rtac subsetI 1); -by (dresolve_tac [Fin_Vfrom_lemma] 1); +by (dtac Fin_Vfrom_lemma 1); by (safe_tac ZF_cs); by (resolve_tac [Vfrom RS ssubst] 1); by (fast_tac (ZF_cs addSDs [ltD]) 1); @@ -645,7 +645,7 @@ "!!i. [| n: nat; Limit(i) |] ==> n -> Vfrom(A,i) <= Vfrom(A,i)"; by (eresolve_tac [nat_fun_subset_Fin RS subset_trans] 1); by (REPEAT (ares_tac [Fin_subset_VLimit, Sigma_subset_VLimit, - nat_subset_VLimit, subset_refl] 1)); + nat_subset_VLimit, subset_refl] 1)); val nat_fun_VLimit = result(); bind_thm ("nat_fun_subset_VLimit", [Pi_mono, nat_fun_VLimit] MRS subset_trans); @@ -672,7 +672,7 @@ goal Univ.thy "!!i. [| W <= Vfrom(A,i); Limit(i) |] ==> W -||> Vfrom(A,i) <= Vfrom(A,i)"; by (eresolve_tac [subset_refl RSN (2, FiniteFun_mono) RS subset_trans] 1); -by (eresolve_tac [FiniteFun_VLimit1] 1); +by (etac FiniteFun_VLimit1 1); val FiniteFun_VLimit = result(); goalw Univ.thy [univ_def] diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/WF.ML --- a/src/ZF/WF.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/WF.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/wf.ML +(* Title: ZF/wf.ML ID: $Id$ - Author: Tobias Nipkow and Lawrence C Paulson + Author: Tobias Nipkow and Lawrence C Paulson Copyright 1992 University of Cambridge For wf.thy. Well-founded Recursion @@ -89,8 +89,8 @@ (*Perform induction on i, then prove the wf(r) subgoal using prems. *) fun wf_ind_tac a prems i = EVERY [res_inst_tac [("a",a)] wf_induct i, - rename_last_tac a ["1"] (i+1), - ares_tac prems i]; + rename_last_tac a ["1"] (i+1), + ares_tac prems i]; (*The form of this rule is designed to match wfI*) val wfr::amem::prems = goal WF.thy @@ -109,8 +109,8 @@ qed "field_Int_square"; val wfr::amem::prems = goalw WF.thy [wf_on_def] - "[| wf[A](r); a:A; \ -\ !!x.[| x: A; ALL y:A. : r --> P(y) |] ==> P(x) \ + "[| wf[A](r); a:A; \ +\ !!x.[| x: A; ALL y:A. : r --> P(y) |] ==> P(x) \ \ |] ==> P(a)"; by (rtac ([wfr, amem, field_Int_square] MRS wf_induct2) 1); by (REPEAT (ares_tac prems 1)); @@ -119,8 +119,8 @@ fun wf_on_ind_tac a prems i = EVERY [res_inst_tac [("a",a)] wf_on_induct i, - rename_last_tac a ["1"] (i+2), - REPEAT (ares_tac prems i)]; + rename_last_tac a ["1"] (i+2), + REPEAT (ares_tac prems i)]; (*If r allows well-founded induction then wf(r)*) val [subs,indhyp] = goal WF.thy @@ -224,7 +224,7 @@ resolve_tac (TrueI::refl::hyps) ORELSE' (cut_facts_tac hyps THEN' DEPTH_SOLVE_1 o (ares_tac [TrueI, ballI] ORELSE' - eresolve_tac [underD, transD, spec RS mp])); + eresolve_tac [underD, transD, spec RS mp])); (*** NOTE! some simplifications need a different solver!! ***) val wf_super_ss = ZF_ss setsolver indhyp_tac; @@ -249,7 +249,7 @@ by (etac is_recfun_type 1); by (ALLGOALS (asm_simp_tac (wf_super_ss addsimps - [ [wfr,transr,recf,recg] MRS is_recfun_equal ]))); + [ [wfr,transr,recf,recg] MRS is_recfun_equal ]))); qed "is_recfun_cut"; (*** Main Existence Lemma ***) @@ -302,7 +302,7 @@ \ wftrec(r,a,H) = H(a, lam x: r-``{a}. wftrec(r,x,H))"; by (rtac (rewrite_rule [is_recfun_def] unfold_the_recfun RS ssubst) 1); by (ALLGOALS (asm_simp_tac - (ZF_ss addsimps [vimage_singleton_iff RS iff_sym, the_recfun_cut]))); + (ZF_ss addsimps [vimage_singleton_iff RS iff_sym, the_recfun_cut]))); qed "wftrec"; (** Removal of the premise trans(r) **) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ZF.ML --- a/src/ZF/ZF.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ZF.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ZF.ML +(* Title: ZF/ZF.ML ID: $Id$ - Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory + Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory Copyright 1994 University of Cambridge Basic introduction and elimination rules for Zermelo-Fraenkel Set Theory @@ -209,9 +209,9 @@ let val substprems = prems RL [subst, ssubst] and iffprems = prems RL [iffD1,iffD2] in [ (rtac equalityI 1), - (REPEAT (eresolve_tac (substprems@[asm_rl, ReplaceE, spec RS mp]) 1 - ORELSE resolve_tac [subsetI, ReplaceI] 1 - ORELSE (resolve_tac iffprems 1 THEN assume_tac 2))) ] + (REPEAT (eresolve_tac (substprems@[asm_rl, ReplaceE, spec RS mp]) 1 + ORELSE resolve_tac [subsetI, ReplaceI] 1 + ORELSE (resolve_tac iffprems 1 THEN assume_tac 2))) ] end); (*** Rules for RepFun ***) @@ -248,7 +248,7 @@ qed_goalw "separation" ZF.thy [Collect_def] "a : {x:A. P(x)} <-> a:A & P(a)" (fn _=> [ (fast_tac (FOL_cs addIs [bexI,ReplaceI] - addSEs [bexE,ReplaceE]) 1) ]); + addSEs [bexE,ReplaceE]) 1) ]); qed_goal "CollectI" ZF.thy "[| a:A; P(a) |] ==> a : {x:A. P(x)}" @@ -296,7 +296,7 @@ qed_goalw "Inter_iff" ZF.thy [Inter_def,Ball_def] "A : Inter(C) <-> (ALL x:C. A: x) & (EX x. x:C)" (fn _=> [ (rtac (separation RS iff_trans) 1), - (fast_tac (FOL_cs addIs [UnionI] addSEs [UnionE]) 1) ]); + (fast_tac (FOL_cs addIs [UnionI] addSEs [UnionE]) 1) ]); (* Intersection is well-behaved only if the family is non-empty! *) qed_goalw "InterI" ZF.thy [Inter_def] @@ -349,8 +349,8 @@ qed_goal "INT_iff" ZF.thy "b : (INT x:A. B(x)) <-> (ALL x:A. b : B(x)) & (EX x. x:A)" (fn _=> [ (simp_tac (FOL_ss addsimps [Inter_def, Ball_def, Bex_def, - separation, Union_iff, RepFun_iff]) 1), - (fast_tac FOL_cs 1) ]); + separation, Union_iff, RepFun_iff]) 1), + (fast_tac FOL_cs 1) ]); qed_goal "INT_I" ZF.thy "[| !!x. x: A ==> b: B(x); a: A |] ==> b: (INT x:A. B(x))" @@ -419,7 +419,7 @@ addSIs [ballI, InterI, CollectI, PowI, empty_subsetI] addIs [bexI, UnionI, ReplaceI, RepFunI] addSEs [bexE, make_elim PowD, UN_E, UnionE, ReplaceE2, RepFunE, - CollectE, emptyE] + CollectE, emptyE] addEs [rev_ballE, InterD, make_elim InterD, subsetD]; (*Standard claset*) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/Zorn.ML --- a/src/ZF/Zorn.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/Zorn.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/Zorn.ML +(* Title: ZF/Zorn.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Proofs from the paper @@ -57,14 +57,14 @@ (*Perform induction on n, then prove the major premise using prems. *) fun TFin_ind_tac a prems i = EVERY [res_inst_tac [("n",a)] TFin_induct i, - rename_last_tac a ["1"] (i+1), - rename_last_tac a ["2"] (i+2), - ares_tac prems i]; + rename_last_tac a ["1"] (i+1), + rename_last_tac a ["2"] (i+2), + ares_tac prems i]; (*** Section 3. Some Properties of the Transfinite Construction ***) bind_thm ("increasing_trans", - TFin_is_subset RSN (3, increasingD2 RSN (2,subset_trans))); + TFin_is_subset RSN (3, increasingD2 RSN (2,subset_trans))); (*Lemma 1 of section 3.1*) val major::prems = goal Zorn.thy @@ -73,7 +73,7 @@ \ |] ==> n<=m | next`m<=n"; by (cut_facts_tac prems 1); by (rtac (major RS TFin_induct) 1); -by (etac Union_lemma0 2); (*or just fast_tac ZF_cs*) +by (etac Union_lemma0 2); (*or just fast_tac ZF_cs*) by (fast_tac (subset_cs addIs [increasing_trans]) 1); qed "TFin_linear_lemma1"; @@ -120,7 +120,7 @@ by (rtac (TFin_linear_lemma2 RSN (3,TFin_linear_lemma1) RS disjE) 1); by (REPEAT (assume_tac 1) THEN etac disjI2 1); by (fast_tac (subset_cs addIs [increasingD2 RS subset_trans, - TFin_is_subset]) 1); + TFin_is_subset]) 1); qed "TFin_subset_linear"; @@ -168,8 +168,8 @@ qed "maxchain_subset_chain"; goal Zorn.thy - "!!S. [| ch : (PROD X:Pow(chain(S)) - {0}. X); \ -\ X : chain(S); X ~: maxchain(S) \ + "!!S. [| ch : (PROD X:Pow(chain(S)) - {0}. X); \ +\ X : chain(S); X ~: maxchain(S) \ \ |] ==> ch ` super(S,X) : super(S,X)"; by (etac apply_type 1); by (rewrite_goals_tac [super_def, maxchain_def]); @@ -177,8 +177,8 @@ qed "choice_super"; goal Zorn.thy - "!!S. [| ch : (PROD X:Pow(chain(S)) - {0}. X); \ -\ X : chain(S); X ~: maxchain(S) \ + "!!S. [| ch : (PROD X:Pow(chain(S)) - {0}. X); \ +\ X : chain(S); X ~: maxchain(S) \ \ |] ==> ch ` super(S,X) ~= X"; by (rtac notI 1); by (dtac choice_super 1); @@ -189,8 +189,8 @@ (*This justifies Definition 4.4*) goal Zorn.thy - "!!S. ch: (PROD X: Pow(chain(S))-{0}. X) ==> \ -\ EX next: increasing(S). ALL X: Pow(S). \ + "!!S. ch: (PROD X: Pow(chain(S))-{0}. X) ==> \ +\ EX next: increasing(S). ALL X: Pow(S). \ \ next`X = if(X: chain(S)-maxchain(S), ch`super(S,X), X)"; by (rtac bexI 1); by (rtac ballI 1); @@ -201,8 +201,8 @@ by (rtac lam_type 1); by (asm_simp_tac (ZF_ss setloop split_tac [expand_if]) 1); by (fast_tac (ZF_cs addSIs [super_subset_chain RS subsetD, - chain_subset_Pow RS subsetD, - choice_super]) 1); + chain_subset_Pow RS subsetD, + choice_super]) 1); (*Now, verify that it increases*) by (asm_simp_tac (ZF_ss addsimps [Pow_iff, subset_refl] setloop split_tac [expand_if]) 1); @@ -215,16 +215,16 @@ (*Lemma 4*) goal Zorn.thy - "!!S. [| c: TFin(S,next); \ -\ ch: (PROD X: Pow(chain(S))-{0}. X); \ -\ next: increasing(S); \ -\ ALL X: Pow(S). next`X = \ -\ if(X: chain(S)-maxchain(S), ch`super(S,X), X) \ + "!!S. [| c: TFin(S,next); \ +\ ch: (PROD X: Pow(chain(S))-{0}. X); \ +\ next: increasing(S); \ +\ ALL X: Pow(S). next`X = \ +\ if(X: chain(S)-maxchain(S), ch`super(S,X), X) \ \ |] ==> c: chain(S)"; by (etac TFin_induct 1); by (asm_simp_tac (ZF_ss addsimps [chain_subset_Pow RS subsetD, - choice_super RS (super_subset_chain RS subsetD)] + choice_super RS (super_subset_chain RS subsetD)] setloop split_tac [expand_if]) 1); by (rewtac chain_def); by (rtac CollectI 1 THEN fast_tac ZF_cs 1); @@ -252,7 +252,7 @@ by (rtac refl 2); by (asm_full_simp_tac (ZF_ss addsimps [subset_refl RS TFin_UnionI RS - (TFin.dom_subset RS subsetD)] + (TFin.dom_subset RS subsetD)] setloop split_tac [expand_if]) 1); by (eresolve_tac [choice_not_equals RS notE] 1); by (REPEAT (assume_tac 1)); @@ -291,11 +291,11 @@ (*Lemma 5*) val major::prems = goal Zorn.thy - "[| n: TFin(S,next); Z <= TFin(S,next); z:Z; ~ Inter(Z) : Z \ + "[| n: TFin(S,next); Z <= TFin(S,next); z:Z; ~ Inter(Z) : Z \ \ |] ==> ALL m:Z. n<=m"; by (cut_facts_tac prems 1); by (rtac (major RS TFin_induct) 1); -by (fast_tac ZF_cs 2); (*second induction step is easy*) +by (fast_tac ZF_cs 2); (*second induction step is easy*) by (rtac ballI 1); by (rtac (bspec RS TFin_subsetD RS disjE) 1); by (REPEAT_SOME (eresolve_tac [asm_rl,subsetD])); @@ -340,7 +340,7 @@ (** Defining the "next" operation for Zermelo's Theorem **) goal AC.thy - "!!S. [| ch : (PROD X:Pow(S) - {0}. X); X<=S; X~=S \ + "!!S. [| ch : (PROD X:Pow(S) - {0}. X); X<=S; X~=S \ \ |] ==> ch ` (S-X) : S-X"; by (etac apply_type 1); by (fast_tac (eq_cs addEs [equalityE]) 1); @@ -348,8 +348,8 @@ (*This justifies Definition 6.1*) goal Zorn.thy - "!!S. ch: (PROD X: Pow(S)-{0}. X) ==> \ -\ EX next: increasing(S). ALL X: Pow(S). \ + "!!S. ch: (PROD X: Pow(S)-{0}. X) ==> \ +\ EX next: increasing(S). ALL X: Pow(S). \ \ next`X = if(X=S, S, cons(ch`(S-X), X))"; by (rtac bexI 1); by (rtac ballI 1); @@ -372,10 +372,10 @@ (*The construction of the injection*) goal Zorn.thy - "!!S. [| ch: (PROD X: Pow(S)-{0}. X); \ -\ next: increasing(S); \ -\ ALL X: Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X)) \ -\ |] ==> (lam x:S. Union({y: TFin(S,next). x~: y})) \ + "!!S. [| ch: (PROD X: Pow(S)-{0}. X); \ +\ next: increasing(S); \ +\ ALL X: Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X)) \ +\ |] ==> (lam x:S. Union({y: TFin(S,next). x~: y})) \ \ : inj(S, TFin(S,next) - {S})"; by (res_inst_tac [("d", "%y. ch`(S-y)")] lam_injective 1); by (rtac DiffI 1); @@ -392,15 +392,15 @@ \ Union({y: TFin(S,next). x~: y})" 1); by (asm_simp_tac (ZF_ss addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset, - Pow_iff, cons_subset_iff, subset_refl, - choice_Diff RS DiffD2] + Pow_iff, cons_subset_iff, subset_refl, + choice_Diff RS DiffD2] setloop split_tac [expand_if]) 2); by (subgoal_tac "x : next ` Union({y: TFin(S,next). x~: y})" 1); by (fast_tac (subset_cs addSIs [Collect_subset RS TFin_UnionI, TFin_nextI]) 2); (*End of the lemmas!*) by (asm_full_simp_tac (ZF_ss addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset, - Pow_iff, cons_subset_iff, subset_refl] + Pow_iff, cons_subset_iff, subset_refl] setloop split_tac [expand_if]) 1); by (REPEAT (eresolve_tac [asm_rl, consE, sym, notE] 1)); qed "choice_imp_injection"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/add_ind_def.ML --- a/src/ZF/add_ind_def.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/add_ind_def.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/add_ind_def.ML +(* Title: ZF/add_ind_def.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Fixedpoint definition module -- for Inductive/Coinductive Definitions @@ -25,42 +25,42 @@ Products are used only to derive "streamlined" induction rules for relations *) -signature FP = (** Description of a fixed point operator **) +signature FP = (** Description of a fixed point operator **) sig - val oper : term (*fixed point operator*) - val bnd_mono : term (*monotonicity predicate*) - val bnd_monoI : thm (*intro rule for bnd_mono*) - val subs : thm (*subset theorem for fp*) - val Tarski : thm (*Tarski's fixed point theorem*) - val induct : thm (*induction/coinduction rule*) + val oper : term (*fixed point operator*) + val bnd_mono : term (*monotonicity predicate*) + val bnd_monoI : thm (*intro rule for bnd_mono*) + val subs : thm (*subset theorem for fp*) + val Tarski : thm (*Tarski's fixed point theorem*) + val induct : thm (*induction/coinduction rule*) end; -signature PR = (** Description of a Cartesian product **) +signature PR = (** Description of a Cartesian product **) sig - val sigma : term (*Cartesian product operator*) - val pair : term (*pairing operator*) - val split_const : term (*splitting operator*) - val fsplit_const : term (*splitting operator for formulae*) - val pair_iff : thm (*injectivity of pairing, using <->*) - val split_eq : thm (*equality rule for split*) - val fsplitI : thm (*intro rule for fsplit*) - val fsplitD : thm (*destruct rule for fsplit*) - val fsplitE : thm (*elim rule; apparently never used*) + val sigma : term (*Cartesian product operator*) + val pair : term (*pairing operator*) + val split_const : term (*splitting operator*) + val fsplit_const : term (*splitting operator for formulae*) + val pair_iff : thm (*injectivity of pairing, using <->*) + val split_eq : thm (*equality rule for split*) + val fsplitI : thm (*intro rule for fsplit*) + val fsplitD : thm (*destruct rule for fsplit*) + val fsplitE : thm (*elim rule; apparently never used*) end; -signature SU = (** Description of a disjoint sum **) +signature SU = (** Description of a disjoint sum **) sig - val sum : term (*disjoint sum operator*) - val inl : term (*left injection*) - val inr : term (*right injection*) - val elim : term (*case operator*) - val case_inl : thm (*inl equality rule for case*) - val case_inr : thm (*inr equality rule for case*) - val inl_iff : thm (*injectivity of inl, using <->*) - val inr_iff : thm (*injectivity of inr, using <->*) - val distinct : thm (*distinctness of inl, inr using <->*) - val distinct' : thm (*distinctness of inr, inl using <->*) - val free_SEs : thm list (*elim rules for SU, and pair_iff!*) + val sum : term (*disjoint sum operator*) + val inl : term (*left injection*) + val inr : term (*right injection*) + val elim : term (*case operator*) + val case_inl : thm (*inl equality rule for case*) + val case_inr : thm (*inr equality rule for case*) + val inl_iff : thm (*injectivity of inl, using <->*) + val inr_iff : thm (*injectivity of inr, using <->*) + val distinct : thm (*distinctness of inl, inr using <->*) + val distinct' : thm (*distinctness of inr, inl using <->*) + val free_SEs : thm list (*elim rules for SU, and pair_iff!*) end; signature ADD_INDUCTIVE_DEF = @@ -89,8 +89,8 @@ val rec_hds = map head_of rec_tms; val _ = assert_all is_Const rec_hds - (fn t => "Recursive set not previously declared as constant: " ^ - Sign.string_of_term sign t); + (fn t => "Recursive set not previously declared as constant: " ^ + Sign.string_of_term sign t); (*Now we know they are all Consts, so get their names, type and params*) val rec_names = map (#1 o dest_Const) rec_hds @@ -102,16 +102,16 @@ local (*Checking the introduction rules*) val intr_sets = map (#2 o rule_concl_msg sign) intr_tms; fun intr_ok set = - case head_of set of Const(a,recT) => a mem rec_names | _ => false; + case head_of set of Const(a,recT) => a mem rec_names | _ => false; in val _ = assert_all intr_ok intr_sets - (fn t => "Conclusion of rule does not name a recursive set: " ^ - Sign.string_of_term sign t); + (fn t => "Conclusion of rule does not name a recursive set: " ^ + Sign.string_of_term sign t); end; val _ = assert_all is_Free rec_params - (fn t => "Param in recursion term not a free variable: " ^ - Sign.string_of_term sign t); + (fn t => "Param in recursion term not a free variable: " ^ + Sign.string_of_term sign t); (*** Construct the lfp definition ***) val mk_variant = variant (foldr add_term_names (intr_tms,[])); @@ -120,36 +120,36 @@ fun dest_tprop (Const("Trueprop",_) $ P) = P | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^ - Sign.string_of_term sign Q); + Sign.string_of_term sign Q); (*Makes a disjunct from an introduction rule*) fun lfp_part intr = (*quantify over rule's free vars except parameters*) let val prems = map dest_tprop (strip_imp_prems intr) - val _ = seq (fn rec_hd => seq (chk_prem rec_hd) prems) rec_hds - val exfrees = term_frees intr \\ rec_params - val zeq = eq_const $ (Free(z',iT)) $ (#1 (rule_concl intr)) + val _ = seq (fn rec_hd => seq (chk_prem rec_hd) prems) rec_hds + val exfrees = term_frees intr \\ rec_params + val zeq = eq_const $ (Free(z',iT)) $ (#1 (rule_concl intr)) in foldr mk_exists (exfrees, fold_bal (app conj) (zeq::prems)) end; (*The Part(A,h) terms -- compose injections to make h*) - fun mk_Part (Bound 0) = Free(X',iT) (*no mutual rec, no Part needed*) + fun mk_Part (Bound 0) = Free(X',iT) (*no mutual rec, no Part needed*) | mk_Part h = Part_const $ Free(X',iT) $ Abs(w',iT,h); (*Access to balanced disjoint sums via injections*) val parts = - map mk_Part (accesses_bal (ap Su.inl, ap Su.inr, Bound 0) - (length rec_tms)); + map mk_Part (accesses_bal (ap Su.inl, ap Su.inr, Bound 0) + (length rec_tms)); (*replace each set by the corresponding Part(A,h)*) val part_intrs = map (subst_free (rec_tms ~~ parts) o lfp_part) intr_tms; val lfp_abs = absfree(X', iT, - mk_Collect(z', dom_sum, fold_bal (app disj) part_intrs)); + mk_Collect(z', dom_sum, fold_bal (app disj) part_intrs)); val lfp_rhs = Fp.oper $ dom_sum $ lfp_abs val _ = seq (fn rec_hd => deny (rec_hd occs lfp_rhs) - "Illegal occurrence of recursion operator") - rec_hds; + "Illegal occurrence of recursion operator") + rec_hds; (*** Make the new theory ***) @@ -163,11 +163,11 @@ (*The individual sets must already be declared*) val axpairs = map mk_defpair - ((big_rec_tm, lfp_rhs) :: - (case parts of - [_] => [] (*no mutual recursion*) - | _ => rec_tms ~~ (*define the sets as Parts*) - map (subst_atomic [(Free(X',iT),big_rec_tm)]) parts)); + ((big_rec_tm, lfp_rhs) :: + (case parts of + [_] => [] (*no mutual recursion*) + | _ => rec_tms ~~ (*define the sets as Parts*) + map (subst_atomic [(Free(X',iT),big_rec_tm)]) parts)); val _ = seq (writeln o Sign.string_of_term sign o #2) axpairs @@ -179,7 +179,7 @@ fun add_constructs_def (rec_names, con_ty_lists) thy = let val _ = writeln" Defining the constructor functions..."; - val case_name = "f"; (*name for case variables*) + val case_name = "f"; (*name for case variables*) (** Define the constructors **) @@ -189,16 +189,16 @@ fun mk_inject n k u = access_bal(ap Su.inl, ap Su.inr, u) n k; - val npart = length rec_names; (*total # of mutually recursive parts*) + val npart = length rec_names; (*total # of mutually recursive parts*) (*Make constructor definition; kpart is # of this mutually recursive part*) fun mk_con_defs (kpart, con_ty_list) = - let val ncon = length con_ty_list (*number of constructors*) - fun mk_def (((id,T,syn), name, args, prems), kcon) = - (*kcon is index of constructor*) - mk_defpair (list_comb (Const(name,T), args), - mk_inject npart kpart - (mk_inject ncon kcon (mk_tuple args))) + let val ncon = length con_ty_list (*number of constructors*) + fun mk_def (((id,T,syn), name, args, prems), kcon) = + (*kcon is index of constructor*) + mk_defpair (list_comb (Const(name,T), args), + mk_inject npart kpart + (mk_inject ncon kcon (mk_tuple args))) in map mk_def (con_ty_list ~~ (1 upto ncon)) end; (** Define the case operator **) @@ -206,25 +206,25 @@ (*Combine split terms using case; yields the case operator for one part*) fun call_case case_list = let fun call_f (free,args) = - ap_split Pr.split_const free (map (#2 o dest_Free) args) + ap_split Pr.split_const free (map (#2 o dest_Free) args) in fold_bal (app Su.elim) (map call_f case_list) end; (** Generating function variables for the case definition - Non-identifiers (e.g. infixes) get a name of the form f_op_nnn. **) + Non-identifiers (e.g. infixes) get a name of the form f_op_nnn. **) (*Treatment of a single constructor*) fun add_case (((id,T,syn), name, args, prems), (opno,cases)) = - if Syntax.is_identifier id - then (opno, - (Free(case_name ^ "_" ^ id, T), args) :: cases) - else (opno+1, - (Free(case_name ^ "_op_" ^ string_of_int opno, T), args) :: - cases) + if Syntax.is_identifier id + then (opno, + (Free(case_name ^ "_" ^ id, T), args) :: cases) + else (opno+1, + (Free(case_name ^ "_op_" ^ string_of_int opno, T), args) :: + cases) (*Treatment of a list of constructors, for one part*) fun add_case_list (con_ty_list, (opno,case_lists)) = - let val (opno',case_list) = foldr add_case (con_ty_list, (opno,[])) - in (opno', case_list :: case_lists) end; + let val (opno',case_list) = foldr add_case (con_ty_list, (opno,[])) + in (opno', case_list :: case_lists) end; (*Treatment of all parts*) val (_, case_lists) = foldr add_case_list (con_ty_lists, (1,[])); @@ -239,18 +239,18 @@ val big_case_args = flat (map (map #1) case_lists); val big_case_tm = - list_comb (Const(big_case_name, big_case_typ), big_case_args); + list_comb (Const(big_case_name, big_case_typ), big_case_args); val big_case_def = mk_defpair - (big_case_tm, fold_bal (app Su.elim) (map call_case case_lists)); + (big_case_tm, fold_bal (app Su.elim) (map call_case case_lists)); (** Build the new theory **) val const_decs = - (big_case_name, big_case_typ, NoSyn) :: map #1 (flat con_ty_lists); + (big_case_name, big_case_typ, NoSyn) :: map #1 (flat con_ty_lists); val axpairs = - big_case_def :: flat (map mk_con_defs ((1 upto npart) ~~ con_ty_lists)) + big_case_def :: flat (map mk_con_defs ((1 upto npart) ~~ con_ty_lists)) in thy |> add_consts_i const_decs |> add_defs_i axpairs end; end; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/constructor.ML --- a/src/ZF/constructor.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/constructor.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/constructor.ML +(* Title: ZF/constructor.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Constructor function module -- for (Co)Datatype Definitions @@ -8,20 +8,20 @@ signature CONSTRUCTOR_ARG = sig - val thy : theory (*parent containing constructor defs*) - val big_rec_name : string (*name of mutually recursive set*) + val thy : theory (*parent containing constructor defs*) + val big_rec_name : string (*name of mutually recursive set*) val con_ty_lists : ((string*typ*mixfix) * - string * term list * term list) list list - (*description of constructors*) + string * term list * term list) list list + (*description of constructors*) end; signature CONSTRUCTOR_RESULT = sig - val con_defs : thm list (*definitions made in thy*) - val case_eqns : thm list (*equations for case operator*) - val free_iffs : thm list (*freeness rewrite rules*) - val free_SEs : thm list (*freeness destruct rules*) - val mk_free : string -> thm (*makes freeness theorems*) + val con_defs : thm list (*definitions made in thy*) + val case_eqns : thm list (*equations for case operator*) + val free_iffs : thm list (*freeness rewrite rules*) + val free_SEs : thm list (*freeness destruct rules*) + val mk_free : string -> thm (*makes freeness theorems*) end; @@ -82,11 +82,11 @@ con1(x)=con1(y) ==> x=y, con1(x)=con1(y) <-> x=y, etc. *) fun mk_free s = prove_goalw Const.thy con_defs s - (fn prems => [cut_facts_tac prems 1, - fast_tac (ZF_cs addSEs free_SEs) 1]); + (fn prems => [cut_facts_tac prems 1, + fast_tac (ZF_cs addSEs free_SEs) 1]); val case_eqns = map prove_case_equation - (flat Const.con_ty_lists ~~ big_case_args ~~ tl con_defs); + (flat Const.con_ty_lists ~~ big_case_args ~~ tl con_defs); end end; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/domrange.ML --- a/src/ZF/domrange.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/domrange.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/domrange +(* Title: ZF/domrange ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1991 University of Cambridge Converse, domain, range of a relation or function @@ -31,7 +31,7 @@ (assume_tac 1) ]); val converse_cs = pair_cs addSIs [converseI] - addSEs [converseD,converseE]; + addSEs [converseD,converseE]; qed_goal "converse_converse" ZF.thy "!!A B r. r<=Sigma(A,B) ==> converse(converse(r)) = r" @@ -137,7 +137,7 @@ qed_goal "image_singleton_iff" ZF.thy "b : r``{a} <-> :r" (fn _ => [ rtac (image_iff RS iff_trans) 1, - fast_tac pair_cs 1 ]); + fast_tac pair_cs 1 ]); qed_goalw "imageI" ZF.thy [image_def] "!!a b r. [| : r; a:A |] ==> b : r``A" @@ -163,7 +163,7 @@ qed_goal "vimage_singleton_iff" ZF.thy "a : r-``{b} <-> :r" (fn _ => [ rtac (vimage_iff RS iff_trans) 1, - fast_tac pair_cs 1 ]); + fast_tac pair_cs 1 ]); qed_goalw "vimageI" ZF.thy [vimage_def] "!!A B r. [| : r; b:B |] ==> a : r-``B" diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/equalities.ML --- a/src/ZF/equalities.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/equalities.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/equalities +(* Title: ZF/equalities ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge Set Theory examples: Union, Intersection, Inclusion, etc. @@ -292,7 +292,7 @@ (*First-order version of the above, for rewriting*) goal ZF.thy "I * (A Un B) = I*A Un I*B"; -by (resolve_tac [SUM_Un_distrib2] 1); +by (rtac SUM_Un_distrib2 1); qed "prod_Un_distrib2"; goal ZF.thy "(SUM i:I Int J. C(i)) = (SUM i:I. C(i)) Int (SUM j:J. C(j))"; @@ -306,7 +306,7 @@ (*First-order version of the above, for rewriting*) goal ZF.thy "I * (A Int B) = I*A Int I*B"; -by (resolve_tac [SUM_Int_distrib2] 1); +by (rtac SUM_Int_distrib2 1); qed "prod_Int_distrib2"; (*Cf Aczel, Non-Well-Founded Sets, page 115*) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Acc.ML --- a/src/ZF/ex/Acc.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Acc.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/acc +(* Title: ZF/ex/acc ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Inductive definition of acc(r) @@ -26,8 +26,8 @@ qed "acc_downward"; val [major,indhyp] = goal Acc.thy - "[| a : acc(r); \ -\ !!x. [| x: acc(r); ALL y. :r --> P(y) |] ==> P(x) \ + "[| a : acc(r); \ +\ !!x. [| x: acc(r); ALL y. :r --> P(y) |] ==> P(x) \ \ |] ==> P(a)"; by (rtac (major RS acc.induct) 1); by (rtac indhyp 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/BT.ML --- a/src/ZF/ex/BT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/BT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/BT.ML +(* Title: ZF/ex/BT.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Datatype definition of binary trees @@ -11,8 +11,8 @@ (*Perform induction on l, then prove the major premise using prems. *) fun bt_ind_tac a prems i = EVERY [res_inst_tac [("x",a)] bt.induct i, - rename_last_tac a ["1","2"] (i+2), - ares_tac prems i]; + rename_last_tac a ["1","2"] (i+2), + ares_tac prems i]; (** Lemmas to justify using "bt" in other recursive type definitions **) @@ -27,7 +27,7 @@ by (rtac lfp_lowerbound 1); by (rtac (A_subset_univ RS univ_mono) 2); by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, - Pair_in_univ]) 1); + Pair_in_univ]) 1); qed "bt_univ"; bind_thm ("bt_subset_univ", ([bt_mono, bt_univ] MRS subset_trans)); @@ -59,11 +59,11 @@ "[| t: bt(A); \ \ c: C(Lf); \ \ !!x y z r s. [| x:A; y:bt(A); z:bt(A); r:C(y); s:C(z) |] ==> \ -\ h(x,y,z,r,s): C(Br(x,y,z)) \ +\ h(x,y,z,r,s): C(Br(x,y,z)) \ \ |] ==> bt_rec(t,c,h) : C(t)"; by (bt_ind_tac "t" prems 1); by (ALLGOALS (asm_simp_tac (ZF_ss addsimps - (prems@[bt_rec_Lf,bt_rec_Br])))); + (prems@[bt_rec_Lf,bt_rec_Br])))); qed "bt_rec_type"; (** Versions for use with definitions **) @@ -119,9 +119,9 @@ addsimps bt.case_eqns addsimps bt_typechecks addsimps [bt_rec_Lf, bt_rec_Br, - n_nodes_Lf, n_nodes_Br, - n_leaves_Lf, n_leaves_Br, - bt_reflect_Lf, bt_reflect_Br]; + n_nodes_Lf, n_nodes_Br, + n_leaves_Lf, n_leaves_Br, + bt_reflect_Lf, bt_reflect_Br]; (*** theorems about n_leaves ***) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Bin.ML --- a/src/ZF/ex/Bin.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Bin.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/Bin.ML +(* Title: ZF/ex/Bin.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge For Bin.thy. Arithmetic on binary integers. @@ -11,8 +11,8 @@ (*Perform induction on l, then prove the major premise using prems. *) fun bin_ind_tac a prems i = EVERY [res_inst_tac [("x",a)] bin.induct i, - rename_last_tac a ["1"] (i+3), - ares_tac prems i]; + rename_last_tac a ["1"] (i+3), + ares_tac prems i]; (** bin_rec -- by Vset recursion **) @@ -44,7 +44,7 @@ by (bin_ind_tac "w" prems 1); by (ALLGOALS (asm_simp_tac (ZF_ss addsimps (prems@[bin_rec_Plus, bin_rec_Minus, - bin_rec_Bcons])))); + bin_rec_Bcons])))); qed "bin_rec_type"; (** Versions for use with definitions **) @@ -68,7 +68,7 @@ qed "def_bin_rec_Bcons"; fun bin_recs def = map standard - ([def] RL [def_bin_rec_Plus, def_bin_rec_Minus, def_bin_rec_Bcons]); + ([def] RL [def_bin_rec_Plus, def_bin_rec_Minus, def_bin_rec_Bcons]); goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Plus,0) = Plus"; by (asm_simp_tac (ZF_ss addsimps (bin.case_eqns @ bool_simps)) 1); @@ -92,8 +92,8 @@ qed "norm_Bcons_Bcons"; val norm_Bcons_simps = [norm_Bcons_Plus_0, norm_Bcons_Plus_1, - norm_Bcons_Minus_0, norm_Bcons_Minus_1, - norm_Bcons_Bcons]; + norm_Bcons_Minus_0, norm_Bcons_Minus_1, + norm_Bcons_Bcons]; (** Type checking **) @@ -102,7 +102,7 @@ goalw Bin.thy [integ_of_bin_def] "!!w. w: bin ==> integ_of_bin(w) : integ"; by (typechk_tac (bin_typechecks0@integ_typechecks@ - nat_typechecks@[bool_into_nat])); + nat_typechecks@[bool_into_nat])); qed "integ_of_bin_type"; goalw Bin.thy [norm_Bcons_def] @@ -130,13 +130,13 @@ goalw Bin.thy [bin_add_def] "!!v w. [| v: bin; w: bin |] ==> bin_add(v,w) : bin"; by (typechk_tac ([norm_Bcons_type, bin_succ_type, bin_pred_type]@ - bin_typechecks0@ bool_typechecks@ZF_typechecks)); + bin_typechecks0@ bool_typechecks@ZF_typechecks)); qed "bin_add_type"; goalw Bin.thy [bin_mult_def] "!!v w. [| v: bin; w: bin |] ==> bin_mult(v,w) : bin"; by (typechk_tac ([norm_Bcons_type, bin_minus_type, bin_add_type]@ - bin_typechecks0@ bool_typechecks)); + bin_typechecks0@ bool_typechecks)); qed "bin_mult_type"; val bin_typechecks = bin_typechecks0 @ @@ -145,8 +145,8 @@ val bin_ss = integ_ss addsimps([bool_1I, bool_0I, - bin_rec_Plus, bin_rec_Minus, bin_rec_Bcons] @ - bin_recs integ_of_bin_def @ bool_simps @ bin_typechecks); + bin_rec_Plus, bin_rec_Minus, bin_rec_Bcons] @ + bin_recs integ_of_bin_def @ bool_simps @ bin_typechecks); val typechecks = bin_typechecks @ integ_typechecks @ nat_typechecks @ [bool_subset_nat RS subsetD]; @@ -178,7 +178,7 @@ (*norm_Bcons preserves the integer value of its argument*) goal Bin.thy - "!!w. [| w: bin; b: bool |] ==> \ + "!!w. [| w: bin; b: bool |] ==> \ \ integ_of_bin(norm_Bcons(w,b)) = integ_of_bin(Bcons(w,b))"; by (etac bin.elim 1); by (asm_simp_tac (ZF_ss addsimps norm_Bcons_simps) 3); @@ -213,7 +213,7 @@ val bin_minus_ss = bin_ss addsimps (bin_recs bin_minus_def @ - [integ_of_bin_succ, integ_of_bin_pred]); + [integ_of_bin_succ, integ_of_bin_pred]); goal Bin.thy "!!w. w: bin ==> integ_of_bin(bin_minus(w)) = $~ integ_of_bin(w)"; @@ -252,9 +252,9 @@ qed "bin_add_Bcons_Bcons"; val bin_add_simps = [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, - bin_add_Bcons_Minus, bin_add_Bcons_Bcons, - integ_of_bin_succ, integ_of_bin_pred, - integ_of_bin_norm_Bcons]; + bin_add_Bcons_Minus, bin_add_Bcons_Bcons, + integ_of_bin_succ, integ_of_bin_pred, + integ_of_bin_norm_Bcons]; val bin_add_ss = bin_ss addsimps ([bool_subset_nat RS subsetD] @ bin_add_simps); @@ -283,11 +283,11 @@ val bin_mult_ss = bin_ss addsimps (bin_recs bin_mult_def @ - [integ_of_bin_minus, integ_of_bin_add, - integ_of_bin_norm_Bcons]); + [integ_of_bin_minus, integ_of_bin_add, + integ_of_bin_norm_Bcons]); val major::prems = goal Bin.thy - "[| v: bin; w: bin |] ==> \ + "[| v: bin; w: bin |] ==> \ \ integ_of_bin(bin_mult(v,w)) = \ \ integ_of_bin(v) $* integ_of_bin(w)"; by (cut_facts_tac prems 1); @@ -373,19 +373,19 @@ val bin_comp_ss = integ_ss addsimps [integ_of_bin_add RS sym, (*invoke bin_add*) - integ_of_bin_minus RS sym, (*invoke bin_minus*) - integ_of_bin_mult RS sym, (*invoke bin_mult*) - bin_succ_Plus, bin_succ_Minus, - bin_succ_Bcons1, bin_succ_Bcons0, - bin_pred_Plus, bin_pred_Minus, - bin_pred_Bcons1, bin_pred_Bcons0, - bin_minus_Plus, bin_minus_Minus, - bin_minus_Bcons1, bin_minus_Bcons0, - bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, - bin_add_Bcons_Minus, bin_add_Bcons_Bcons0, - bin_add_Bcons_Bcons10, bin_add_Bcons_Bcons11, - bin_mult_Plus, bin_mult_Minus, - bin_mult_Bcons1, bin_mult_Bcons0] @ + integ_of_bin_minus RS sym, (*invoke bin_minus*) + integ_of_bin_mult RS sym, (*invoke bin_mult*) + bin_succ_Plus, bin_succ_Minus, + bin_succ_Bcons1, bin_succ_Bcons0, + bin_pred_Plus, bin_pred_Minus, + bin_pred_Bcons1, bin_pred_Bcons0, + bin_minus_Plus, bin_minus_Minus, + bin_minus_Bcons1, bin_minus_Bcons0, + bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, + bin_add_Bcons_Minus, bin_add_Bcons_Bcons0, + bin_add_Bcons_Bcons10, bin_add_Bcons_Bcons11, + bin_mult_Plus, bin_mult_Minus, + bin_mult_Bcons1, bin_mult_Bcons0] @ norm_Bcons_simps setsolver (type_auto_tac ([bool_1I, bool_0I] @ bin_typechecks0)); @@ -395,63 +395,63 @@ (*All runtimes below are on a SPARCserver 10*) goal Bin.thy "#13 $+ #19 = #32"; -by (simp_tac bin_comp_ss 1); (*0.4 secs*) +by (simp_tac bin_comp_ss 1); (*0.4 secs*) result(); bin_add(binary_of_int 13, binary_of_int 19); goal Bin.thy "#1234 $+ #5678 = #6912"; -by (simp_tac bin_comp_ss 1); (*1.3 secs*) +by (simp_tac bin_comp_ss 1); (*1.3 secs*) result(); bin_add(binary_of_int 1234, binary_of_int 5678); goal Bin.thy "#1359 $+ #~2468 = #~1109"; -by (simp_tac bin_comp_ss 1); (*1.2 secs*) +by (simp_tac bin_comp_ss 1); (*1.2 secs*) result(); bin_add(binary_of_int 1359, binary_of_int ~2468); goal Bin.thy "#93746 $+ #~46375 = #47371"; -by (simp_tac bin_comp_ss 1); (*1.9 secs*) +by (simp_tac bin_comp_ss 1); (*1.9 secs*) result(); bin_add(binary_of_int 93746, binary_of_int ~46375); goal Bin.thy "$~ #65745 = #~65745"; -by (simp_tac bin_comp_ss 1); (*0.4 secs*) +by (simp_tac bin_comp_ss 1); (*0.4 secs*) result(); bin_minus(binary_of_int 65745); (* negation of ~54321 *) goal Bin.thy "$~ #~54321 = #54321"; -by (simp_tac bin_comp_ss 1); (*0.5 secs*) +by (simp_tac bin_comp_ss 1); (*0.5 secs*) result(); bin_minus(binary_of_int ~54321); goal Bin.thy "#13 $* #19 = #247"; -by (simp_tac bin_comp_ss 1); (*0.7 secs*) +by (simp_tac bin_comp_ss 1); (*0.7 secs*) result(); bin_mult(binary_of_int 13, binary_of_int 19); goal Bin.thy "#~84 $* #51 = #~4284"; -by (simp_tac bin_comp_ss 1); (*1.3 secs*) +by (simp_tac bin_comp_ss 1); (*1.3 secs*) result(); bin_mult(binary_of_int ~84, binary_of_int 51); (*The worst case for 8-bit operands *) goal Bin.thy "#255 $* #255 = #65025"; -by (simp_tac bin_comp_ss 1); (*4.3 secs*) +by (simp_tac bin_comp_ss 1); (*4.3 secs*) result(); bin_mult(binary_of_int 255, binary_of_int 255); goal Bin.thy "#1359 $* #~2468 = #~3354012"; -by (simp_tac bin_comp_ss 1); (*6.1 secs*) +by (simp_tac bin_comp_ss 1); (*6.1 secs*) result(); bin_mult(binary_of_int 1359, binary_of_int ~2468); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Brouwer.ML --- a/src/ZF/ex/Brouwer.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Brouwer.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/Brouwer.ML +(* Title: ZF/ex/Brouwer.ML ID: $ $ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Infinite branching datatype definitions @@ -21,11 +21,11 @@ (*A nicer induction rule than the standard one*) val major::prems = goal Brouwer.thy - "[| b: brouwer; \ -\ P(Zero); \ -\ !!b. [| b: brouwer; P(b) |] ==> P(Suc(b)); \ -\ !!h. [| h: nat -> brouwer; ALL i:nat. P(h`i) \ -\ |] ==> P(Lim(h)) \ + "[| b: brouwer; \ +\ P(Zero); \ +\ !!b. [| b: brouwer; P(b) |] ==> P(Suc(b)); \ +\ !!h. [| h: nat -> brouwer; ALL i:nat. P(h`i) \ +\ |] ==> P(Lim(h)) \ \ |] ==> P(b)"; by (rtac (major RS brouwer.induct) 1); by (REPEAT_SOME (ares_tac prems)); @@ -45,9 +45,9 @@ (*A nicer induction rule than the standard one*) val major::prems = goal Brouwer.thy - "[| w: Well(A,B); \ -\ !!a f. [| a: A; f: B(a) -> Well(A,B); ALL y: B(a). P(f`y) \ -\ |] ==> P(Sup(a,f)) \ + "[| w: Well(A,B); \ +\ !!a f. [| a: A; f: B(a) -> Well(A,B); ALL y: B(a). P(f`y) \ +\ |] ==> P(Sup(a,f)) \ \ |] ==> P(w)"; by (rtac (major RS Well.induct) 1); by (REPEAT_SOME (ares_tac prems)); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/CoUnit.ML --- a/src/ZF/ex/CoUnit.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/CoUnit.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/CoUnit.ML +(* Title: ZF/ex/CoUnit.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Trivial codatatype definitions, one of which goes wrong! @@ -59,13 +59,13 @@ by (etac counit2.elim 1); by (rewrite_goals_tac counit2.con_defs); val lleq_cs = subset_cs - addSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono] + addSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono] addSEs [Ord_in_Ord, Pair_inject]; by (fast_tac lleq_cs 1); qed "counit2_Int_Vset_subset_lemma"; val counit2_Int_Vset_subset = standard - (counit2_Int_Vset_subset_lemma RS spec RS spec RS mp); + (counit2_Int_Vset_subset_lemma RS spec RS spec RS mp); goal CoUnit.thy "!!x y. [| x: counit2; y: counit2 |] ==> x=y"; by (rtac equalityI 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Comb.ML --- a/src/ZF/ex/Comb.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Comb.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/comb.ML +(* Title: ZF/ex/comb.ML ID: $Id$ - Author: Lawrence C Paulson + Author: Lawrence C Paulson Copyright 1993 University of Cambridge Combinatory Logic example: the Church-Rosser Theorem @@ -42,10 +42,10 @@ (r_into_rtrancl RS (trans_rtrancl RS transD))); val reduction_rls = [reduction_refl, contract.K, contract.S, - contract.K RS rtrancl_into_rtrancl2, - contract.S RS rtrancl_into_rtrancl2, - contract.Ap1 RS rtrancl_into_rtrancl2, - contract.Ap2 RS rtrancl_into_rtrancl2]; + contract.K RS rtrancl_into_rtrancl2, + contract.S RS rtrancl_into_rtrancl2, + contract.Ap1 RS rtrancl_into_rtrancl2, + contract.Ap2 RS rtrancl_into_rtrancl2]; (*Example only: not used*) goalw Comb.thy [I_def] "!!p. p:comb ==> I#p ---> p"; @@ -65,10 +65,10 @@ val contract_cs = ZF_cs addSIs comb.intrs - addIs contract.intrs - addSEs [contract_combD1,contract_combD2] (*type checking*) - addSEs [K_contractE, S_contractE, Ap_contractE] - addSEs comb.free_SEs; + addIs contract.intrs + addSEs [contract_combD1,contract_combD2] (*type checking*) + addSEs [K_contractE, S_contractE, Ap_contractE] + addSEs comb.free_SEs; goalw Comb.thy [I_def] "!!r. I -1-> r ==> P"; by (fast_tac contract_cs 1); @@ -129,7 +129,7 @@ goal Comb.thy "field(parcontract) = comb"; by (fast_tac (ZF_cs addIs [equalityI, parcontract.K] - addSEs [parcontract_combE2]) 1); + addSEs [parcontract_combE2]) 1); qed "field_parcontract_eq"; (*Derive a case for each combinator constructor*) @@ -139,10 +139,10 @@ val parcontract_cs = ZF_cs addSIs comb.intrs - addIs parcontract.intrs - addSEs [Ap_E, K_parcontractE, S_parcontractE, - Ap_parcontractE] - addSEs [parcontract_combD1, parcontract_combD2] (*type checking*) + addIs parcontract.intrs + addSEs [Ap_E, K_parcontractE, S_parcontractE, + Ap_parcontractE] + addSEs [parcontract_combD1, parcontract_combD2] (*type checking*) addSEs comb.free_SEs; (*** Basic properties of parallel contraction ***) @@ -176,18 +176,18 @@ by (etac trancl_induct 1); by (fast_tac (ZF_cs addIs [r_into_trancl]) 1); by (slow_best_tac (ZF_cs addSDs [spec RS mp] - addIs [r_into_trancl, trans_trancl RS transD]) 1); + addIs [r_into_trancl, trans_trancl RS transD]) 1); qed "diamond_trancl_lemma"; val diamond_lemmaE = diamond_trancl_lemma RS spec RS mp RS exE; val [major] = goal Comb.thy "diamond(r) ==> diamond(r^+)"; -bw diamond_def; (*unfold only in goal, not in premise!*) +by (rewtac diamond_def); (*unfold only in goal, not in premise!*) by (rtac (impI RS allI RS allI) 1); by (etac trancl_induct 1); by (ALLGOALS (slow_best_tac (ZF_cs addIs [r_into_trancl, trans_trancl RS transD] - addEs [major RS diamond_lemmaE]))); + addEs [major RS diamond_lemmaE]))); qed "diamond_trancl"; @@ -204,7 +204,7 @@ by (etac rtrancl_induct 1); by (fast_tac (parcontract_cs addIs [r_into_trancl]) 1); by (fast_tac (ZF_cs addIs [contract_imp_parcontract, - r_into_trancl, trans_trancl RS transD]) 1); + r_into_trancl, trans_trancl RS transD]) 1); qed "reduce_imp_parreduce"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Data.ML --- a/src/ZF/ex/Data.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Data.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/Data.ML +(* Title: ZF/ex/Data.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Sample datatype definition. @@ -28,7 +28,7 @@ by (rtac lfp_lowerbound 1); by (rtac ([A_subset_univ, Un_upper1] MRS subset_trans RS univ_mono) 2); by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ, - Pair_in_univ]) 1); + Pair_in_univ]) 1); qed "data_univ"; bind_thm ("data_subset_univ", ([data_mono, data_univ] MRS subset_trans)); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Enum.ML --- a/src/ZF/ex/Enum.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Enum.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/Enum +(* Title: ZF/ex/Enum ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Example of a BIG enumeration type diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Integ.ML --- a/src/ZF/ex/Integ.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Integ.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/integ.ML +(* Title: ZF/ex/integ.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge For integ.thy. The integers as equivalence classes over nat*nat. @@ -70,7 +70,7 @@ val intrel_ss = arith_ss addsimps [equiv_intrel RS eq_equiv_class_iff, intrel_iff, - add_0_right, add_succ_right] + add_0_right, add_succ_right] addcongs [conj_cong]; val eq_intrelD = equiv_intrel RSN (2,eq_equiv_class); @@ -104,7 +104,7 @@ goalw Integ.thy [integ_def,zminus_def] "!!z. z : integ ==> $~z : integ"; by (typechk_tac [split_type, SigmaI, zminus_ize UN_equiv_class_type, - quotientI]); + quotientI]); qed "zminus_type"; goalw Integ.thy [integ_def,zminus_def] @@ -113,7 +113,7 @@ by (safe_tac intrel_cs); (*The setloop is only needed because assumptions are in the wrong order!*) by (asm_full_simp_tac (intrel_ss addsimps add_ac - setloop dtac eq_intrelD) 1); + setloop dtac eq_intrelD) 1); qed "zminus_inject"; goalw Integ.thy [zminus_def] @@ -139,7 +139,7 @@ by (dres_inst_tac [("psi", "?lhs=?rhs")] asm_rl 1); by (dres_inst_tac [("psi", "?lhs zmagnitude(z) : nat"; by (typechk_tac [split_type, zmagnitude_ize UN_equiv_class_type, - add_type, diff_type]); + add_type, diff_type]); qed "zmagnitude_type"; goalw Integ.thy [zmagnitude_def] @@ -203,7 +203,7 @@ (** Congruence property for addition **) goalw Integ.thy [congruent2_def] - "congruent2(intrel, %z1 z2. \ + "congruent2(intrel, %z1 z2. \ \ let =z1; =z2 \ \ in intrel``{})"; (*Proof via congruent2_commuteI seems longer*) @@ -228,9 +228,9 @@ qed "zadd_type"; goalw Integ.thy [zadd_def] - "!!x1 y1. [| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ -\ (intrel``{}) $+ (intrel``{}) = \ -\ intrel `` {}"; + "!!x1 y1. [| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ +\ (intrel``{}) $+ (intrel``{}) = \ +\ intrel `` {}"; by (asm_simp_tac (ZF_ss addsimps [zadd_ize UN_equiv_class2, SigmaI]) 1); by (simp_tac (ZF_ss addsimps [Let_def]) 1); qed "zadd"; @@ -299,8 +299,8 @@ (** Congruence property for multiplication **) goal Integ.thy - "congruent2(intrel, %p1 p2. \ -\ split(%x1 y1. split(%x2 y2. \ + "congruent2(intrel, %p1 p2. \ +\ split(%x1 y1. split(%x2 y2. \ \ intrel``{}, p2), p1))"; by (rtac (equiv_intrel RS congruent2_commuteI) 1); by (safe_tac intrel_cs); @@ -321,13 +321,13 @@ goalw Integ.thy [integ_def,zmult_def] "!!z w. [| z: integ; w: integ |] ==> z $* w : integ"; by (REPEAT (ares_tac [zmult_ize UN_equiv_class_type2, - split_type, add_type, mult_type, - quotientI, SigmaI] 1)); + split_type, add_type, mult_type, + quotientI, SigmaI] 1)); qed "zmult_type"; goalw Integ.thy [zmult_def] - "!!x1 x2. [| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ -\ (intrel``{}) $* (intrel``{}) = \ + "!!x1 x2. [| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \ +\ (intrel``{}) $* (intrel``{}) = \ \ intrel `` {}"; by (asm_simp_tac (ZF_ss addsimps [zmult_ize UN_equiv_class2, SigmaI]) 1); qed "zmult"; @@ -367,7 +367,7 @@ by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); by (asm_simp_tac (intrel_ss addsimps ([zmult, add_mult_distrib_left, - add_mult_distrib] @ add_ac @ mult_ac)) 1); + add_mult_distrib] @ add_ac @ mult_ac)) 1); qed "zmult_assoc"; (*For AC rewriting*) @@ -375,7 +375,7 @@ "!!z1 z2 z3. [| z1:integ; z2:integ; z3: integ |] ==> \ \ z1$*(z2$*z3) = z2$*(z1$*z3)" (fn _ => [asm_simp_tac (ZF_ss addsimps [zmult_assoc RS sym, - zmult_commute]) 1]); + zmult_commute]) 1]); (*Integer multiplication is an AC operator*) val zmult_ac = [zmult_assoc, zmult_commute, zmult_left_commute]; @@ -386,7 +386,7 @@ by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1)); by (asm_simp_tac (intrel_ss addsimps ([zadd, zmult, add_mult_distrib] @ - add_ac @ mult_ac)) 1); + add_ac @ mult_ac)) 1); qed "zadd_zmult_distrib"; val integ_typechecks = @@ -401,5 +401,5 @@ val integ_ss = arith_ss addsimps (zadd_simps @ zminus_simps @ zmult_simps @ - [zmagnitude_znat, zmagnitude_zminus_znat] @ - integ_typechecks); + [zmagnitude_znat, zmagnitude_zminus_znat] @ + integ_typechecks); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/LList.ML --- a/src/ZF/ex/LList.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/LList.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/LList.ML +(* Title: ZF/ex/LList.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Codatatype definition of Lazy Lists @@ -33,10 +33,10 @@ (** Closure of quniv(A) under llist -- why so complex? Its a gfp... **) val quniv_cs = subset_cs addSIs [QPair_Int_Vset_subset_UN RS subset_trans, - QPair_subset_univ, - empty_subsetI, one_in_quniv RS qunivD] + QPair_subset_univ, + empty_subsetI, one_in_quniv RS qunivD] addIs [Int_lower1 RS subset_trans] - addSDs [qunivD] + addSDs [qunivD] addSEs [Ord_in_Ord]; goal LList.thy @@ -65,7 +65,7 @@ (*** Lazy List Equality: lleq ***) val lleq_cs = subset_cs - addSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono] + addSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono] addSEs [Ord_in_Ord, Pair_inject]; (*Lemma for proving finality. Unfold the lazy list; use induction hypothesis*) @@ -80,7 +80,7 @@ qed "lleq_Int_Vset_subset_lemma"; bind_thm ("lleq_Int_Vset_subset", - (lleq_Int_Vset_subset_lemma RS spec RS spec RS mp)); + (lleq_Int_Vset_subset_lemma RS spec RS spec RS mp)); (*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*) @@ -118,7 +118,7 @@ by (rtac bnd_monoI 1); by (REPEAT (ares_tac [subset_refl, QInr_mono, QPair_mono] 2)); by (REPEAT (ares_tac [subset_refl, A_subset_univ, - QInr_subset_univ, QPair_subset_univ] 1)); + QInr_subset_univ, QPair_subset_univ] 1)); qed "lconst_fun_bnd_mono"; (* lconst(a) = LCons(a,lconst(a)) *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Limit.ML --- a/src/ZF/ex/Limit.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Limit.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/Limit +(* Title: ZF/ex/Limit ID: $Id$ - Author: Sten Agerholm + Author: Sten Agerholm The inverse limit construction. *) @@ -8,7 +8,7 @@ open Limit; (*----------------------------------------------------------------------*) -(* Useful goal commands. *) +(* Useful goal commands. *) (*----------------------------------------------------------------------*) val brr = fn thl => fn n => by(REPEAT(ares_tac thl n)); @@ -16,17 +16,17 @@ fun rotate n i = EVERY(replicate n (etac revcut_rl i)); (*----------------------------------------------------------------------*) -(* Preliminary theorems. *) +(* Preliminary theorems. *) (*----------------------------------------------------------------------*) val theI2 = prove_goal ZF.thy (* From Larry *) "[| EX! x. P(x); !!x. P(x) ==> Q(x) |] ==> Q(THE x.P(x))" (fn prems => [ resolve_tac prems 1, - rtac theI 1, - resolve_tac prems 1 ]); + rtac theI 1, + resolve_tac prems 1 ]); (*----------------------------------------------------------------------*) -(* Basic results. *) +(* Basic results. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [set_def] @@ -45,7 +45,7 @@ val rel_E = result(); (*----------------------------------------------------------------------*) -(* I/E/D rules for po and cpo. *) +(* I/E/D rules for po and cpo. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [po_def] @@ -118,7 +118,7 @@ val cpo_islub = result(); (*----------------------------------------------------------------------*) -(* Theorems about isub and islub. *) +(* Theorems about isub and islub. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [islub_def] (* islub_isub *) @@ -181,7 +181,7 @@ val islub_unique = result(); (*----------------------------------------------------------------------*) -(* lub gives the least upper bound of chains. *) +(* lub gives the least upper bound of chains. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [lub_def] @@ -196,7 +196,7 @@ val cpo_lub = result(); (*----------------------------------------------------------------------*) -(* Theorems about chains. *) +(* Theorems about chains. *) (*----------------------------------------------------------------------*) val chainI = prove_goalw Limit.thy [chain_def] @@ -251,7 +251,7 @@ val chain_rel_gen = result(); (*----------------------------------------------------------------------*) -(* Theorems about pcpos and bottom. *) +(* Theorems about pcpos and bottom. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [pcpo_def] (* pcpoI *) @@ -309,7 +309,7 @@ val bot_unique = result(); (*----------------------------------------------------------------------*) -(* Constant chains and lubs and cpos. *) +(* Constant chains and lubs and cpos. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [chain_def] (* chain_const *) @@ -341,7 +341,7 @@ val lub_const = result(); (*----------------------------------------------------------------------*) -(* Taking the suffix of chains has no effect on ub's. *) +(* Taking the suffix of chains has no effect on ub's. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [isub_def,suffix_def] (* isub_suffix *) @@ -368,7 +368,7 @@ val lub_suffix = result(); (*----------------------------------------------------------------------*) -(* Dominate and subchain. *) +(* Dominate and subchain. *) (*----------------------------------------------------------------------*) val dominateI = prove_goalw Limit.thy [dominate_def] @@ -442,7 +442,7 @@ val dominate_islub_eq = result(); (*----------------------------------------------------------------------*) -(* Matrix. *) +(* Matrix. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [matrix_def] (* matrix_fun *) @@ -660,7 +660,7 @@ val lub_matrix_diag_sym = result(); (*----------------------------------------------------------------------*) -(* I/E/D rules for mono and cont. *) +(* I/E/D rules for mono and cont. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [mono_def] (* monoI *) @@ -724,7 +724,7 @@ val cont_lub = result(); (*----------------------------------------------------------------------*) -(* Continuity and chains. *) +(* Continuity and chains. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [] (* mono_chain *) @@ -753,7 +753,7 @@ val cont_chain = result(); (*----------------------------------------------------------------------*) -(* I/E/D rules about (set+rel) cf, the continuous function space. *) +(* I/E/D rules about (set+rel) cf, the continuous function space. *) (*----------------------------------------------------------------------*) (* The following development more difficult with cpo-as-relation approach. *) @@ -787,7 +787,7 @@ val rel_cf = result(); (*----------------------------------------------------------------------*) -(* Theorems about the continuous function space. *) +(* Theorems about the continuous function space. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [] (* chain_cf *) @@ -960,7 +960,7 @@ val bot_cf = result(); (*----------------------------------------------------------------------*) -(* Identity and composition. *) +(* Identity and composition. *) (*----------------------------------------------------------------------*) val id_thm = prove_goalw Perm.thy [id_def] "x:X ==> (id(X)`x) = x" @@ -1049,7 +1049,7 @@ val comp_lubs = result(); (*----------------------------------------------------------------------*) -(* A couple (out of many) theorems about arithmetic. *) +(* A couple (out of many) theorems about arithmetic. *) (*----------------------------------------------------------------------*) val prems = goal Arith.thy (* le_cases *) @@ -1065,7 +1065,7 @@ val lt_cases = result(); (*----------------------------------------------------------------------*) -(* Theorems about projpair. *) +(* Theorems about projpair. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [projpair_def] (* projpairI *) @@ -1109,12 +1109,12 @@ val projpairDs = [projpair_e_cont,projpair_p_cont,projpair_eq,projpair_rel]; (*----------------------------------------------------------------------*) -(* NB! projpair_e_cont and projpair_p_cont cannot be used repeatedly *) +(* NB! projpair_e_cont and projpair_p_cont cannot be used repeatedly *) (* at the same time since both match a goal of the form f:cont(X,Y).*) (*----------------------------------------------------------------------*) (*----------------------------------------------------------------------*) -(* Uniqueness of embedding projection pairs. *) +(* Uniqueness of embedding projection pairs. *) (*----------------------------------------------------------------------*) val id_comp = fun_is_rel RS left_comp_id; @@ -1242,7 +1242,7 @@ (*----------------------------------------------------------------------*) -(* The identity embedding. *) +(* The identity embedding. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [projpair_def] (* projpair_id *) @@ -1264,7 +1264,7 @@ val Rp_id = result(); (*----------------------------------------------------------------------*) -(* Composition preserves embeddings. *) +(* Composition preserves embeddings. *) (*----------------------------------------------------------------------*) (* Considerably shorter, only partly due to a simpler comp_assoc. *) @@ -1305,7 +1305,7 @@ val Rp_comp = lemma RS Rp_unique; (*----------------------------------------------------------------------*) -(* Infinite cartesian product. *) +(* Infinite cartesian product. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [set_def,iprod_def] (* iprodI *) @@ -1425,7 +1425,7 @@ val lub_iprod = result(); (*----------------------------------------------------------------------*) -(* The notion of subcpo. *) +(* The notion of subcpo. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [subcpo_def] (* subcpoI *) @@ -1472,7 +1472,7 @@ brr(isubD1::prems)1; brr((hd prems RS subcpo_relD1)::chain_in::isubD1::isubD2::prems)1; val ub_subcpo = result(); - + (* STRIP_TAC and HOL resolution is efficient sometimes. The following theorem is proved easily in HOL without intro and elim rules. *) @@ -1508,7 +1508,7 @@ val lub_subcpo = result(); (*----------------------------------------------------------------------*) -(* Making subcpos using mkcpo. *) +(* Making subcpos using mkcpo. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [set_def,mkcpo_def] (* mkcpoI *) @@ -1587,7 +1587,7 @@ val subcpo_mkcpo = result(); (*----------------------------------------------------------------------*) -(* Embedding projection chains of cpos. *) +(* Embedding projection chains of cpos. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [emb_chain_def] (* emb_chainI *) @@ -1606,7 +1606,7 @@ (fn prems => [fast_tac ZF_cs 1]); (*----------------------------------------------------------------------*) -(* Dinf, the inverse Limit. *) +(* Dinf, the inverse Limit. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [Dinf_def] (* DinfI *) @@ -1693,7 +1693,7 @@ val lub_Dinf = result(); (*----------------------------------------------------------------------*) -(* Generalising embedddings D_m -> D_{m+1} to embeddings D_m -> D_n, *) +(* Generalising embedddings D_m -> D_{m+1} to embeddings D_m -> D_n, *) (* defined as eps(DD,ee,m,n), via e_less and e_gr. *) (*----------------------------------------------------------------------*) @@ -2257,7 +2257,7 @@ val eps_split_right = result(); (*----------------------------------------------------------------------*) -(* That was eps: D_m -> D_n, NEXT rho_emb: D_n -> Dinf. *) +(* That was eps: D_m -> D_n, NEXT rho_emb: D_n -> Dinf. *) (*----------------------------------------------------------------------*) (* Considerably shorter. *) @@ -2479,7 +2479,7 @@ val rho_proj_cont = result(); (*----------------------------------------------------------------------*) -(* Commutivity and universality. *) +(* Commutivity and universality. *) (*----------------------------------------------------------------------*) val prems = goalw Limit.thy [commute_def] (* commuteI *) diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/ListN.ML --- a/src/ZF/ex/ListN.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/ListN.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/listn +(* Title: ZF/ex/listn ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Inductive definition of lists of n elements @@ -24,7 +24,7 @@ by (rtac iffI 1); by (etac listn_induct 1); by (safe_tac (ZF_cs addSIs (list_typechecks @ - [length_Nil, length_Cons, list_into_listn]))); + [length_Nil, length_Cons, list_into_listn]))); qed "listn_iff"; goal ListN.thy "listn(A)``{n} = {l:list(A). length(l)=n}"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Ntree.ML --- a/src/ZF/ex/Ntree.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Ntree.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/Ntree.ML +(* Title: ZF/ex/Ntree.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Datatype definition n-ary branching trees @@ -22,9 +22,9 @@ (*A nicer induction rule than the standard one*) val major::prems = goal Ntree.thy - "[| t: ntree(A); \ + "[| t: ntree(A); \ \ !!x n h. [| x: A; n: nat; h: n -> ntree(A); ALL i:n. P(h`i) \ -\ |] ==> P(Branch(x,h)) \ +\ |] ==> P(Branch(x,h)) \ \ |] ==> P(t)"; by (rtac (major RS ntree.induct) 1); by (etac UN_E 1); @@ -35,14 +35,14 @@ (*Induction on ntree(A) to prove an equation*) val major::prems = goal Ntree.thy - "[| t: ntree(A); f: ntree(A)->B; g: ntree(A)->B; \ + "[| t: ntree(A); f: ntree(A)->B; g: ntree(A)->B; \ \ !!x n h. [| x: A; n: nat; h: n -> ntree(A); f O h = g O h |] ==> \ -\ f ` Branch(x,h) = g ` Branch(x,h) \ +\ f ` Branch(x,h) = g ` Branch(x,h) \ \ |] ==> f`t=g`t"; by (rtac (major RS ntree_induct) 1); by (REPEAT_SOME (ares_tac prems)); by (cut_facts_tac prems 1); -br fun_extension 1; +by (rtac fun_extension 1); by (REPEAT_SOME (ares_tac [comp_fun])); by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply]) 1); qed "ntree_induct_eqn"; @@ -77,10 +77,10 @@ (*A nicer induction rule than the standard one*) val major::prems = goal Ntree.thy - "[| t: maptree(A); \ -\ !!x n h. [| x: A; h: maptree(A) -||> maptree(A); \ -\ ALL y: field(h). P(y) \ -\ |] ==> P(Sons(x,h)) \ + "[| t: maptree(A); \ +\ !!x n h. [| x: A; h: maptree(A) -||> maptree(A); \ +\ ALL y: field(h). P(y) \ +\ |] ==> P(Sons(x,h)) \ \ |] ==> P(t)"; by (rtac (major RS maptree.induct) 1); by (REPEAT_SOME (ares_tac prems)); @@ -102,14 +102,14 @@ (*A nicer induction rule than the standard one*) val major::prems = goal Ntree.thy - "[| t: maptree2(A,B); \ + "[| t: maptree2(A,B); \ \ !!x n h. [| x: A; h: B -||> maptree2(A,B); ALL y:range(h). P(y) \ -\ |] ==> P(Sons2(x,h)) \ +\ |] ==> P(Sons2(x,h)) \ \ |] ==> P(t)"; by (rtac (major RS maptree2.induct) 1); by (REPEAT_SOME (ares_tac prems)); by (eresolve_tac [[subset_refl, Collect_subset] MRS - FiniteFun_mono RS subsetD] 1); + FiniteFun_mono RS subsetD] 1); by (dresolve_tac [FiniteFun.dom_subset RS subsetD] 1); by (dresolve_tac [Fin.dom_subset RS subsetD] 1); by (fast_tac ZF_cs 1); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Primrec.ML --- a/src/ZF/ex/Primrec.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Primrec.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/Primrec +(* Title: ZF/ex/Primrec ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Primitive Recursive Functions @@ -25,8 +25,8 @@ val pr_ss = arith_ss addsimps list.case_eqns addsimps [list_rec_Nil, list_rec_Cons, - drop_0, drop_Nil, drop_succ_Cons, - map_Nil, map_Cons] + drop_0, drop_Nil, drop_succ_Cons, + map_Nil, map_Cons] setsolver (type_auto_tac pr_typechecks); goalw Primrec.thy [SC_def] @@ -68,7 +68,7 @@ val pr_ss = pr_ss setsolver (type_auto_tac ([primrec_into_fun] @ - pr_typechecks @ primrec.intrs)); + pr_typechecks @ primrec.intrs)); goalw Primrec.thy [ACK_def] "!!i. i:nat ==> ACK(i): primrec"; by (etac nat_induct 1); @@ -112,7 +112,7 @@ val ack_ss = pr_ss addsimps [ack_0, ack_succ_0, ack_succ_succ, - ack_type, nat_into_Ord]; + ack_type, nat_into_Ord]; (*PROPERTY A 4*) goal Primrec.thy "!!i. i:nat ==> ALL j:nat. j < ack(i,j)"; @@ -121,9 +121,9 @@ by (rtac ballI 1); by (eres_inst_tac [("n","j")] nat_induct 1); by (DO_GOAL [rtac (nat_0I RS nat_0_le RS lt_trans), - asm_simp_tac ack_ss] 1); + asm_simp_tac ack_ss] 1); by (DO_GOAL [etac (succ_leI RS lt_trans1), - asm_simp_tac ack_ss] 1); + asm_simp_tac ack_ss] 1); qed "lt_ack2_lemma"; bind_thm ("lt_ack2", (lt_ack2_lemma RS bspec)); @@ -273,15 +273,15 @@ (** COMP case **) goal Primrec.thy - "!!fs. fs : list({f: primrec . \ -\ EX kf:nat. ALL l:list(nat). \ -\ f`l < ack(kf, list_add(l))}) \ -\ ==> EX k:nat. ALL l: list(nat). \ + "!!fs. fs : list({f: primrec . \ +\ EX kf:nat. ALL l:list(nat). \ +\ f`l < ack(kf, list_add(l))}) \ +\ ==> EX k:nat. ALL l: list(nat). \ \ list_add(map(%f. f ` l, fs)) < ack(k, list_add(l))"; by (etac list.induct 1); by (DO_GOAL [res_inst_tac [("x","0")] bexI, - asm_simp_tac (ack2_ss addsimps [lt_ack1, nat_0_le]), - resolve_tac nat_typechecks] 1); + asm_simp_tac (ack2_ss addsimps [lt_ack1, nat_0_le]), + resolve_tac nat_typechecks] 1); by (safe_tac ZF_cs); by (asm_simp_tac ack2_ss 1); by (rtac (ballI RS bexI) 1); @@ -292,11 +292,11 @@ qed "COMP_map_lemma"; goalw Primrec.thy [COMP_def] - "!!g. [| g: primrec; kg: nat; \ -\ ALL l:list(nat). g`l < ack(kg, list_add(l)); \ -\ fs : list({f: primrec . \ -\ EX kf:nat. ALL l:list(nat). \ -\ f`l < ack(kf, list_add(l))}) \ + "!!g. [| g: primrec; kg: nat; \ +\ ALL l:list(nat). g`l < ack(kg, list_add(l)); \ +\ fs : list({f: primrec . \ +\ EX kf:nat. ALL l:list(nat). \ +\ f`l < ack(kf, list_add(l))}) \ \ |] ==> EX k:nat. ALL l: list(nat). COMP(g,fs)`l < ack(k, list_add(l))"; by (asm_simp_tac ZF_ss 1); by (forward_tac [list_CollectD] 1); @@ -312,11 +312,11 @@ (** PREC case **) goalw Primrec.thy [PREC_def] - "!!f g. [| ALL l:list(nat). f`l #+ list_add(l) < ack(kf, list_add(l)); \ -\ ALL l:list(nat). g`l #+ list_add(l) < ack(kg, list_add(l)); \ -\ f: primrec; kf: nat; \ -\ g: primrec; kg: nat; \ -\ l: list(nat) \ + "!!f g. [| ALL l:list(nat). f`l #+ list_add(l) < ack(kf, list_add(l)); \ +\ ALL l:list(nat). g`l #+ list_add(l) < ack(kg, list_add(l)); \ +\ f: primrec; kf: nat; \ +\ g: primrec; kg: nat; \ +\ l: list(nat) \ \ |] ==> PREC(f,g)`l #+ list_add(l) < ack(succ(kf#+kg), list_add(l))"; by (etac list.elim 1); by (asm_simp_tac (ack2_ss addsimps [[nat_le_refl, lt_ack2] MRS lt_trans]) 1); @@ -325,8 +325,8 @@ by (eres_inst_tac [("n","a")] nat_induct 1); (*base case*) by (DO_GOAL [asm_simp_tac ack2_ss, rtac lt_trans, etac bspec, - assume_tac, rtac (add_le_self RS ack_lt_mono1), - REPEAT o ares_tac (ack_typechecks)] 1); + assume_tac, rtac (add_le_self RS ack_lt_mono1), + REPEAT o ares_tac (ack_typechecks)] 1); (*ind step*) by (asm_simp_tac (ack2_ss addsimps [add_succ_right]) 1); by (rtac (succ_leI RS lt_trans1) 1); @@ -343,19 +343,19 @@ qed "PREC_case_lemma"; goal Primrec.thy - "!!f g. [| f: primrec; kf: nat; \ -\ g: primrec; kg: nat; \ -\ ALL l:list(nat). f`l < ack(kf, list_add(l)); \ -\ ALL l:list(nat). g`l < ack(kg, list_add(l)) \ -\ |] ==> EX k:nat. ALL l: list(nat). \ -\ PREC(f,g)`l< ack(k, list_add(l))"; + "!!f g. [| f: primrec; kf: nat; \ +\ g: primrec; kg: nat; \ +\ ALL l:list(nat). f`l < ack(kf, list_add(l)); \ +\ ALL l:list(nat). g`l < ack(kg, list_add(l)) \ +\ |] ==> EX k:nat. ALL l: list(nat). \ +\ PREC(f,g)`l< ack(k, list_add(l))"; by (rtac (ballI RS bexI) 1); by (rtac ([add_le_self, PREC_case_lemma] MRS lt_trans1) 1); by (REPEAT (SOMEGOAL (FIRST' [test_assume_tac, - match_tac (ack_typechecks), - rtac (ack_add_bound2 RS ballI) THEN' etac bspec]))); + match_tac (ack_typechecks), + rtac (ack_add_bound2 RS ballI) THEN' etac bspec]))); qed "PREC_case"; goal Primrec.thy @@ -364,7 +364,7 @@ by (safe_tac ZF_cs); by (DEPTH_SOLVE (ares_tac ([SC_case, CONST_case, PROJ_case, COMP_case, PREC_case, - bexI, ballI] @ nat_typechecks) 1)); + bexI, ballI] @ nat_typechecks) 1)); qed "ack_bounds_primrec"; goal Primrec.thy diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/PropLog.ML --- a/src/ZF/ex/PropLog.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/PropLog.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/prop-log.ML +(* Title: ZF/ex/prop-log.ML ID: $Id$ - Author: Tobias Nipkow & Lawrence C Paulson + Author: Tobias Nipkow & Lawrence C Paulson Copyright 1992 University of Cambridge For ex/prop-log.thy. Inductive definition of propositional logic. @@ -47,7 +47,7 @@ goalw PropLog.thy [is_true_def] "is_true(#v,t) <-> v:t"; by (simp_tac (prop_rec_ss addsimps [one_not_0 RS not_sym] - setloop (split_tac [expand_if])) 1); + setloop (split_tac [expand_if])) 1); qed "is_true_Var"; goalw PropLog.thy [is_true_def] @@ -72,7 +72,7 @@ val prop_ss = prop_rec_ss addsimps prop.intrs addsimps [is_true_Fls, is_true_Var, is_true_Imp, - hyps_Fls, hyps_Var, hyps_Imp]; + hyps_Fls, hyps_Var, hyps_Imp]; (*** Proof theory of propositional logic ***) @@ -177,9 +177,9 @@ by (rtac (major RS prop.induct) 1); by (ALLGOALS (asm_simp_tac (prop_ss addsimps [thms_I, thms.H]))); by (safe_tac (ZF_cs addSEs [Fls_Imp RS weaken_left_Un1, - Fls_Imp RS weaken_left_Un2])); + Fls_Imp RS weaken_left_Un2])); by (ALLGOALS (fast_tac (ZF_cs addIs [weaken_left_Un1, weaken_left_Un2, - weaken_right, Imp_Fls]))); + weaken_right, Imp_Fls]))); qed "hyps_thms_if"; (*Key lemma for completeness; yields a set of assumptions satisfying p*) @@ -252,7 +252,7 @@ "p: prop ==> hyps(p,t) : Fin(UN v:nat. {#v, #v=>Fls})"; by (rtac (major RS prop.induct) 1); by (asm_simp_tac (prop_ss addsimps (Fin.intrs @ [UN_I, cons_iff]) - setloop (split_tac [expand_if])) 2); + setloop (split_tac [expand_if])) 2); by (ALLGOALS (asm_simp_tac (prop_ss addsimps [Un_0, Fin.emptyI, Fin_UnI]))); qed "hyps_finite"; @@ -307,7 +307,7 @@ val [finite] = goal PropLog.thy "H: Fin(prop) ==> H |- p <-> H |= p & p:prop"; by (fast_tac (ZF_cs addSEs [soundness, finite RS completeness, - thms_in_pl]) 1); + thms_in_pl]) 1); qed "thms_iff"; writeln"Reached end of file."; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/ROOT.ML --- a/src/ZF/ex/ROOT.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/ROOT.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,12 +1,12 @@ -(* Title: ZF/ex/ROOT +(* Title: ZF/ex/ROOT ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Executes miscellaneous examples for Zermelo-Fraenkel Set Theory *) -ZF_build_completed; (*Make examples fail if ZF did*) +ZF_build_completed; (*Make examples fail if ZF did*) writeln"Root file for ZF Set Theory examples"; proof_timing := true; @@ -19,22 +19,22 @@ time_use_thy "Bin"; (** Datatypes **) -time_use_thy "BT"; (*binary trees*) -time_use_thy "Data"; (*Sample datatype*) -time_use_thy "Term"; (*terms: recursion over the list functor*) -time_use_thy "TF"; (*trees/forests: mutual recursion*) -time_use_thy "Ntree"; (*variable-branching trees; function demo*) -time_use_thy "Brouwer"; (*Infinite-branching trees*) -time_use_thy "Enum"; (*Enormous enumeration type*) +time_use_thy "BT"; (*binary trees*) +time_use_thy "Data"; (*Sample datatype*) +time_use_thy "Term"; (*terms: recursion over the list functor*) +time_use_thy "TF"; (*trees/forests: mutual recursion*) +time_use_thy "Ntree"; (*variable-branching trees; function demo*) +time_use_thy "Brouwer"; (*Infinite-branching trees*) +time_use_thy "Enum"; (*Enormous enumeration type*) (** Inductive definitions **) -time_use_thy "Rmap"; (*mapping a relation over a list*) -time_use_thy "PropLog"; (*completeness of propositional logic*) +time_use_thy "Rmap"; (*mapping a relation over a list*) +time_use_thy "PropLog"; (*completeness of propositional logic*) (*two Coq examples by Christine Paulin-Mohring*) time_use_thy "ListN"; time_use_thy "Acc"; -time_use_thy "Comb"; (*Combinatory Logic example*) -time_use_thy "Primrec"; (*Primitive recursive functions*) +time_use_thy "Comb"; (*Combinatory Logic example*) +time_use_thy "Primrec"; (*Primitive recursive functions*) (** CoDatatypes **) time_use_thy "LList"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Ramsey.ML --- a/src/ZF/ex/Ramsey.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Ramsey.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/ramsey.ML +(* Title: ZF/ex/ramsey.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge Ramsey's Theorem (finite exponent 2 version) @@ -94,7 +94,7 @@ by (fast_tac (ZF_cs addSIs [Atleast0]) 1); by (asm_simp_tac arith_ss 1); by (rtac ballI 1); -by (rename_tac "n" 1); (*simplifier does NOT preserve bound names!*) +by (rename_tac "n" 1); (*simplifier does NOT preserve bound names!*) by (nat_ind_tac "n" [] 1); by (fast_tac (ZF_cs addSIs [Atleast0]) 1); by (asm_simp_tac (arith_ss addsimps [add_succ_right]) 1); @@ -157,7 +157,7 @@ \ Atleast(j,I) |] ==> \ \ Indept(cons(a,I), V, E) & Atleast(succ(j), cons(a,I))"; by (cut_facts_tac prems 1); -by (fast_tac (ZF_cs addSEs [Atleast_succI]) 1); (*34 secs*) +by (fast_tac (ZF_cs addSEs [Atleast_succI]) 1); (*34 secs*) qed "Indept_succ"; val prems = goalw Ramsey.thy [Symmetric_def,Clique_def] @@ -184,14 +184,14 @@ by (fast_tac ZF_cs 1); by (fast_tac (ZF_cs addEs [Clique_superset]) 1); (*easy -- given a Clique*) by (safe_tac ZF_cs); -by (eresolve_tac (swapify [exI]) 1); (*ignore main EX quantifier*) -by (REPEAT (ares_tac [Indept_succ] 1)); (*make a bigger Indept*) +by (eresolve_tac (swapify [exI]) 1); (*ignore main EX quantifier*) +by (REPEAT (ares_tac [Indept_succ] 1)); (*make a bigger Indept*) (*case n*) by (rtac (ram2 RS spec RS spec RS mp RS disjE) 1); by (fast_tac ZF_cs 1); by (safe_tac ZF_cs); by (rtac exI 1); -by (REPEAT (ares_tac [Clique_succ] 1)); (*make a bigger Clique*) +by (REPEAT (ares_tac [Clique_succ] 1)); (*make a bigger Clique*) by (fast_tac (ZF_cs addEs [Indept_superset]) 1); (*easy -- given an Indept*) qed "Ramsey_step_lemma"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Rmap.ML --- a/src/ZF/ex/Rmap.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Rmap.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/Rmap +(* Title: ZF/ex/Rmap ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Inductive definition of an operator to "map" a relation over a list @@ -12,7 +12,7 @@ by (rtac lfp_mono 1); by (REPEAT (rtac rmap.bnd_mono 1)); by (REPEAT (ares_tac ([Sigma_mono, list_mono, domain_mono, range_mono] @ - basic_monos) 1)); + basic_monos) 1)); qed "rmap_mono"; bind_thm ("rmap_induct", @@ -22,17 +22,17 @@ and Cons_rmap_case = rmap.mk_cases list.con_defs " : rmap(r)"; val rmap_cs = ZF_cs addIs rmap.intrs - addSEs [Nil_rmap_case, Cons_rmap_case]; + addSEs [Nil_rmap_case, Cons_rmap_case]; goal Rmap.thy "!!r. r <= A*B ==> rmap(r) <= list(A)*list(B)"; by (rtac (rmap.dom_subset RS subset_trans) 1); by (REPEAT (ares_tac [domain_rel_subset, range_rel_subset, - Sigma_mono, list_mono] 1)); + Sigma_mono, list_mono] 1)); qed "rmap_rel_type"; goal Rmap.thy "!!r. A <= domain(r) ==> list(A) <= domain(rmap(r))"; -by (resolve_tac [subsetI] 1); +by (rtac subsetI 1); by (etac list.induct 1); by (ALLGOALS (fast_tac rmap_cs)); qed "rmap_total"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/TF.ML --- a/src/ZF/ex/TF.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/TF.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/tf.ML +(* Title: ZF/ex/tf.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Trees & forests, a mutually recursive type definition. @@ -90,11 +90,11 @@ val major::prems = goal TF.thy "[| z: tree_forest(A); \ -\ !!x f r. [| x: A; f: forest(A); r: C(f) \ -\ |] ==> b(x,f,r): C(Tcons(x,f)); \ -\ c : C(Fnil); \ +\ !!x f r. [| x: A; f: forest(A); r: C(f) \ +\ |] ==> b(x,f,r): C(Tcons(x,f)); \ +\ c : C(Fnil); \ \ !!t f r1 r2. [| t: tree(A); f: forest(A); r1: C(t); r2: C(f) \ -\ |] ==> d(t,f,r1,r2): C(Fcons(t,f)) \ +\ |] ==> d(t,f,r1,r2): C(Fcons(t,f)) \ \ |] ==> TF_rec(z,b,c,d) : C(z)"; by (rtac (major RS tree_forest.induct) 1); by (ALLGOALS (asm_simp_tac (TF_rec_ss addsimps prems))); @@ -102,12 +102,12 @@ (*Mutually recursive version*) val prems = goal TF.thy - "[| !!x f r. [| x: A; f: forest(A); r: D(f) \ -\ |] ==> b(x,f,r): C(Tcons(x,f)); \ -\ c : D(Fnil); \ + "[| !!x f r. [| x: A; f: forest(A); r: D(f) \ +\ |] ==> b(x,f,r): C(Tcons(x,f)); \ +\ c : D(Fnil); \ \ !!t f r1 r2. [| t: tree(A); f: forest(A); r1: C(t); r2: D(f) \ -\ |] ==> d(t,f,r1,r2): D(Fcons(t,f)) \ -\ |] ==> (ALL t:tree(A). TF_rec(t,b,c,d) : C(t)) & \ +\ |] ==> d(t,f,r1,r2): D(Fcons(t,f)) \ +\ |] ==> (ALL t:tree(A). TF_rec(t,b,c,d) : C(t)) & \ \ (ALL f: forest(A). TF_rec(f,b,c,d) : D(f))"; by (rewtac Ball_def); by (rtac tree_forest.mutual_induct 1); @@ -136,13 +136,13 @@ qed "def_TF_rec_Fcons"; fun TF_recs def = map standard - ([def] RL [def_TF_rec_Tcons, def_TF_rec_Fnil, def_TF_rec_Fcons]); + ([def] RL [def_TF_rec_Tcons, def_TF_rec_Fnil, def_TF_rec_Fcons]); (** list_of_TF and TF_of_list **) val [list_of_TF_Tcons, list_of_TF_Fnil, list_of_TF_Fcons] = - TF_recs list_of_TF_def; + TF_recs list_of_TF_def; goalw TF.thy [list_of_TF_def] "!!z A. z: tree_forest(A) ==> list_of_TF(z) : list(tree(A))"; @@ -183,7 +183,7 @@ (** TF_preorder **) val [TF_preorder_Tcons, TF_preorder_Fnil, TF_preorder_Fcons] = - TF_recs TF_preorder_def; + TF_recs TF_preorder_def; goalw TF.thy [TF_preorder_def] "!!z A. z: tree_forest(A) ==> TF_preorder(z) : list(A)"; @@ -215,7 +215,7 @@ (*essentially the same as list induction*) val major::prems = goal TF.thy - "[| f: forest(A); \ + "[| f: forest(A); \ \ R(Fnil); \ \ !!t f. [| t: tree(A); f: forest(A); R(f) |] ==> R(Fcons(t,f)) \ \ |] ==> R(f)"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/Term.ML --- a/src/ZF/ex/Term.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/Term.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/Term.ML +(* Title: ZF/ex/Term.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Datatype definition of terms over an alphabet. @@ -90,10 +90,10 @@ (*Slightly odd typing condition on r in the second premise!*) val major::prems = goal Term.thy - "[| t: term(A); \ -\ !!x zs r. [| x: A; zs: list(term(A)); \ -\ r: list(UN t:term(A). C(t)) |] \ -\ ==> d(x, zs, r): C(Apply(x,zs)) \ + "[| t: term(A); \ +\ !!x zs r. [| x: A; zs: list(term(A)); \ +\ r: list(UN t:term(A). C(t)) |] \ +\ ==> d(x, zs, r): C(Apply(x,zs)) \ \ |] ==> term_rec(t,d) : C(t)"; by (rtac (major RS term.induct) 1); by (forward_tac [list_CollectD] 1); @@ -113,9 +113,9 @@ qed "def_term_rec"; val prems = goal Term.thy - "[| t: term(A); \ + "[| t: term(A); \ \ !!x zs r. [| x: A; zs: list(term(A)); r: list(C) |] \ -\ ==> d(x, zs, r): C \ +\ ==> d(x, zs, r): C \ \ |] ==> term_rec(t,d) : C"; by (REPEAT (ares_tac (term_rec_type::prems) 1)); by (etac (subset_refl RS UN_least RS list_mono RS subsetD) 1); @@ -208,7 +208,7 @@ goal Term.thy "!!t A. t: term(A) ==> term_size(reflect(t)) = term_size(t)"; by (etac term_induct_eqn 1); by (asm_simp_tac (term_ss addsimps [rev_map_distrib RS sym, map_compose, - list_add_rev]) 1); + list_add_rev]) 1); qed "term_size_reflect"; goal Term.thy "!!t A. t: term(A) ==> term_size(t) = length(preorder(t))"; @@ -222,7 +222,7 @@ goal Term.thy "!!t A. t: term(A) ==> reflect(reflect(t)) = t"; by (etac term_induct_eqn 1); by (asm_simp_tac (term_ss addsimps [rev_map_distrib, map_compose, - map_ident, rev_rev_ident]) 1); + map_ident, rev_rev_ident]) 1); qed "reflect_reflect_ident"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/misc.ML --- a/src/ZF/ex/misc.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/misc.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/misc +(* Title: ZF/ex/misc ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Miscellaneous examples for Zermelo-Fraenkel Set Theory @@ -82,7 +82,7 @@ fun forw_iterate tyrls rls facts 0 = facts | forw_iterate tyrls rls facts n = let val facts' = - gen_union eq_thm (forw_typechk (tyrls@facts) (facts RL rls), facts); + gen_union eq_thm (forw_typechk (tyrls@facts) (facts RL rls), facts); in forw_iterate tyrls rls facts' (n-1) end; val pastre_rls = @@ -93,49 +93,49 @@ pastre_rls [fact1,fact2,fact3] 4; val prems = goalw Perm.thy [bij_def] - "[| (h O g O f): inj(A,A); \ -\ (f O h O g): surj(B,B); \ -\ (g O f O h): surj(C,C); \ + "[| (h O g O f): inj(A,A); \ +\ (f O h O g): surj(B,B); \ +\ (g O f O h): surj(C,C); \ \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); qed "pastre1"; val prems = goalw Perm.thy [bij_def] - "[| (h O g O f): surj(A,A); \ -\ (f O h O g): inj(B,B); \ -\ (g O f O h): surj(C,C); \ + "[| (h O g O f): surj(A,A); \ +\ (f O h O g): inj(B,B); \ +\ (g O f O h): surj(C,C); \ \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); qed "pastre2"; val prems = goalw Perm.thy [bij_def] - "[| (h O g O f): surj(A,A); \ -\ (f O h O g): surj(B,B); \ -\ (g O f O h): inj(C,C); \ + "[| (h O g O f): surj(A,A); \ +\ (f O h O g): surj(B,B); \ +\ (g O f O h): inj(C,C); \ \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); qed "pastre3"; val prems = goalw Perm.thy [bij_def] - "[| (h O g O f): surj(A,A); \ -\ (f O h O g): inj(B,B); \ -\ (g O f O h): inj(C,C); \ + "[| (h O g O f): surj(A,A); \ +\ (f O h O g): inj(B,B); \ +\ (g O f O h): inj(C,C); \ \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); qed "pastre4"; val prems = goalw Perm.thy [bij_def] - "[| (h O g O f): inj(A,A); \ -\ (f O h O g): surj(B,B); \ -\ (g O f O h): inj(C,C); \ + "[| (h O g O f): inj(A,A); \ +\ (f O h O g): surj(B,B); \ +\ (g O f O h): inj(C,C); \ \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); qed "pastre5"; val prems = goalw Perm.thy [bij_def] - "[| (h O g O f): inj(A,A); \ -\ (f O h O g): inj(B,B); \ -\ (g O f O h): surj(C,C); \ + "[| (h O g O f): inj(A,A); \ +\ (f O h O g): inj(B,B); \ +\ (g O f O h): surj(C,C); \ \ f: A->B; g: B->C; h: C->A |] ==> h: bij(C,A)"; by (REPEAT (resolve_tac (IntI :: pastre_facts prems) 1)); qed "pastre6"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ex/twos_compl.ML --- a/src/ZF/ex/twos_compl.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ex/twos_compl.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ex/twos-compl.ML +(* Title: ZF/ex/twos-compl.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge ML code for Arithmetic on binary integers; the model for theory Bin diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/func.ML --- a/src/ZF/func.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/func.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/func +(* Title: ZF/func ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1991 University of Cambridge Functions in Zermelo-Fraenkel Set Theory @@ -112,7 +112,7 @@ "f: Pi(A,B) ==> : f <-> a:A & f`a = b"; by (cut_facts_tac [major RS fun_is_rel] 1); by (fast_tac (ZF_cs addSIs [major RS apply_Pair, - major RSN (2,apply_equality)]) 1); + major RSN (2,apply_equality)]) 1); qed "apply_iff"; (*Refining one Pi type to another*) @@ -207,7 +207,7 @@ "[| f : Pi(A,B); g: Pi(A,D); \ \ !!x. x:A ==> f`x = g`x |] ==> f=g"; by (REPEAT (ares_tac (prems @ (prems RL [sym]) @ - [subset_refl,equalityI,fun_subset]) 1)); + [subset_refl,equalityI,fun_subset]) 1)); qed "fun_extension"; goal ZF.thy "!!f A B. f : Pi(A,B) ==> (lam x:A. f`x) = f"; @@ -275,7 +275,7 @@ val [prem] = goalw ZF.thy [restrict_def] "A<=C ==> restrict(lam x:C. b(x), A) = (lam x:A. b(x))"; by (rtac (refl RS lam_cong) 1); -by (etac (prem RS subsetD RS beta) 1); (*easier than calling simp_tac*) +by (etac (prem RS subsetD RS beta) 1); (*easier than calling simp_tac*) qed "restrict_lam_eq"; @@ -302,8 +302,8 @@ (** The Union of 2 disjoint functions is a function **) val Un_rls = [Un_subset_iff, domain_Un_eq, SUM_Un_distrib1, prod_Un_distrib2, - Un_upper1 RSN (2, subset_trans), - Un_upper2 RSN (2, subset_trans)]; + Un_upper1 RSN (2, subset_trans), + Un_upper2 RSN (2, subset_trans)]; goal ZF.thy "!!f. [| f: A->B; g: C->D; A Int C = 0 |] ==> \ @@ -315,7 +315,7 @@ (*Solve the two cases that contradict A Int C = 0*) by (deepen_tac ZF_cs 2 2); by (deepen_tac ZF_cs 2 2); -bw function_def; +by (rewtac function_def); by (REPEAT_FIRST (dtac (spec RS spec))); by (deepen_tac ZF_cs 1 1); by (deepen_tac ZF_cs 1 1); @@ -388,10 +388,10 @@ by (rtac UN_I 1 THEN assume_tac 1); by (rtac (apply_funtype RS UN_I) 1 THEN REPEAT (ares_tac [consI1] 1)); by (simp_tac (FOL_ss addsimps cons_iff::mem_simps) 1); -by (resolve_tac [fun_extension] 1 THEN REPEAT (ares_tac [fun_extend] 1)); +by (rtac fun_extension 1 THEN REPEAT (ares_tac [fun_extend] 1)); by (etac consE 1); by (ALLGOALS (asm_simp_tac (FOL_ss addsimps [restrict, fun_extend_apply1, - fun_extend_apply2]))); + fun_extend_apply2]))); qed "cons_fun_eq"; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/ind_syntax.ML --- a/src/ZF/ind_syntax.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/ind_syntax.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/ind-syntax.ML +(* Title: ZF/ind-syntax.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Abstract Syntax functions for Inductive Definitions @@ -53,22 +53,22 @@ (*simple error-checking in the premises of an inductive definition*) fun chk_prem rec_hd (Const("op &",_) $ _ $ _) = - error"Premises may not be conjuctive" + error"Premises may not be conjuctive" | chk_prem rec_hd (Const("op :",_) $ t $ X) = - deny (Logic.occs(rec_hd,t)) "Recursion term on left of member symbol" + deny (Logic.occs(rec_hd,t)) "Recursion term on left of member symbol" | chk_prem rec_hd t = - deny (Logic.occs(rec_hd,t)) "Recursion term in side formula"; + deny (Logic.occs(rec_hd,t)) "Recursion term in side formula"; (*Return the conclusion of a rule, of the form t:X*) fun rule_concl rl = let val Const("Trueprop",_) $ (Const("op :",_) $ t $ X) = - Logic.strip_imp_concl rl + Logic.strip_imp_concl rl in (t,X) end; (*As above, but return error message if bad*) fun rule_concl_msg sign rl = rule_concl rl handle Bind => error ("Ill-formed conclusion of introduction rule: " ^ - Sign.string_of_term sign rl); + Sign.string_of_term sign rl); (*For deriving cases rules. CollectD2 discards the domain, which is redundant; read_instantiate replaces a propositional variable by a formula variable*) @@ -85,10 +85,10 @@ (*read a constructor specification*) fun read_construct sign (id, sprems, syn) = let val prems = map (readtm sign oT) sprems - val args = map (#1 o dest_mem) prems - val T = (map (#2 o dest_Free) args) ---> iT - handle TERM _ => error - "Bad variable in constructor specification" + val args = map (#1 o dest_mem) prems + val T = (map (#2 o dest_Free) args) ---> iT + handle TERM _ => error + "Bad variable in constructor specification" val name = Syntax.const_name id syn (*handle infix constructors*) in ((id,T,syn), name, args, prems) end; @@ -97,17 +97,17 @@ (*convert constructor specifications into introduction rules*) fun mk_intr_tms (rec_tm, constructs) = let fun mk_intr ((id,T,syn), name, args, prems) = - Logic.list_implies - (map mk_tprop prems, - mk_tprop (mem_const $ list_comb(Const(name,T), args) $ rec_tm)) + Logic.list_implies + (map mk_tprop prems, + mk_tprop (mem_const $ list_comb(Const(name,T), args) $ rec_tm)) in map mk_intr constructs end; val mk_all_intr_tms = flat o map mk_intr_tms o op ~~; -val Un = Const("op Un", [iT,iT]--->iT) -and empty = Const("0", iT) -and univ = Const("univ", iT-->iT) -and quniv = Const("quniv", iT-->iT); +val Un = Const("op Un", [iT,iT]--->iT) +and empty = Const("0", iT) +and univ = Const("univ", iT-->iT) +and quniv = Const("quniv", iT-->iT); (*Make a datatype's domain: form the union of its set parameters*) fun union_params rec_tm = @@ -135,9 +135,9 @@ (*Includes rules for succ and Pair since they are common constructions*) val elim_rls = [asm_rl, FalseE, succ_neq_0, sym RS succ_neq_0, - Pair_neq_0, sym RS Pair_neq_0, Pair_inject, - make_elim succ_inject, - refl_thin, conjE, exE, disjE]; + Pair_neq_0, sym RS Pair_neq_0, Pair_inject, + make_elim succ_inject, + refl_thin, conjE, exE, disjE]; (*Turns iff rules into safe elimination rules*) fun mk_free_SEs iffs = map (gen_make_elim [conjE,FalseE]) (iffs RL [iffD1]); diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/indrule.ML --- a/src/ZF/indrule.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/indrule.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/indrule.ML +(* Title: ZF/indrule.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Induction rule module -- for Inductive/Coinductive Definitions @@ -10,8 +10,8 @@ signature INDRULE = sig - val induct : thm (*main induction rule*) - val mutual_induct : thm (*mutual induction rule*) + val induct : thm (*main induction rule*) + val mutual_induct : thm (*mutual induction rule*) end; @@ -32,7 +32,7 @@ (*** Prove the main induction rule ***) -val pred_name = "P"; (*name for predicate variables*) +val pred_name = "P"; (*name for predicate variables*) val big_rec_def::part_rec_defs = Intr_elim.defs; @@ -40,20 +40,20 @@ ind_alist = [(rec_tm1,pred1),...] -- associates predicates with rec ops prem is a premise of an intr rule*) fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $ - (Const("op :",_)$t$X), iprems) = + (Const("op :",_)$t$X), iprems) = (case gen_assoc (op aconv) (ind_alist, X) of - Some pred => prem :: Ind_Syntax.mk_tprop (pred $ t) :: iprems - | None => (*possibly membership in M(rec_tm), for M monotone*) - let fun mk_sb (rec_tm,pred) = - (rec_tm, Ind_Syntax.Collect_const$rec_tm$pred) - in subst_free (map mk_sb ind_alist) prem :: iprems end) + Some pred => prem :: Ind_Syntax.mk_tprop (pred $ t) :: iprems + | None => (*possibly membership in M(rec_tm), for M monotone*) + let fun mk_sb (rec_tm,pred) = + (rec_tm, Ind_Syntax.Collect_const$rec_tm$pred) + in subst_free (map mk_sb ind_alist) prem :: iprems end) | add_induct_prem ind_alist (prem,iprems) = prem :: iprems; (*Make a premise of the induction rule.*) fun induct_prem ind_alist intr = let val quantfrees = map dest_Free (term_frees intr \\ rec_params) val iprems = foldr (add_induct_prem ind_alist) - (Logic.strip_imp_prems intr,[]) + (Logic.strip_imp_prems intr,[]) val (t,X) = Ind_Syntax.rule_concl intr val (Some pred) = gen_assoc (op aconv) (ind_alist, X) val concl = Ind_Syntax.mk_tprop (pred $ t) @@ -64,8 +64,8 @@ Intro rules with extra Vars in premises still cause some backtracking *) fun ind_tac [] 0 = all_tac | ind_tac(prem::prems) i = - DEPTH_SOLVE_1 (ares_tac [prem, refl] i) THEN - ind_tac prems (i-1); + DEPTH_SOLVE_1 (ares_tac [prem, refl] i) THEN + ind_tac prems (i-1); val pred = Free(pred_name, Ind_Syntax.iT --> Ind_Syntax.oT); @@ -76,15 +76,15 @@ prove_goalw_cterm part_rec_defs (cterm_of sign (Logic.list_implies (ind_prems, - Ind_Syntax.mk_tprop (Ind_Syntax.mk_all_imp(big_rec_tm,pred))))) + Ind_Syntax.mk_tprop (Ind_Syntax.mk_all_imp(big_rec_tm,pred))))) (fn prems => [rtac (impI RS allI) 1, - DETERM (etac Intr_elim.raw_induct 1), - (*Push Part inside Collect*) - asm_full_simp_tac (FOL_ss addsimps [Part_Collect]) 1, - REPEAT (FIRSTGOAL (eresolve_tac [CollectE, exE, conjE, disjE] ORELSE' - hyp_subst_tac)), - ind_tac (rev prems) (length prems) ]); + DETERM (etac Intr_elim.raw_induct 1), + (*Push Part inside Collect*) + asm_full_simp_tac (FOL_ss addsimps [Part_Collect]) 1, + REPEAT (FIRSTGOAL (eresolve_tac [CollectE, exE, conjE, disjE] ORELSE' + hyp_subst_tac)), + ind_tac (rev prems) (length prems) ]); (*** Prove the simultaneous induction rule ***) @@ -111,11 +111,11 @@ val pfree = Free(pred_name ^ "_" ^ rec_name, T) val frees = mk_frees "za" (binder_types T) val qconcl = - foldr Ind_Syntax.mk_all (frees, - Ind_Syntax.imp $ - (Ind_Syntax.mem_const $ foldr1 (app Pr.pair) frees $ - rec_tm) - $ (list_comb (pfree,frees))) + foldr Ind_Syntax.mk_all (frees, + Ind_Syntax.imp $ + (Ind_Syntax.mem_const $ foldr1 (app Pr.pair) frees $ + rec_tm) + $ (list_comb (pfree,frees))) in (Ind_Syntax.ap_split Pr.fsplit_const pfree (binder_types T), qconcl) end; @@ -130,20 +130,20 @@ (*To instantiate the main induction rule*) val induct_concl = Ind_Syntax.mk_tprop(Ind_Syntax.mk_all_imp(big_rec_tm, - Abs("z", Ind_Syntax.iT, - fold_bal (app Ind_Syntax.conj) - (map mk_rec_imp (Inductive.rec_tms~~preds))))) + Abs("z", Ind_Syntax.iT, + fold_bal (app Ind_Syntax.conj) + (map mk_rec_imp (Inductive.rec_tms~~preds))))) and mutual_induct_concl = Ind_Syntax.mk_tprop(fold_bal (app Ind_Syntax.conj) qconcls); val lemma = (*makes the link between the two induction rules*) prove_goalw_cterm part_rec_defs - (cterm_of sign (Logic.mk_implies (induct_concl,mutual_induct_concl))) - (fn prems => - [cut_facts_tac prems 1, - REPEAT (eresolve_tac [asm_rl, conjE, PartE, mp] 1 - ORELSE resolve_tac [allI, impI, conjI, Part_eqI] 1 - ORELSE dresolve_tac [spec, mp, Pr.fsplitD] 1)]); + (cterm_of sign (Logic.mk_implies (induct_concl,mutual_induct_concl))) + (fn prems => + [cut_facts_tac prems 1, + REPEAT (eresolve_tac [asm_rl, conjE, PartE, mp] 1 + ORELSE resolve_tac [allI, impI, conjI, Part_eqI] 1 + ORELSE dresolve_tac [spec, mp, Pr.fsplitD] 1)]); (*Mutual induction follows by freeness of Inl/Inr.*) @@ -167,42 +167,42 @@ | mutual_ind_tac(prem::prems) i = DETERM (SELECT_GOAL - ( - (*Simplify the assumptions and goal by unfolding Part and - using freeness of the Sum constructors; proves all but one + ( + (*Simplify the assumptions and goal by unfolding Part and + using freeness of the Sum constructors; proves all but one conjunct by contradiction*) - rewrite_goals_tac all_defs THEN - simp_tac (mut_ss addsimps [Part_iff]) 1 THEN - IF_UNSOLVED (*simp_tac may have finished it off!*) - ((*simplify assumptions, but don't accept new rewrite rules!*) - asm_full_simp_tac (mut_ss setmksimps (fn _=>[])) 1 THEN - (*unpackage and use "prem" in the corresponding place*) - REPEAT (rtac impI 1) THEN - rtac (rewrite_rule all_defs prem) 1 THEN - (*prem must not be REPEATed below: could loop!*) - DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' - eresolve_tac (conjE::mp::cmonos)))) - ) i) + rewrite_goals_tac all_defs THEN + simp_tac (mut_ss addsimps [Part_iff]) 1 THEN + IF_UNSOLVED (*simp_tac may have finished it off!*) + ((*simplify assumptions, but don't accept new rewrite rules!*) + asm_full_simp_tac (mut_ss setmksimps (fn _=>[])) 1 THEN + (*unpackage and use "prem" in the corresponding place*) + REPEAT (rtac impI 1) THEN + rtac (rewrite_rule all_defs prem) 1 THEN + (*prem must not be REPEATed below: could loop!*) + DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' + eresolve_tac (conjE::mp::cmonos)))) + ) i) THEN mutual_ind_tac prems (i-1); val _ = writeln " Proving the mutual induction rule..."; val mutual_induct_fsplit = prove_goalw_cterm [] - (cterm_of sign - (Logic.list_implies - (map (induct_prem (Inductive.rec_tms~~preds)) Inductive.intr_tms, - mutual_induct_concl))) - (fn prems => - [rtac (quant_induct RS lemma) 1, - mutual_ind_tac (rev prems) (length prems)]); + (cterm_of sign + (Logic.list_implies + (map (induct_prem (Inductive.rec_tms~~preds)) Inductive.intr_tms, + mutual_induct_concl))) + (fn prems => + [rtac (quant_induct RS lemma) 1, + mutual_ind_tac (rev prems) (length prems)]); (*Attempts to remove all occurrences of fsplit*) val fsplit_tac = REPEAT (SOMEGOAL (FIRST' [rtac Pr.fsplitI, - dtac Pr.fsplitD, - etac Pr.fsplitE, (*apparently never used!*) - bound_hyp_subst_tac])) + dtac Pr.fsplitD, + etac Pr.fsplitE, (*apparently never used!*) + bound_hyp_subst_tac])) THEN prune_params_tac in @@ -214,7 +214,7 @@ recursion or because it involves tuples. This saves storage.*) val mutual_induct = if length Intr_elim.rec_names > 1 orelse - is_sigma Inductive.dom_sum + is_sigma Inductive.dom_sum then rule_by_tactic fsplit_tac mutual_induct_fsplit else TrueI; end diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/intr_elim.ML --- a/src/ZF/intr_elim.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/intr_elim.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,43 +1,43 @@ -(* Title: ZF/intr_elim.ML +(* Title: ZF/intr_elim.ML ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Introduction/elimination rule module -- for Inductive/Coinductive Definitions *) -signature INDUCTIVE_ARG = (** Description of a (co)inductive def **) +signature INDUCTIVE_ARG = (** Description of a (co)inductive def **) sig val thy : theory (*new theory with inductive defs*) - val monos : thm list (*monotonicity of each M operator*) - val con_defs : thm list (*definitions of the constructors*) - val type_intrs : thm list (*type-checking intro rules*) - val type_elims : thm list (*type-checking elim rules*) + val monos : thm list (*monotonicity of each M operator*) + val con_defs : thm list (*definitions of the constructors*) + val type_intrs : thm list (*type-checking intro rules*) + val type_elims : thm list (*type-checking elim rules*) end; -signature INDUCTIVE_I = (** Terms read from the theory section **) +signature INDUCTIVE_I = (** Terms read from the theory section **) sig - val rec_tms : term list (*the recursive sets*) - val dom_sum : term (*their common domain*) - val intr_tms : term list (*terms for the introduction rules*) + val rec_tms : term list (*the recursive sets*) + val dom_sum : term (*their common domain*) + val intr_tms : term list (*terms for the introduction rules*) end; signature INTR_ELIM = sig val thy : theory (*copy of input theory*) - val defs : thm list (*definitions made in thy*) - val bnd_mono : thm (*monotonicity for the lfp definition*) - val dom_subset : thm (*inclusion of recursive set in dom*) - val intrs : thm list (*introduction rules*) - val elim : thm (*case analysis theorem*) - val mk_cases : thm list -> string -> thm (*generates case theorems*) + val defs : thm list (*definitions made in thy*) + val bnd_mono : thm (*monotonicity for the lfp definition*) + val dom_subset : thm (*inclusion of recursive set in dom*) + val intrs : thm list (*introduction rules*) + val elim : thm (*case analysis theorem*) + val mk_cases : thm list -> string -> thm (*generates case theorems*) end; -signature INTR_ELIM_AUX = (** Used to make induction rules **) +signature INTR_ELIM_AUX = (** Used to make induction rules **) sig - val raw_induct : thm (*raw induction rule from Fp.induct*) - val rec_names : string list (*names of recursive sets*) + val raw_induct : thm (*raw induction rule from Fp.induct*) + val rec_names : string list (*names of recursive sets*) end; (*prove intr/elim rules for a fixedpoint definition*) @@ -52,7 +52,7 @@ val _ = deny (big_rec_name mem map ! (stamps_of_thy Inductive.thy)) ("Definition " ^ big_rec_name ^ - " would clash with the theory of the same name!"); + " would clash with the theory of the same name!"); (*fetch fp definitions from the theory*) val big_rec_def::part_rec_defs = @@ -73,7 +73,7 @@ (cterm_of sign (Ind_Syntax.mk_tprop (Fp.bnd_mono $ dom_sum $ fp_abs))) (fn _ => [rtac (Collect_subset RS bnd_monoI) 1, - REPEAT (ares_tac (basic_monos @ Inductive.monos) 1)]); + REPEAT (ares_tac (basic_monos @ Inductive.monos) 1)]); val dom_subset = standard (big_rec_def RS Fp.subs); @@ -110,20 +110,20 @@ REPEAT (rtac refl 2), (*Typechecking; this can fail*) REPEAT (FIRSTGOAL ( dresolve_tac rec_typechecks - ORELSE' eresolve_tac (asm_rl::PartE::SigmaE2:: - Inductive.type_elims) - ORELSE' hyp_subst_tac)), + ORELSE' eresolve_tac (asm_rl::PartE::SigmaE2:: + Inductive.type_elims) + ORELSE' hyp_subst_tac)), DEPTH_SOLVE (swap_res_tac (SigmaI::subsetI::Inductive.type_intrs) 1)]; (*combines disjI1 and disjI2 to access the corresponding nested disjunct...*) val mk_disj_rls = let fun f rl = rl RS disjI1 - and g rl = rl RS disjI2 + and g rl = rl RS disjI2 in accesses_bal(f, g, asm_rl) end; val intrs = map (uncurry (prove_goalw_cterm part_rec_defs)) (map (cterm_of sign) Inductive.intr_tms ~~ - map intro_tacsf (mk_disj_rls(length Inductive.intr_tms))); + map intro_tacsf (mk_disj_rls(length Inductive.intr_tms))); (********) val _ = writeln " Proving the elimination rule..."; @@ -131,7 +131,7 @@ (*Breaks down logical connectives in the monotonic function*) val basic_elim_tac = REPEAT (SOMEGOAL (eresolve_tac (Ind_Syntax.elim_rls @ Su.free_SEs) - ORELSE' bound_hyp_subst_tac)) + ORELSE' bound_hyp_subst_tac)) THEN prune_params_tac (*Mutual recursion: collapse references to Part(D,h)*) THEN fold_tac part_rec_defs; @@ -145,7 +145,7 @@ and intrs = intrs; val elim = rule_by_tactic basic_elim_tac - (unfold RS Ind_Syntax.equals_CollectD); + (unfold RS Ind_Syntax.equals_CollectD); (*Applies freeness of the given constructors, which *must* be unfolded by the given defs. Cannot simply use the local con_defs because @@ -153,9 +153,9 @@ String s should have the form t:Si where Si is an inductive set*) fun mk_cases defs s = rule_by_tactic (rewrite_goals_tac defs THEN - basic_elim_tac THEN - fold_tac defs) - (assume_read Inductive.thy s RS elim); + basic_elim_tac THEN + fold_tac defs) + (assume_read Inductive.thy s RS elim); val raw_induct = standard ([big_rec_def, bnd_mono] MRS Fp.induct) and rec_names = rec_names diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/mono.ML --- a/src/ZF/mono.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/mono.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/mono +(* Title: ZF/mono ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Monotonicity of various operations (for lattice properties see subset.ML) @@ -187,5 +187,5 @@ (*Used in intr_elim.ML and in individual datatype definitions*) val basic_monos = [subset_refl, imp_refl, disj_mono, conj_mono, - ex_mono, Collect_mono, Part_mono, in_mono]; + ex_mono, Collect_mono, Part_mono, in_mono]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/pair.ML --- a/src/ZF/pair.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/pair.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/pair +(* Title: ZF/pair ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge Ordered pairs in Zermelo-Fraenkel Set Theory @@ -80,7 +80,7 @@ (*Also provable via rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac) - THEN prune_params_tac) + THEN prune_params_tac) (read_instantiate [("c","")] SigmaE); *) qed_goal "SigmaE2" ZF.thy "[| : Sigma(A,B); \ @@ -105,7 +105,7 @@ val pair_cs = upair_cs addSIs [SigmaI] addSEs [SigmaE2, SigmaE, Pair_inject, make_elim succ_inject, - Pair_neq_0, sym RS Pair_neq_0, succ_neq_0, sym RS succ_neq_0]; + Pair_neq_0, sym RS Pair_neq_0, succ_neq_0, sym RS succ_neq_0]; (*** Projections: fst, snd ***) @@ -166,8 +166,8 @@ qed "splitI"; val major::sigma::prems = goalw ZF.thy [split_def] - "[| split(R,z); z:Sigma(A,B); \ -\ !!x y. [| z = ; R(x,y) |] ==> P \ + "[| split(R,z); z:Sigma(A,B); \ +\ !!x y. [| z = ; R(x,y) |] ==> P \ \ |] ==> P"; by (rtac (sigma RS SigmaE) 1); by (cut_facts_tac [major] 1); @@ -184,7 +184,7 @@ fun prove_fun s = (writeln s; prove_goal ZF.thy s (fn prems => [ (cut_facts_tac prems 1), - (fast_tac (pair_cs addSIs [equalityI]) 1) ])); + (fast_tac (pair_cs addSIs [equalityI]) 1) ])); (*INCLUDED IN ZF_ss*) val mem_simps = map prove_fun diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/simpdata.ML --- a/src/ZF/simpdata.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/simpdata.ML Tue Jan 30 13:42:57 1996 +0100 @@ -40,35 +40,35 @@ (*Analyse a theorem to atomic rewrite rules*) fun atomize (conn_pairs, mem_pairs) th = let fun tryrules pairs t = - case head_of t of - Const(a,_) => - (case assoc(pairs,a) of - Some rls => flat (map (atomize (conn_pairs, mem_pairs)) - ([th] RL rls)) - | None => [th]) - | _ => [th] + case head_of t of + Const(a,_) => + (case assoc(pairs,a) of + Some rls => flat (map (atomize (conn_pairs, mem_pairs)) + ([th] RL rls)) + | None => [th]) + | _ => [th] in case concl_of th of - Const("Trueprop",_) $ P => - (case P of - Const("op :",_) $ a $ b => tryrules mem_pairs b - | Const("True",_) => [] - | Const("False",_) => [] - | A => tryrules conn_pairs A) + Const("Trueprop",_) $ P => + (case P of + Const("op :",_) $ a $ b => tryrules mem_pairs b + | Const("True",_) => [] + | Const("False",_) => [] + | A => tryrules conn_pairs A) | _ => [th] end; (*Analyse a rigid formula*) val ZF_conn_pairs = - [("Ball", [bspec]), - ("All", [spec]), - ("op -->", [mp]), - ("op &", [conjunct1,conjunct2])]; + [("Ball", [bspec]), + ("All", [spec]), + ("op -->", [mp]), + ("op &", [conjunct1,conjunct2])]; (*Analyse a:b, where b is rigid*) val ZF_mem_pairs = - [("Collect", [CollectD1,CollectD2]), - ("op -", [DiffD1,DiffD2]), - ("op Int", [IntD1,IntD2])]; + [("Collect", [CollectD1,CollectD2]), + ("op -", [DiffD1,DiffD2]), + ("op Int", [IntD1,IntD2])]; val ZF_simps = [empty_subsetI, consI1, succI1, ball_simp, if_true, if_false, @@ -86,5 +86,5 @@ val ZF_ss = FOL_ss setmksimps (map mk_meta_eq o ZF_atomize o gen_all) addsimps (ZF_simps @ mem_simps @ bquant_simps @ - Collect_simps @ UnInt_simps) + Collect_simps @ UnInt_simps) addcongs ZF_congs; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/subset.ML --- a/src/ZF/subset.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/subset.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/subset +(* Title: ZF/subset ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1991 University of Cambridge Derived rules involving subsets @@ -193,7 +193,7 @@ (*A more powerful claset for subset reasoning*) val subset_cs = subset0_cs addSIs [subset_refl,cons_subsetI,subset_consI,Union_least,UN_least,Un_least, - Inter_greatest,Int_greatest,RepFun_subset] + Inter_greatest,Int_greatest,RepFun_subset] addSIs [Un_upper1,Un_upper2,Int_lower1,Int_lower2] addIs [Union_upper,Inter_lower] addSEs [cons_subsetE]; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/thy_syntax.ML --- a/src/ZF/thy_syntax.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/thy_syntax.ML Tue Jan 30 13:42:57 1996 +0100 @@ -13,50 +13,50 @@ fun inductive_decl co = let open ThyParse fun mk_intr_name (s,_) = (*the "op" cancels any infix status*) - if Syntax.is_identifier s then "op " ^ s else "_" + if Syntax.is_identifier s then "op " ^ s else "_" fun mk_params ((((((rec_tms, sdom_sum), ipairs), - monos), con_defs), type_intrs), type_elims) = + monos), con_defs), type_intrs), type_elims) = let val big_rec_name = space_implode "_" - (map (scan_to_id o trim) rec_tms) - and srec_tms = mk_list rec_tms - and sintrs = mk_big_list (map snd ipairs) - val stri_name = big_rec_name ^ "_Intrnl" + (map (scan_to_id o trim) rec_tms) + and srec_tms = mk_list rec_tms + and sintrs = mk_big_list (map snd ipairs) + val stri_name = big_rec_name ^ "_Intrnl" in - (";\n\n\ - \structure " ^ stri_name ^ " =\n\ - \ struct\n\ - \ val _ = writeln \"" ^ co ^ - "Inductive definition " ^ big_rec_name ^ "\"\n\ - \ val rec_tms\t= map (readtm (sign_of thy) Ind_Syntax.iT) " - ^ srec_tms ^ "\n\ - \ and dom_sum\t= readtm (sign_of thy) Ind_Syntax.iT " ^ sdom_sum ^ "\n\ - \ and intr_tms\t= map (readtm (sign_of thy) propT)\n" - ^ sintrs ^ "\n\ - \ end;\n\n\ - \val thy = thy |> " ^ co ^ "Ind.add_fp_def_i \n (" ^ - stri_name ^ ".rec_tms, " ^ - stri_name ^ ".dom_sum, " ^ - stri_name ^ ".intr_tms)" - , - "structure " ^ big_rec_name ^ " =\n\ - \ let\n\ - \ val _ = writeln \"Proofs for " ^ co ^ - "Inductive definition " ^ big_rec_name ^ "\"\n\ - \ structure Result = " ^ co ^ "Ind_section_Fun\n\ - \\t (open " ^ stri_name ^ "\n\ - \\t val thy\t\t= thy\n\ - \\t val monos\t\t= " ^ monos ^ "\n\ - \\t val con_defs\t\t= " ^ con_defs ^ "\n\ - \\t val type_intrs\t= " ^ type_intrs ^ "\n\ - \\t val type_elims\t= " ^ type_elims ^ ")\n\ - \ in\n\ - \ struct\n\ - \ val " ^ mk_list (map mk_intr_name ipairs) ^ " = Result.intrs;\n\ - \ open Result\n\ - \ end\n\ - \ end;\n\n\ - \structure " ^ stri_name ^ " = struct end;\n\n" - ) + (";\n\n\ + \structure " ^ stri_name ^ " =\n\ + \ struct\n\ + \ val _ = writeln \"" ^ co ^ + "Inductive definition " ^ big_rec_name ^ "\"\n\ + \ val rec_tms\t= map (readtm (sign_of thy) Ind_Syntax.iT) " + ^ srec_tms ^ "\n\ + \ and dom_sum\t= readtm (sign_of thy) Ind_Syntax.iT " ^ sdom_sum ^ "\n\ + \ and intr_tms\t= map (readtm (sign_of thy) propT)\n" + ^ sintrs ^ "\n\ + \ end;\n\n\ + \val thy = thy |> " ^ co ^ "Ind.add_fp_def_i \n (" ^ + stri_name ^ ".rec_tms, " ^ + stri_name ^ ".dom_sum, " ^ + stri_name ^ ".intr_tms)" + , + "structure " ^ big_rec_name ^ " =\n\ + \ let\n\ + \ val _ = writeln \"Proofs for " ^ co ^ + "Inductive definition " ^ big_rec_name ^ "\"\n\ + \ structure Result = " ^ co ^ "Ind_section_Fun\n\ + \\t (open " ^ stri_name ^ "\n\ + \\t val thy\t\t= thy\n\ + \\t val monos\t\t= " ^ monos ^ "\n\ + \\t val con_defs\t\t= " ^ con_defs ^ "\n\ + \\t val type_intrs\t= " ^ type_intrs ^ "\n\ + \\t val type_elims\t= " ^ type_elims ^ ")\n\ + \ in\n\ + \ struct\n\ + \ val " ^ mk_list (map mk_intr_name ipairs) ^ " = Result.intrs;\n\ + \ open Result\n\ + \ end\n\ + \ end;\n\n\ + \structure " ^ stri_name ^ " = struct end;\n\n" + ) end val domains = "domains" $$-- enum1 "+" string --$$ "<=" -- !! string val ipairs = "intrs" $$-- repeat1 (ident -- !! string) @@ -75,42 +75,42 @@ val mk_data = mk_list o map mk_const o snd val mk_scons = mk_big_list o map mk_data fun mk_intr_name s = (*the "op" cancels any infix status*) - if Syntax.is_identifier s then "op " ^ s ^ "_I" else "_" + if Syntax.is_identifier s then "op " ^ s ^ "_I" else "_" fun mk_params ((((dom, rec_pairs), monos), type_intrs), type_elims) = let val rec_names = map (scan_to_id o trim o fst) rec_pairs - val big_rec_name = space_implode "_" rec_names - and srec_tms = mk_list (map fst rec_pairs) + val big_rec_name = space_implode "_" rec_names + and srec_tms = mk_list (map fst rec_pairs) and scons = mk_scons rec_pairs and sdom_sum = - if dom = "" then (*default domain: univ or quniv*) - "Ind_Syntax." ^ co ^ "data_domain rec_tms" - else "readtm (sign_of thy) Ind_Syntax.iT " ^ dom + if dom = "" then (*default domain: univ or quniv*) + "Ind_Syntax." ^ co ^ "data_domain rec_tms" + else "readtm (sign_of thy) Ind_Syntax.iT " ^ dom val stri_name = big_rec_name ^ "_Intrnl" val con_names = flat (map (map (trim o #1 o #1) o snd) rec_pairs) in - (";\n\n\ + (";\n\n\ \structure " ^ stri_name ^ " =\n\ \ struct\n\ \ val _ = writeln \"" ^ co ^ - "Datatype definition " ^ big_rec_name ^ "\"\n\ + "Datatype definition " ^ big_rec_name ^ "\"\n\ \ val rec_tms\t= map (readtm (sign_of thy) Ind_Syntax.iT) " ^ srec_tms ^ "\n\ \ val dom_sum\t= " ^ sdom_sum ^ "\n\ \ and con_ty_lists\t= Ind_Syntax.read_constructs (sign_of thy)\n" - ^ scons ^ "\n\ + ^ scons ^ "\n\ \ val intr_tms\t= Ind_Syntax.mk_all_intr_tms (rec_tms, con_ty_lists)\n\ \ end;\n\n\ \val thy = thy |> " ^ co ^ "Ind.add_constructs_def(" ^ - mk_list (map quote rec_names) ^ ", " ^ + mk_list (map quote rec_names) ^ ", " ^ stri_name ^ ".con_ty_lists) \n\ \ |> " ^ co ^ "Ind.add_fp_def_i \n (" ^ - stri_name ^ ".rec_tms, " ^ + stri_name ^ ".rec_tms, " ^ stri_name ^ ".dom_sum, " ^ stri_name ^ ".intr_tms)" , - "structure " ^ big_rec_name ^ " =\n\ + "structure " ^ big_rec_name ^ " =\n\ \ let\n\ \ val _ = writeln \"Proofs for " ^ co ^ - "Datatype definition " ^ big_rec_name ^ "\"\n\ + "Datatype definition " ^ big_rec_name ^ "\"\n\ \ structure Result = " ^ co ^ "Data_section_Fun\n\ \\t (open " ^ stri_name ^ "\n\ \\t val thy\t\t= thy\n\ @@ -125,8 +125,8 @@ \ end\n\ \ end;\n\n\ \structure " ^ stri_name ^ " = struct end;\n\n" - ) - end + ) + end fun optstring s = optional (s $$-- string) "\"[]\"" >> trim val string_list = "(" $$-- list1 string --$$ ")" || ThyParse.empty; val construct = name -- string_list -- opt_mixfix; @@ -140,13 +140,13 @@ (** Section specifications **) val user_keywords = ["inductive", "coinductive", "datatype", "codatatype", - "and", "|", "<=", "domains", "intrs", "monos", - "con_defs", "type_intrs", "type_elims"]; + "and", "|", "<=", "domains", "intrs", "monos", + "con_defs", "type_intrs", "type_elims"]; val user_sections = [("inductive", inductive_decl ""), - ("coinductive", inductive_decl "Co"), - ("datatype", datatype_decl ""), - ("codatatype", datatype_decl "Co")]; + ("coinductive", inductive_decl "Co"), + ("datatype", datatype_decl ""), + ("codatatype", datatype_decl "Co")]; end; diff -r 5a6f2aabd538 -r 6bcb44e4d6e5 src/ZF/upair.ML --- a/src/ZF/upair.ML Mon Jan 29 14:16:13 1996 +0100 +++ b/src/ZF/upair.ML Tue Jan 30 13:42:57 1996 +0100 @@ -1,6 +1,6 @@ -(* Title: ZF/upair +(* Title: ZF/upair ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1991 University of Cambridge UNORDERED pairs in Zermelo-Fraenkel Set Theory @@ -14,9 +14,9 @@ (*** Lemmas about power sets ***) -val Pow_bottom = empty_subsetI RS PowI; (* 0 : Pow(B) *) -val Pow_top = subset_refl RS PowI; (* A : Pow(A) *) -val Pow_neq_0 = Pow_top RSN (2,equals0D); (* Pow(a)=0 ==> P *) +val Pow_bottom = empty_subsetI RS PowI; (* 0 : Pow(B) *) +val Pow_top = subset_refl RS PowI; (* A : Pow(A) *) +val Pow_neq_0 = Pow_top RSN (2,equals0D); (* Pow(a)=0 ==> P *) (*** Unordered pairs - Upair ***) @@ -176,7 +176,7 @@ "[| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a" (fn [pa,eq] => [ (fast_tac (lemmas_cs addSIs [singletonI,pa] addIs [equalityI] - addEs [eq RS subst]) 1) ]); + addEs [eq RS subst]) 1) ]); (* Only use this if you already know EX!x. P(x) *) qed_goal "the_equality2" ZF.thy @@ -194,8 +194,8 @@ qed_goal "theI2" ZF.thy "[| EX! x. P(x); !!x. P(x) ==> Q(x) |] ==> Q(THE x.P(x))" (fn prems => [ resolve_tac prems 1, - rtac theI 1, - resolve_tac prems 1 ]); + rtac theI 1, + resolve_tac prems 1 ]); (*the_cong is no longer necessary: if (ALL y.P(y)<->Q(y)) then (THE x.P(x)) rewrites to (THE x. Q(x)) *) @@ -238,8 +238,8 @@ qed_goal "expand_if" ZF.thy "P(if(Q,x,y)) <-> ((Q --> P(x)) & (~Q --> P(y)))" (fn _=> [ (excluded_middle_tac "Q" 1), - (asm_simp_tac if_ss 1), - (asm_simp_tac if_ss 1) ]); + (asm_simp_tac if_ss 1), + (asm_simp_tac if_ss 1) ]); qed_goal "if_iff" ZF.thy "a: if(P,x,y) <-> P & a:x | ~P & a:y" (fn _=> [ (simp_tac (if_ss setloop split_tac [expand_if]) 1) ]); @@ -247,7 +247,7 @@ qed_goal "if_type" ZF.thy "[| P ==> a: A; ~P ==> b: A |] ==> if(P,a,b): A" (fn prems=> [ (simp_tac - (if_ss addsimps prems setloop split_tac [expand_if]) 1) ]); + (if_ss addsimps prems setloop split_tac [expand_if]) 1) ]); (*** Foundation lemmas ***) @@ -259,7 +259,7 @@ (etac equals0D 1), (rtac consI1 1), (fast_tac (lemmas_cs addIs [consI1,consI2] - addSEs [consE,equalityE]) 1) ]); + addSEs [consE,equalityE]) 1) ]); (*was called mem_anti_refl*) qed_goal "mem_irrefl" ZF.thy "a:a ==> P"