# HG changeset patch # User regensbu # Date 804436120 -7200 # Node ID 74be52691d626a198205930b424389cf7fabb5c6 # Parent cbd32a0f2f41995b278edd26208b6e3e689b39e5 The curried version of HOLCF is now just called HOLCF. The old uncurried version is no longer supported diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cfun1.ML --- a/src/HOLCF/Cfun1.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cfun1.ML Thu Jun 29 16:28:40 1995 +0200 @@ -16,7 +16,7 @@ (fn prems => [ (rtac (mem_Collect_eq RS ssubst) 1), - (rtac contX_id 1) + (rtac cont_id 1) ]); @@ -24,14 +24,14 @@ (* less_cfun is a partial order on type 'a -> 'b *) (* ------------------------------------------------------------------------ *) -qed_goalw "refl_less_cfun" Cfun1.thy [less_cfun_def] "less_cfun(f,f)" +qed_goalw "refl_less_cfun" Cfun1.thy [less_cfun_def] "less_cfun f f" (fn prems => [ (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), @@ -44,7 +44,7 @@ ]); 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), @@ -57,14 +57,14 @@ (* ------------------------------------------------------------------------ *) 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) ]); -qed_goal "cfun_fun_cong" Cfun1.thy "f=g ==> f[x] = g[x]" +qed_goal "cfun_fun_cong" Cfun1.thy "f=g ==> f`x = g`x" (fn prems => [ (cut_facts_tac prems 1), @@ -72,7 +72,7 @@ (rtac refl 1) ]); -qed_goal "cfun_arg_cong" Cfun1.thy "x=y ==> f[x] = f[y]" +qed_goal "cfun_arg_cong" Cfun1.thy "x=y ==> f`x = f`y" (fn prems => [ (cut_facts_tac prems 1), @@ -86,7 +86,7 @@ (* additional lemma about the isomorphism between -> and Cfun *) (* ------------------------------------------------------------------------ *) -qed_goal "Abs_Cfun_inverse2" Cfun1.thy "contX(f) ==> fapp(fabs(f)) = f" +qed_goal "Abs_Cfun_inverse2" Cfun1.thy "cont(f) ==> fapp(fabs(f)) = f" (fn prems => [ (cut_facts_tac prems 1), @@ -100,7 +100,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "Cfunapp2" Cfun1.thy - "contX(f) ==> (fabs(f))[x] = f(x)" + "cont(f) ==> (fabs f)`x = f x" (fn prems => [ (cut_facts_tac prems 1), @@ -112,7 +112,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "beta_cfun" Cfun1.thy - "contX(c1) ==> (LAM x .c1(x))[u] = c1(u)" + "cont(c1) ==> (LAM x .c1 x)`u = c1 u" (fn prems => [ (cut_facts_tac prems 1), @@ -120,3 +120,6 @@ (atac 1) ]); + + + diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cfun1.thy --- a/src/HOLCF/Cfun1.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cfun1.thy Thu Jun 29 16:28:40 1995 +0200 @@ -18,8 +18,7 @@ consts Cfun :: "('a => 'b)set" - fapp :: "('a -> 'b)=>('a => 'b)" ("(_[_])" [1000,0] 1000) - (* usually Rep_Cfun *) + fapp :: "('a -> 'b)=>('a => 'b)" (* usually Rep_Cfun *) (* application *) fabs :: "('a => 'b)=>('a -> 'b)" (binder "LAM " 10) @@ -28,18 +27,24 @@ less_cfun :: "[('a -> 'b),('a -> 'b)]=>bool" -rules +syntax "@fapp" :: "('a -> 'b)=>('a => 'b)" ("_`_" [999,1000] 999) + +translations "f`x" == "fapp f x" - Cfun_def "Cfun == {f. contX(f)}" +defs + Cfun_def "Cfun == {f. cont(f)}" + +rules (*faking a type definition... *) (* -> is isomorphic to Cfun *) - Rep_Cfun "fapp(fo):Cfun" - Rep_Cfun_inverse "fabs(fapp(fo)) = fo" - Abs_Cfun_inverse "f:Cfun ==> fapp(fabs(f))=f" + Rep_Cfun "fapp fo : Cfun" + Rep_Cfun_inverse "fabs (fapp fo) = fo" + Abs_Cfun_inverse "f:Cfun ==> fapp(fabs f) = f" +defs (*defining the abstract constants*) - less_cfun_def "less_cfun(fo1,fo2) == ( fapp(fo1) << fapp(fo2) )" + less_cfun_def "less_cfun fo1 fo2 == ( fapp fo1 << fapp fo2 )" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cfun2.ML --- a/src/HOLCF/Cfun2.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cfun2.ML Thu Jun 29 16:28:40 1995 +0200 @@ -29,7 +29,7 @@ [ (rtac (less_cfun RS ssubst) 1), (rtac (Abs_Cfun_inverse2 RS ssubst) 1), - (rtac contX_const 1), + (rtac cont_const 1), (fold_goals_tac [UU_fun_def]), (rtac minimal_fun 1) ]); @@ -37,35 +37,34 @@ (* ------------------------------------------------------------------------ *) (* fapp yields continuous functions in 'a => 'b *) (* this is continuity of fapp in its 'second' argument *) -(* contX_fapp2 ==> monofun_fapp2 & contlub_fapp2 *) +(* cont_fapp2 ==> monofun_fapp2 & contlub_fapp2 *) (* ------------------------------------------------------------------------ *) -qed_goal "contX_fapp2" Cfun2.thy "contX(fapp(fo))" +qed_goal "cont_fapp2" Cfun2.thy "cont(fapp(fo))" (fn prems => [ - (res_inst_tac [("P","contX")] CollectD 1), + (res_inst_tac [("P","cont")] CollectD 1), (fold_goals_tac [Cfun_def]), (rtac Rep_Cfun 1) ]); -val monofun_fapp2 = contX_fapp2 RS contX2mono; +val monofun_fapp2 = cont_fapp2 RS cont2mono; (* monofun(fapp(?fo1)) *) -val contlub_fapp2 = contX_fapp2 RS contX2contlub; +val contlub_fapp2 = cont_fapp2 RS cont2contlub; (* contlub(fapp(?fo1)) *) (* ------------------------------------------------------------------------ *) -(* expanded thms contX_fapp2, contlub_fapp2 *) -(* looks nice with mixfix syntac _[_] *) +(* expanded thms cont_fapp2, contlub_fapp2 *) +(* looks nice with mixfix syntac *) (* ------------------------------------------------------------------------ *) -val contX_cfun_arg = (contX_fapp2 RS contXE RS spec RS mp); -(* is_chain(?x1) ==> range(%i. ?fo3[?x1(i)]) <<| ?fo3[lub(range(?x1))] *) +val cont_cfun_arg = (cont_fapp2 RS contE RS spec RS mp); +(* is_chain(?x1) ==> range (%i. ?fo3`(?x1 i)) <<| ?fo3`(lub (range ?x1)) *) val contlub_cfun_arg = (contlub_fapp2 RS contlubE RS spec RS mp); -(* is_chain(?x1) ==> ?fo4[lub(range(?x1))] = lub(range(%i. ?fo4[?x1(i)])) *) - +(* is_chain(?x1) ==> ?fo4`(lub (range ?x1)) = lub (range (%i. ?fo4`(?x1 i))) *) (* ------------------------------------------------------------------------ *) @@ -84,7 +83,7 @@ (* monotonicity of application fapp in mixfix syntax [_]_ *) (* ------------------------------------------------------------------------ *) -qed_goal "monofun_cfun_fun" Cfun2.thy "f1 << f2 ==> f1[x] << f2[x]" +qed_goal "monofun_cfun_fun" Cfun2.thy "f1 << f2 ==> f1`x << f2`x" (fn prems => [ (cut_facts_tac prems 1), @@ -95,14 +94,14 @@ val monofun_cfun_arg = (monofun_fapp2 RS monofunE RS spec RS spec RS mp); -(* ?x2 << ?x1 ==> ?fo5[?x2] << ?fo5[?x1] *) +(* ?x2 << ?x1 ==> ?fo5`?x2 << ?fo5`?x1 *) (* ------------------------------------------------------------------------ *) (* monotonicity of fapp in both arguments in mixfix syntax [_]_ *) (* ------------------------------------------------------------------------ *) qed_goal "monofun_cfun" Cfun2.thy - "[|f1< f1[x1] << f2[x2]" + "[|f1< f1`x1 << f2`x2" (fn prems => [ (cut_facts_tac prems 1), @@ -118,7 +117,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "ch2ch_fappR" Cfun2.thy - "is_chain(Y) ==> is_chain(%i. f[Y(i)])" + "is_chain(Y) ==> is_chain(%i. f`(Y i))" (fn prems => [ (cut_facts_tac prems 1), @@ -127,7 +126,7 @@ val ch2ch_fappL = (monofun_fapp1 RS ch2ch_MF2L); -(* is_chain(?F) ==> is_chain(%i. ?F(i)[?x]) *) +(* is_chain(?F) ==> is_chain (%i. ?F i`?x) *) (* ------------------------------------------------------------------------ *) @@ -136,7 +135,7 @@ (* ------------------------------------------------------------------------ *) 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), @@ -153,8 +152,8 @@ 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)]))))" +\ 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), @@ -169,12 +168,12 @@ (* the lub of a chain of cont. functions is continuous *) (* ------------------------------------------------------------------------ *) -qed_goal "contX_lubcfun" Cfun2.thy - "is_chain(F) ==> contX(% x.lub(range(% j.F(j)[x])))" +qed_goal "cont_lubcfun" Cfun2.thy + "is_chain(F) ==> cont(% x.lub(range(% j.F(j)`x)))" (fn prems => [ (cut_facts_tac prems 1), - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (etac lub_cfun_mono 1), (rtac contlubI 1), (strip_tac 1), @@ -189,7 +188,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "lub_cfun" Cfun2.thy - "is_chain(CCF) ==> range(CCF) <<| fabs(% x.lub(range(% i.CCF(i)[x])))" + "is_chain(CCF) ==> range(CCF) <<| (LAM x.lub(range(% i.CCF(i)`x)))" (fn prems => [ (cut_facts_tac prems 1), @@ -199,13 +198,13 @@ (rtac allI 1), (rtac (less_cfun RS ssubst) 1), (rtac (Abs_Cfun_inverse2 RS ssubst) 1), - (etac contX_lubcfun 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 contX_lubcfun 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) @@ -213,7 +212,7 @@ val thelub_cfun = (lub_cfun RS thelubI); (* -is_chain(?CCF1) ==> lub(range(?CCF1)) = fabs(%x. lub(range(%i. ?CCF1(i)[x]))) +is_chain(?CCF1) ==> lub (range ?CCF1) = (LAM x. lub (range (%i. ?CCF1 i`x))) *) qed_goal "cpo_cfun" Cfun2.thy @@ -230,7 +229,7 @@ (* Extensionality in 'a -> 'b *) (* ------------------------------------------------------------------------ *) -qed_goal "ext_cfun" Cfun1.thy "(!!x. f[x] = g[x]) ==> f = g" +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), @@ -245,7 +244,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "semi_monofun_fabs" Cfun2.thy - "[|contX(f);contX(g);f<fabs(f)<fabs(f)< [ (rtac (less_cfun RS iffD2) 1), @@ -260,14 +259,14 @@ (* Extenionality wrt. << in 'a -> 'b *) (* ------------------------------------------------------------------------ *) -qed_goal "less_cfun2" Cfun2.thy "(!!x. f[x] << g[x]) ==> f << g" +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 contX_fapp2 1), - (rtac contX_fapp2 1), + (rtac cont_fapp2 1), + (rtac cont_fapp2 1), (rtac (less_fun RS iffD2) 1), (rtac allI 1), (resolve_tac prems 1) diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cfun2.thy --- a/src/HOLCF/Cfun2.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cfun2.thy Thu Jun 29 16:28:40 1995 +0200 @@ -21,7 +21,7 @@ inst_cfun_po "((op <<)::['a->'b,'a->'b]=>bool) = less_cfun" -(* definitions *) +defs (* The least element in type 'a->'b *) UU_cfun_def "UU_cfun == fabs(% x.UU)" diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cfun3.ML --- a/src/HOLCF/Cfun3.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cfun3.ML Thu Jun 29 16:28:40 1995 +0200 @@ -20,7 +20,7 @@ (rtac (thelub_cfun RS ssubst) 1), (atac 1), (rtac (Cfunapp2 RS ssubst) 1), - (etac contX_lubcfun 1), + (etac cont_lubcfun 1), (rtac (thelub_fun RS ssubst) 1), (etac (monofun_fapp1 RS ch2ch_monofun) 1), (rtac refl 1) @@ -28,25 +28,25 @@ (* ------------------------------------------------------------------------ *) -(* the contX property for fapp in its first argument *) +(* the cont property for fapp in its first argument *) (* ------------------------------------------------------------------------ *) -qed_goal "contX_fapp1" Cfun3.thy "contX(fapp)" +qed_goal "cont_fapp1" Cfun3.thy "cont(fapp)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_fapp1 1), (rtac contlub_fapp1 1) ]); (* ------------------------------------------------------------------------ *) -(* contlub, contX properties of fapp in its first argument in mixfix _[_] *) +(* contlub, cont properties of fapp in its first argument in mixfix _[_] *) (* ------------------------------------------------------------------------ *) qed_goal "contlub_cfun_fun" Cfun3.thy "is_chain(FY) ==>\ -\ lub(range(FY))[x] = lub(range(%i.FY(i)[x]))" +\ lub(range FY)`x = lub(range (%i.FY(i)`x))" (fn prems => [ (cut_facts_tac prems 1), @@ -58,9 +58,9 @@ ]); -qed_goal "contX_cfun_fun" Cfun3.thy +qed_goal "cont_cfun_fun" Cfun3.thy "is_chain(FY) ==>\ -\ range(%i.FY(i)[x]) <<| lub(range(FY))[x]" +\ range(%i.FY(i)`x) <<| lub(range FY)`x" (fn prems => [ (cut_facts_tac prems 1), @@ -71,26 +71,26 @@ (* ------------------------------------------------------------------------ *) -(* contlub, contX properties of fapp in both argument in mixfix _[_] *) +(* contlub, cont properties of fapp in both argument in mixfix _[_] *) (* ------------------------------------------------------------------------ *) qed_goal "contlub_cfun" Cfun3.thy "[|is_chain(FY);is_chain(TY)|] ==>\ -\ lub(range(FY))[lub(range(TY))] = lub(range(%i.FY(i)[TY(i)]))" +\ (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 contX_fapp1 1), + (rtac cont_fapp1 1), (rtac allI 1), - (rtac contX_fapp2 1), + (rtac cont_fapp2 1), (atac 1), (atac 1) ]); -qed_goal "contX_cfun" Cfun3.thy +qed_goal "cont_cfun" Cfun3.thy "[|is_chain(FY);is_chain(TY)|] ==>\ -\ range(%i.FY(i)[TY(i)]) <<| lub(range(FY))[lub(range(TY))]" +\ range(%i.(FY i)`(TY i)) <<| (lub (range FY))`(lub(range TY))" (fn prems => [ (cut_facts_tac prems 1), @@ -106,32 +106,32 @@ (* ------------------------------------------------------------------------ *) -(* contX2contX lemma for fapp *) +(* cont2cont lemma for fapp *) (* ------------------------------------------------------------------------ *) -qed_goal "contX2contX_fapp" Cfun3.thy - "[|contX(%x.ft(x));contX(%x.tt(x))|] ==> contX(%x.(ft(x))[tt(x)])" +qed_goal "cont2cont_fapp" Cfun3.thy + "[|cont(%x.ft x);cont(%x.tt x)|] ==> cont(%x. (ft x)`(tt x))" (fn prems => [ (cut_facts_tac prems 1), - (rtac contX2contX_app2 1), - (rtac contX2contX_app2 1), - (rtac contX_const 1), - (rtac contX_fapp1 1), + (rtac cont2cont_app2 1), + (rtac cont2cont_app2 1), + (rtac cont_const 1), + (rtac cont_fapp1 1), (atac 1), - (rtac contX_fapp2 1), + (rtac cont_fapp2 1), (atac 1) ]); (* ------------------------------------------------------------------------ *) -(* contX2mono Lemma for %x. LAM y. c1(x,y) *) +(* cont2mono Lemma for %x. LAM y. c1(x)(y) *) (* ------------------------------------------------------------------------ *) -qed_goal "contX2mono_LAM" Cfun3.thy - "[|!x.contX(c1(x)); !y.monofun(%x.c1(x,y))|] ==>\ -\ monofun(%x. LAM y. c1(x,y))" +qed_goal "cont2mono_LAM" Cfun3.thy + "[|!x.cont(c1 x); !y.monofun(%x.c1 x y)|] ==>\ +\ monofun(%x. LAM y. c1 x y)" (fn prems => [ (cut_facts_tac prems 1), @@ -148,70 +148,70 @@ ]); (* ------------------------------------------------------------------------ *) -(* contX2contX Lemma for %x. LAM y. c1(x,y) *) +(* cont2cont Lemma for %x. LAM y. c1 x y) *) (* ------------------------------------------------------------------------ *) -qed_goal "contX2contX_LAM" Cfun3.thy - "[| !x.contX(c1(x)); !y.contX(%x.c1(x,y)) |] ==> contX(%x. LAM y. c1(x,y))" +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 monocontlub2contX 1), - (etac contX2mono_LAM 1), - (rtac (contX2mono RS allI) 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 (contX2mono_LAM RS ch2ch_monofun) 1), + (rtac (cont2mono_LAM RS ch2ch_monofun) 1), (atac 1), - (rtac (contX2mono RS allI) 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 (contX2contlub RS contlubE + (rtac (cont2contlub RS contlubE RS spec RS mp ) 1), (etac spec 1), (atac 1) ]); (* ------------------------------------------------------------------------ *) -(* elimination of quantifier in premisses of contX2contX_LAM yields good *) -(* lemma for the contX tactic *) +(* elimination of quantifier in premisses of cont2cont_LAM yields good *) +(* lemma for the cont tactic *) (* ------------------------------------------------------------------------ *) -val contX2contX_LAM2 = (allI RSN (2,(allI RS contX2contX_LAM))); +val cont2cont_LAM2 = (allI RSN (2,(allI RS cont2cont_LAM))); (* - [| !!x. contX(?c1.0(x)); !!y. contX(%x. ?c1.0(x,y)) |] ==> - contX(%x. LAM y. ?c1.0(x,y)) +[| !!x. cont (?c1.0 x); + !!y. cont (%x. ?c1.0 x y) |] ==> cont (%x. LAM y. ?c1.0 x y) *) (* ------------------------------------------------------------------------ *) -(* contX2contX tactic *) +(* cont2cont tactic *) (* ------------------------------------------------------------------------ *) -val contX_lemmas = [contX_const, contX_id, contX_fapp2, - contX2contX_fapp,contX2contX_LAM2]; +val cont_lemmas = [cont_const, cont_id, cont_fapp2, + cont2cont_fapp,cont2cont_LAM2]; -val contX_tac = (fn i => (resolve_tac contX_lemmas i)); +val cont_tac = (fn i => (resolve_tac cont_lemmas i)); -val contX_tacR = (fn i => (REPEAT (contX_tac i))); +val cont_tacR = (fn i => (REPEAT (cont_tac i))); (* ------------------------------------------------------------------------ *) (* function application _[_] is strict in its first arguments *) (* ------------------------------------------------------------------------ *) -qed_goal "strict_fapp1" Cfun3.thy "(UU::'a->'b)[x] = (UU::'b)" +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), - (contX_tac 1), + (cont_tac 1), (rtac refl 1) ]); @@ -224,19 +224,15 @@ "Istrictify(f)(UU)= (UU)" (fn prems => [ - (rtac select_equality 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) + (simp_tac HOL_ss 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), - (rtac select_equality 1), - (fast_tac HOL_cs 1), - (fast_tac HOL_cs 1) + (asm_simp_tac HOL_ss 1) ]); qed_goal "monofun_Istrictify1" Cfun3.thy "monofun(Istrictify)" @@ -295,7 +291,7 @@ (rtac (thelub_cfun RS ssubst) 1), (atac 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_lubcfun 1), + (rtac cont_lubcfun 1), (atac 1), (rtac refl 1), (hyp_subst_tac 1), @@ -326,7 +322,7 @@ (rtac Istrictify1 1), (rtac (Istrictify2 RS ssubst) 1), (atac 1), - (res_inst_tac [("s","lub(range(%i. f[Y(i)]))")] trans 1), + (res_inst_tac [("s","lub(range(%i. f`(Y i)))")] trans 1), (rtac contlub_cfun_arg 1), (atac 1), (rtac lub_equal2 1), @@ -347,38 +343,38 @@ ]); -val contX_Istrictify1 = (contlub_Istrictify1 RS - (monofun_Istrictify1 RS monocontlub2contX)); +val cont_Istrictify1 = (contlub_Istrictify1 RS + (monofun_Istrictify1 RS monocontlub2cont)); -val contX_Istrictify2 = (contlub_Istrictify2 RS - (monofun_Istrictify2 RS monocontlub2contX)); +val cont_Istrictify2 = (contlub_Istrictify2 RS + (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), - (contX_tac 1), - (rtac contX_Istrictify2 1), - (rtac contX2contX_CF1L 1), - (rtac contX_Istrictify1 1), + (cont_tac 1), + (rtac cont_Istrictify2 1), + (rtac cont2cont_CF1L 1), + (rtac cont_Istrictify1 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Istrictify2 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), - (contX_tac 1), - (rtac contX_Istrictify2 1), - (rtac contX2contX_CF1L 1), - (rtac contX_Istrictify1 1), + (cont_tac 1), + (rtac cont_Istrictify2 1), + (rtac cont2cont_CF1L 1), + (rtac cont_Istrictify1 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Istrictify2 1), + (rtac cont_Istrictify2 1), (rtac Istrictify2 1), (resolve_tac prems 1) ]); @@ -392,12 +388,12 @@ strictify2]; (* ------------------------------------------------------------------------ *) -(* use contX_tac as autotac. *) +(* use cont_tac as autotac. *) (* ------------------------------------------------------------------------ *) val Cfun_ss = HOL_ss addsimps Cfun_rews setsolver (fn thms => (resolve_tac (TrueI::refl::thms)) ORELSE' atac ORELSE' - (fn i => DEPTH_SOLVE_1 (contX_tac i)) + (fn i => DEPTH_SOLVE_1 (cont_tac i)) ); diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cfun3.thy --- a/src/HOLCF/Cfun3.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cfun3.thy Thu Jun 29 16:28:40 1995 +0200 @@ -19,13 +19,11 @@ inst_cfun_pcpo "(UU::'a->'b) = UU_cfun" -Istrictify_def "Istrictify(f,x) == (@z. - ( x=UU --> z = UU) - & (~x=UU --> z = f[x]))" +defs -strictify_def "strictify == (LAM f x.Istrictify(f,x))" +Istrictify_def "Istrictify f x == if x=UU then UU else f`x" + +strictify_def "strictify == (LAM f x.Istrictify f x)" end - - diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cont.ML --- a/src/HOLCF/Cont.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cont.ML Thu Jun 29 16:28:40 1995 +0200 @@ -31,16 +31,16 @@ ]); -qed_goalw "contXI" Cont.thy [contX] - "! Y. is_chain(Y) --> range(% i.f(Y(i))) <<| f(lub(range(Y))) ==> contX(f)" +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) ]); -qed_goalw "contXE" Cont.thy [contX] - "contX(f) ==> ! Y. is_chain(Y) --> range(% i.f(Y(i))) <<| f(lub(range(Y)))" +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), @@ -66,7 +66,7 @@ (* ------------------------------------------------------------------------ *) (* the main purpose of cont.thy is to show: *) -(* monofun(f) & contlub(f) <==> contX(f) *) +(* monofun(f) & contlub(f) <==> cont(f) *) (* ------------------------------------------------------------------------ *) (* ------------------------------------------------------------------------ *) @@ -100,11 +100,11 @@ ]); (* ------------------------------------------------------------------------ *) -(* left to right: monofun(f) & contlub(f) ==> contX(f) *) +(* left to right: monofun(f) & contlub(f) ==> cont(f) *) (* ------------------------------------------------------------------------ *) -qed_goalw "monocontlub2contX" Cont.thy [contX] - "[|monofun(f);contlub(f)|] ==> contX(f)" +qed_goalw "monocontlub2cont" Cont.thy [cont] + "[|monofun(f);contlub(f)|] ==> cont(f)" (fn prems => [ (cut_facts_tac prems 1), @@ -120,49 +120,49 @@ (* first a lemma about binary chains *) (* ------------------------------------------------------------------------ *) -qed_goal "binchain_contX" Cont.thy -"[| contX(f); x << y |] ==> range(%i. f(if(i = 0,x,y))) <<| f(y)" +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 (contXE RS spec RS mp) 2), + (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: contX(f) ==> monofun(f) & contlub(f) *) -(* part1: contX(f) ==> monofun(f *) +(* right to left: cont(f) ==> monofun(f) & contlub(f) *) +(* part1: cont(f) ==> monofun(f *) (* ------------------------------------------------------------------------ *) -qed_goalw "contX2mono" Cont.thy [monofun] - "contX(f) ==> monofun(f)" +qed_goalw "cont2mono" Cont.thy [monofun] + "cont(f) ==> monofun(f)" (fn prems => [ (cut_facts_tac prems 1), (strip_tac 1), - (res_inst_tac [("s","if(0 = 0,x,y)")] subst 1), - (rtac (binchain_contX RS is_ub_lub) 2), + (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 nat_ss 1) ]); (* ------------------------------------------------------------------------ *) -(* right to left: contX(f) ==> monofun(f) & contlub(f) *) -(* part2: contX(f) ==> contlub(f) *) +(* right to left: cont(f) ==> monofun(f) & contlub(f) *) +(* part2: cont(f) ==> contlub(f) *) (* ------------------------------------------------------------------------ *) -qed_goalw "contX2contlub" Cont.thy [contlub] - "contX(f) ==> contlub(f)" +qed_goalw "cont2contlub" Cont.thy [contlub] + "cont(f) ==> contlub(f)" (fn prems => [ (cut_facts_tac prems 1), (strip_tac 1), (rtac (thelubI RS sym) 1), - (etac (contXE RS spec RS mp) 1), + (etac (contE RS spec RS mp) 1), (atac 1) ]); @@ -172,7 +172,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "ch2ch_MF2L" Cont.thy -"[|monofun(MF2); is_chain(F)|] ==> is_chain(%i. MF2(F(i),x))" +"[|monofun(MF2); is_chain(F)|] ==> is_chain(%i. MF2 (F i) x)" (fn prems => [ (cut_facts_tac prems 1), @@ -182,7 +182,7 @@ qed_goal "ch2ch_MF2R" Cont.thy -"[|monofun(MF2(f)); is_chain(Y)|] ==> is_chain(%i. MF2(f,Y(i)))" +"[|monofun(MF2(f)); is_chain(Y)|] ==> is_chain(%i. MF2 f (Y i))" (fn prems => [ (cut_facts_tac prems 1), @@ -210,7 +210,7 @@ "[|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(%j. lub(range(%i. MF2 (F j) (Y i))))" (fn prems => [ (cut_facts_tac prems 1), @@ -230,7 +230,7 @@ "[|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(%i. lub(range(%j. MF2 (F j) (Y i))))" (fn prems => [ (cut_facts_tac prems 1), @@ -250,7 +250,7 @@ "[|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))))" +\ monofun(% x.lub(range(% j.MF2 (F j) (x))))" (fn prems => [ (cut_facts_tac prems 1), @@ -270,8 +270,8 @@ "[|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))))))" +\ 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), @@ -374,24 +374,24 @@ (* ------------------------------------------------------------------------ *) qed_goal "contlub_CF2" Cont.thy -"[|contX(CF2);!f.contX(CF2(f));is_chain(FY);is_chain(TY)|] ==>\ +"[|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 contX2contlub RS contlubE RS spec RS mp RS ssubst) 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 contX2mono RS ch2ch_monofun) 1), + (rtac ((hd prems) RS cont2mono RS ch2ch_monofun) 1), (atac 1), (rtac trans 1), - (rtac (((hd (tl prems)) RS spec RS contX2contlub) RS contlubE RS spec RS mp RS ext RS arg_cong RS arg_cong) 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 contX2mono 1), + (etac cont2mono 1), (rtac allI 1), (etac allE 1), - (etac contX2mono 1), + (etac cont2mono 1), (REPEAT (atac 1)) ]); @@ -434,7 +434,7 @@ (* ------------------------------------------------------------------------ *) 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), @@ -444,29 +444,29 @@ (atac 1) ]); -qed_goal "contX2contX_CF1L" Cont.thy - "[|contX(c1)|] ==> contX(%x. c1(x,y))" +qed_goal "cont2cont_CF1L" Cont.thy + "[|cont(c1)|] ==> cont(%x. c1 x y)" (fn prems => [ (cut_facts_tac prems 1), - (rtac monocontlub2contX 1), - (etac (contX2mono RS mono2mono_MF1L) 1), + (rtac monocontlub2cont 1), + (etac (cont2mono RS mono2mono_MF1L) 1), (rtac contlubI 1), (strip_tac 1), - (rtac ((hd prems) RS contX2contlub RS + (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 contX2mono 1), + (etac cont2mono 1), (atac 1), (rtac refl 1) ]); -(********* Note "(%x.%y.c1(x,y)) = c1" ***********) +(********* 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), @@ -478,23 +478,23 @@ (atac 1) ]); -qed_goal "contX2contX_CF1L_rev" Cont.thy - "!y.contX(%x.c1(x,y)) ==> contX(c1)" +qed_goal "cont2cont_CF1L_rev" Cont.thy + "!y.cont(%x.c1 x y) ==> cont(c1)" (fn prems => [ (cut_facts_tac prems 1), - (rtac monocontlub2contX 1), - (rtac (contX2mono RS allI RS mono2mono_MF1L_rev ) 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 (contX2mono RS allI RS mono2mono_MF1L_rev RS ch2ch_monofun) 1), + (rtac (cont2mono RS allI RS mono2mono_MF1L_rev RS ch2ch_monofun) 1), (etac spec 1), (atac 1), (rtac - ((hd prems) RS spec RS contX2contlub RS contlubE RS spec RS mp) 1), + ((hd prems) RS spec RS cont2contlub RS contlubE RS spec RS mp) 1), (atac 1) ]); @@ -505,17 +505,17 @@ (* ------------------------------------------------------------------------ *) qed_goal "contlub_abstraction" Cont.thy -"[|is_chain(Y::nat=>'a);!y.contX(%x.(c::'a=>'b=>'c)(x,y))|] ==>\ -\ (%y.lub(range(%i.c(Y(i),y)))) = (lub(range(%i.%y.c(Y(i),y))))" +"[|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 (contX2contlub RS contlubE RS spec RS mp) 2), + (rtac (cont2contlub RS contlubE RS spec RS mp) 2), (atac 3), - (etac contX2contX_CF1L_rev 2), + (etac cont2cont_CF1L_rev 2), (rtac ext 1), - (rtac (contX2contlub RS contlubE RS spec RS mp RS sym) 1), + (rtac (cont2contlub RS contlubE RS spec RS mp RS sym) 1), (etac spec 1), (atac 1) ]); @@ -539,58 +539,58 @@ ]); -qed_goal "contX2contlub_app" Cont.thy -"[|contX(ft);!x.contX(ft(x));contX(tt)|] ==> contlub(%x.(ft(x))(tt(x)))" +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 contX2contlub 1), + (etac cont2contlub 1), (atac 1), (rtac contlub_CF2 1), (REPEAT (atac 1)), - (etac (contX2mono RS ch2ch_monofun) 1), + (etac (cont2mono RS ch2ch_monofun) 1), (atac 1) ]); -qed_goal "contX2contX_app" Cont.thy -"[|contX(ft);!x.contX(ft(x));contX(tt)|] ==>\ -\ contX(%x.(ft(x))(tt(x)))" +qed_goal "cont2cont_app" Cont.thy +"[|cont(ft);!x.cont(ft(x));cont(tt)|] ==>\ +\ cont(%x.(ft(x))(tt(x)))" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac mono2mono_app 1), - (rtac contX2mono 1), + (rtac cont2mono 1), (resolve_tac prems 1), (strip_tac 1), - (rtac contX2mono 1), + (rtac cont2mono 1), (cut_facts_tac prems 1), (etac spec 1), - (rtac contX2mono 1), + (rtac cont2mono 1), (resolve_tac prems 1), - (rtac contX2contlub_app 1), + (rtac cont2contlub_app 1), (resolve_tac prems 1), (resolve_tac prems 1), (resolve_tac prems 1) ]); -val contX2contX_app2 = (allI RSN (2,contX2contX_app)); -(* [| contX(?ft); !!x. contX(?ft(x)); contX(?tt) |] ==> *) -(* contX(%x. ?ft(x,?tt(x))) *) +val cont2cont_app2 = (allI RSN (2,cont2cont_app)); +(* [| cont ?ft; !!x. cont (?ft x); cont ?tt |] ==> *) +(* cont (%x. ?ft x (?tt x)) *) (* ------------------------------------------------------------------------ *) (* The identity function is continuous *) (* ------------------------------------------------------------------------ *) -qed_goal "contX_id" Cont.thy "contX(% x.x)" +qed_goal "cont_id" Cont.thy "cont(% x.x)" (fn prems => [ - (rtac contXI 1), + (rtac contI 1), (strip_tac 1), (etac thelubE 1), (rtac refl 1) @@ -602,7 +602,7 @@ (* constant functions are continuous *) (* ------------------------------------------------------------------------ *) -qed_goalw "contX_const" Cont.thy [contX] "contX(%x.c)" +qed_goalw "cont_const" Cont.thy [cont] "cont(%x.c)" (fn prems => [ (strip_tac 1), @@ -617,13 +617,13 @@ ]); -qed_goal "contX2contX_app3" Cont.thy - "[|contX(f);contX(t) |] ==> contX(%x. f(t(x)))" +qed_goal "cont2cont_app3" Cont.thy + "[|cont(f);cont(t) |] ==> cont(%x. f(t(x)))" (fn prems => [ (cut_facts_tac prems 1), - (rtac contX2contX_app2 1), - (rtac contX_const 1), + (rtac cont2cont_app2 1), + (rtac cont_const 1), (atac 1), (atac 1) ]); diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cont.thy --- a/src/HOLCF/Cont.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cont.thy Thu Jun 29 16:28:40 1995 +0200 @@ -21,21 +21,21 @@ consts monofun :: "('a::po => 'b::po) => bool" (* monotonicity *) contlub :: "('a => 'b) => bool" (* first cont. def *) - contX :: "('a => 'b) => bool" (* secnd cont. def *) + cont :: "('a => 'b) => bool" (* secnd cont. def *) -rules +defs monofun "monofun(f) == ! x y. x << y --> f(x) << f(y)" -contlub "contlub(f) == ! Y. is_chain(Y) --> - f(lub(range(Y))) = lub(range(% i.f(Y(i))))" +contlub "contlub(f) == ! Y. is_chain(Y) --> \ +\ f(lub(range(Y))) = lub(range(% i.f(Y(i))))" -contX "contX(f) == ! Y. is_chain(Y) --> - range(% i.f(Y(i))) <<| f(lub(range(Y)))" +cont "cont(f) == ! Y. is_chain(Y) --> \ +\ range(% i.f(Y(i))) <<| f(lub(range(Y)))" (* ------------------------------------------------------------------------ *) (* the main purpose of cont.thy is to show: *) -(* monofun(f) & contlub(f) <==> contX(f) *) +(* monofun(f) & contlub(f) <==> cont(f) *) (* ------------------------------------------------------------------------ *) end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cprod1.ML --- a/src/HOLCF/Cprod1.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cprod1.ML Thu Jun 29 16:28:40 1995 +0200 @@ -9,14 +9,14 @@ open Cprod1; qed_goalw "less_cprod1b" Cprod1.thy [less_cprod_def] - "less_cprod(p1,p2) = ( fst(p1) << fst(p2) & snd(p1) << snd(p2))" + "less_cprod p1 p2 = ( fst(p1) << fst(p2) & snd(p1) << snd(p2))" (fn prems => [ (rtac refl 1) ]); qed_goalw "less_cprod2a" Cprod1.thy [less_cprod_def] - "less_cprod(,) ==> x = UU & y = UU" + "less_cprod (x,y) (UU,UU) ==> x = UU & y = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -33,7 +33,7 @@ ]); qed_goal "less_cprod2b" Cprod1.thy - "less_cprod(p,) ==> p=" + "less_cprod p (UU,UU) ==> p = (UU,UU)" (fn prems => [ (cut_facts_tac prems 1), @@ -44,7 +44,7 @@ ]); qed_goalw "less_cprod2c" Cprod1.thy [less_cprod_def] - "less_cprod(,) ==> x1 << x2 & y1 << y2" + "less_cprod (x1,y1) (x2,y2) ==> x1 << x2 & y1 << y2" (fn prems => [ (cut_facts_tac prems 1), @@ -64,7 +64,7 @@ (* less_cprod is a partial order on 'a * 'b *) (* ------------------------------------------------------------------------ *) -qed_goalw "refl_less_cprod" Cprod1.thy [less_cprod_def] "less_cprod(p,p)" +qed_goalw "refl_less_cprod" Cprod1.thy [less_cprod_def] "less_cprod p p" (fn prems => [ (res_inst_tac [("p","p")] PairE 1), @@ -74,7 +74,7 @@ ]); qed_goal "antisym_less_cprod" Cprod1.thy - "[|less_cprod(p1,p2);less_cprod(p2,p1)|] ==> p1=p2" + "[|less_cprod p1 p2;less_cprod p2 p1|] ==> p1=p2" (fn prems => [ (cut_facts_tac prems 1), @@ -92,7 +92,7 @@ qed_goal "trans_less_cprod" Cprod1.thy - "[|less_cprod(p1,p2);less_cprod(p2,p3)|] ==> less_cprod(p1,p3)" + "[|less_cprod p1 p2;less_cprod p2 p3|] ==> less_cprod p1 p3" (fn prems => [ (cut_facts_tac prems 1), @@ -115,3 +115,5 @@ (atac 1) ]); + + diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cprod1.thy --- a/src/HOLCF/Cprod1.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cprod1.thy Thu Jun 29 16:28:40 1995 +0200 @@ -14,9 +14,9 @@ consts less_cprod :: "[('a::pcpo * 'b::pcpo),('a * 'b)] => bool" -rules +defs - less_cprod_def "less_cprod(p1,p2) == ( fst(p1) << fst(p2) & + less_cprod_def "less_cprod p1 p2 == ( fst(p1) << fst(p2) & snd(p1) << snd(p2))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cprod2.ML --- a/src/HOLCF/Cprod2.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cprod2.ML Thu Jun 29 16:28:40 1995 +0200 @@ -9,7 +9,7 @@ open Cprod2; qed_goal "less_cprod3a" Cprod2.thy - "p1= ==> p1 << p2" + "p1=(UU,UU) ==> p1 << p2" (fn prems => [ (cut_facts_tac prems 1), @@ -31,7 +31,7 @@ ]); qed_goal "less_cprod4a" Cprod2.thy - " << ==> x1=UU & x2=UU" + "(x1,x2) << (UU,UU) ==> x1=UU & x2=UU" (fn prems => [ (cut_facts_tac prems 1), @@ -40,7 +40,7 @@ ]); qed_goal "less_cprod4b" Cprod2.thy - "p << ==> p = " + "p << (UU,UU) ==> p = (UU,UU)" (fn prems => [ (cut_facts_tac prems 1), @@ -49,7 +49,7 @@ ]); qed_goal "less_cprod4c" Cprod2.thy - " << ==> xa< xa< [ (cut_facts_tac prems 1), @@ -62,7 +62,7 @@ (* type cprod is pointed *) (* ------------------------------------------------------------------------ *) -qed_goal "minimal_cprod" Cprod2.thy "< [ (rtac less_cprod3a 1), @@ -94,7 +94,7 @@ ]); qed_goal "monofun_pair" Cprod2.thy - "[|x1< << " + "[|x1< (x1,y1) << (x2,y2)" (fn prems => [ (cut_facts_tac prems 1), @@ -140,7 +140,7 @@ qed_goal "lub_cprod" Cprod2.thy " is_chain(S) ==> range(S) <<| \ -\ < lub(range(%i.fst(S(i)))),lub(range(%i.snd(S(i))))> " +\ (lub(range(%i.fst(S i))),lub(range(%i.snd(S i)))) " (fn prems => [ (cut_facts_tac prems 1), @@ -166,9 +166,12 @@ ]); val thelub_cprod = (lub_cprod RS thelubI); -(* "is_chain(?S1) ==> lub(range(?S1)) = *) -(* " *) +(* +"is_chain ?S1 ==> + lub (range ?S1) = + (lub (range (%i. fst (?S1 i))), lub (range (%i. snd (?S1 i))))" : thm +*) qed_goal "cpo_cprod" Cprod2.thy "is_chain(S::nat=>'a*'b)==>? x.range(S)<<| x" @@ -179,3 +182,4 @@ (etac lub_cprod 1) ]); + diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cprod3.ML --- a/src/HOLCF/Cprod3.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cprod3.ML Thu Jun 29 16:28:40 1995 +0200 @@ -9,13 +9,13 @@ open Cprod3; (* ------------------------------------------------------------------------ *) -(* continuity of <_,_> , fst, snd *) +(* continuity of (_,_) , fst, snd *) (* ------------------------------------------------------------------------ *) qed_goal "Cprod3_lemma1" Cprod3.thy "is_chain(Y::(nat=>'a)) ==>\ -\ =\ -\ ))),lub(range(%i. snd()))>" +\ (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), @@ -51,8 +51,8 @@ qed_goal "Cprod3_lemma2" Cprod3.thy "is_chain(Y::(nat=>'a)) ==>\ -\ <(x::'b),lub(range(Y))> =\ -\ ))),lub(range(%i. snd()))>" +\ ((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), @@ -80,18 +80,18 @@ (etac Cprod3_lemma2 1) ]); -qed_goal "contX_pair1" Cprod3.thy "contX(Pair)" +qed_goal "cont_pair1" Cprod3.thy "cont(Pair)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_pair1 1), (rtac contlub_pair1 1) ]); -qed_goal "contX_pair2" Cprod3.thy "contX(Pair(x))" +qed_goal "cont_pair2" Cprod3.thy "cont(Pair(x))" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_pair2 1), (rtac contlub_pair2 1) ]); @@ -116,18 +116,18 @@ (simp_tac prod_ss 1) ]); -qed_goal "contX_fst" Cprod3.thy "contX(fst)" +qed_goal "cont_fst" Cprod3.thy "cont(fst)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_fst 1), (rtac contlub_fst 1) ]); -qed_goal "contX_snd" Cprod3.thy "contX(snd)" +qed_goal "cont_snd" Cprod3.thy "cont(snd)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_snd 1), (rtac contlub_snd 1) ]); @@ -144,21 +144,21 @@ (* ------------------------------------------------------------------------ *) qed_goalw "beta_cfun_cprod" Cprod3.thy [cpair_def] - "(LAM x y.)[a][b] = " + "(LAM x y.(x,y))`a`b = (a,b)" (fn prems => [ (rtac (beta_cfun RS ssubst) 1), - (contX_tac 1), - (rtac contX_pair2 1), - (rtac contX2contX_CF1L 1), - (rtac contX_pair1 1), + (cont_tac 1), + (rtac cont_pair2 1), + (rtac cont2cont_CF1L 1), + (rtac cont_pair1 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_pair2 1), + (rtac cont_pair2 1), (rtac refl 1) ]); qed_goalw "inject_cpair" Cprod3.thy [cpair_def] - " (a#b)=(aa#ba) ==> a=aa & b=ba" + " = ==> a=aa & b=ba" (fn prems => [ (cut_facts_tac prems 1), @@ -168,7 +168,7 @@ (fast_tac HOL_cs 1) ]); -qed_goalw "inst_cprod_pcpo2" Cprod3.thy [cpair_def] "UU = (UU#UU)" +qed_goalw "inst_cprod_pcpo2" Cprod3.thy [cpair_def] "UU = " (fn prems => [ (rtac sym 1), @@ -179,7 +179,7 @@ ]); qed_goal "defined_cpair_rev" Cprod3.thy - "(a#b) = UU ==> a = UU & b = UU" + " = UU ==> a = UU & b = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -188,7 +188,7 @@ ]); qed_goalw "Exh_Cprod2" Cprod3.thy [cpair_def] - "? a b. z=(a#b) " + "? a b. z=" (fn prems => [ (rtac PairE 1), @@ -198,7 +198,7 @@ ]); qed_goalw "cprodE" Cprod3.thy [cpair_def] -"[|!!x y. [|p=(x#y) |] ==> Q|] ==> Q" +"[|!!x y. [|p= |] ==> Q|] ==> Q" (fn prems => [ (rtac PairE 1), @@ -207,57 +207,57 @@ ]); qed_goalw "cfst2" Cprod3.thy [cfst_def,cpair_def] - "cfst[x#y]=x" + "cfst`=x" (fn prems => [ (cut_facts_tac prems 1), (rtac (beta_cfun_cprod RS ssubst) 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_fst 1), + (rtac cont_fst 1), (simp_tac prod_ss 1) ]); qed_goalw "csnd2" Cprod3.thy [csnd_def,cpair_def] - "csnd[x#y]=y" + "csnd`=y" (fn prems => [ (cut_facts_tac prems 1), (rtac (beta_cfun_cprod RS ssubst) 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_snd 1), + (rtac cont_snd 1), (simp_tac prod_ss 1) ]); qed_goalw "surjective_pairing_Cprod2" Cprod3.thy - [cfst_def,csnd_def,cpair_def] "(cfst[p] # csnd[p]) = p" + [cfst_def,csnd_def,cpair_def] " = p" (fn prems => [ (rtac (beta_cfun_cprod RS ssubst) 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_snd 1), + (rtac cont_snd 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_fst 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]< [ (rtac (beta_cfun RS ssubst) 1), - (rtac contX_snd 1), + (rtac cont_snd 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_snd 1), + (rtac cont_snd 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_fst 1), + (rtac cont_fst 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_fst 1), + (rtac cont_fst 1), (rtac less_cprod3b 1) ]); qed_goalw "less_cprod5c" Cprod3.thy [cfst_def,csnd_def,cpair_def] - "xa#ya << x#y ==>xa< << ==> xa< [ (cut_facts_tac prems 1), @@ -270,39 +270,41 @@ 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)])))" +\ <(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 contX_snd 1), + (rtac cont_snd 1), (rtac (beta_cfun RS ext RS ssubst) 1), - (rtac contX_fst 1), + (rtac cont_fst 1), (rtac lub_cprod 1), (atac 1) ]); val thelub_cprod2 = (lub_cprod2 RS thelubI); -(* "is_chain(?S1) ==> lub(range(?S1)) = *) -(* lub(range(%i. cfst[?S1(i)]))#lub(range(%i. csnd[?S1(i)]))" *) - +(* +is_chain ?S1 ==> + lub (range ?S1) = + " +*) qed_goalw "csplit2" Cprod3.thy [csplit_def] - "csplit[f][x#y]=f[x][y]" + "csplit`f` = f`x`y" (fn prems => [ (rtac (beta_cfun RS ssubst) 1), - (contX_tacR 1), + (cont_tacR 1), (simp_tac Cfun_ss 1), (simp_tac (Cfun_ss addsimps [cfst2,csnd2]) 1) ]); qed_goalw "csplit3" Cprod3.thy [csplit_def] - "csplit[cpair][z]=z" + "csplit`cpair`z=z" (fn prems => [ (rtac (beta_cfun RS ssubst) 1), - (contX_tacR 1), + (cont_tacR 1), (simp_tac (Cfun_ss addsimps [surjective_pairing_Cprod2]) 1) ]); @@ -313,3 +315,4 @@ val Cprod_rews = [cfst2,csnd2,csplit2]; val Cprod_ss = Cfun_ss addsimps Cprod_rews; + diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Cprod3.thy --- a/src/HOLCF/Cprod3.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Cprod3.thy Thu Jun 29 16:28:40 1995 +0200 @@ -18,18 +18,23 @@ csnd :: "('a*'b)->'b" csplit :: "('a->'b->'c)->('a*'b)->'c" -syntax "@cpair" :: "'a => 'b => ('a*'b)" ("_#_" [101,100] 100) +syntax + "@ctuple" :: "['a, args] => 'a * 'b" ("(1<_,/ _>)") + -translations "x#y" == "cpair[x][y]" +translations + "" == ">" + "" == "cpair`x`y" rules -inst_cprod_pcpo "(UU::'a*'b) = " +inst_cprod_pcpo "(UU::'a*'b) = (UU,UU)" -cpair_def "cpair == (LAM x y.)" +defs +cpair_def "cpair == (LAM x y.(x,y))" cfst_def "cfst == (LAM p.fst(p))" csnd_def "csnd == (LAM p.snd(p))" -csplit_def "csplit == (LAM f p.f[cfst[p]][csnd[p]])" +csplit_def "csplit == (LAM f p.f`(cfst`p)`(csnd`p))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Dlist.ML --- a/src/HOLCF/Dlist.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Dlist.ML Thu Jun 29 16:28:40 1995 +0200 @@ -22,7 +22,7 @@ (* Properties of dlist_copy *) (* ------------------------------------------------------------------------*) -val temp = prove_goalw Dlist.thy [dlist_copy_def] "dlist_copy[f][UU]=UU" +val temp = prove_goalw Dlist.thy [dlist_copy_def] "dlist_copy`f`UU=UU" (fn prems => [ (asm_simp_tac (HOLCF_ss addsimps @@ -33,7 +33,7 @@ val temp = prove_goalw Dlist.thy [dlist_copy_def,dnil_def] - "dlist_copy[f][dnil]=dnil" + "dlist_copy`f`dnil=dnil" (fn prems => [ (asm_simp_tac (HOLCF_ss addsimps @@ -44,7 +44,7 @@ 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), @@ -64,12 +64,12 @@ (* ------------------------------------------------------------------------*) 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 HOLCF_ss 1), (rtac (dlist_rep_iso RS subst) 1), - (res_inst_tac [("p","dlist_rep[l]")] ssumE 1), + (res_inst_tac [("p","dlist_rep`l")] ssumE 1), (rtac disjI1 1), (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1), (rtac disjI2 1), @@ -89,7 +89,7 @@ qed_goal "dlistE" Dlist.thy -"[| l=UU ==> Q; l=dnil ==> Q;!!x xl.[|l=dcons[x][xl];x~=UU;xl~=UU|]==>Q|]==>Q" +"[| 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), @@ -108,7 +108,7 @@ (* Properties of dlist_when *) (* ------------------------------------------------------------------------*) -val temp = prove_goalw Dlist.thy [dlist_when_def] "dlist_when[f1][f2][UU]=UU" +val temp = prove_goalw Dlist.thy [dlist_when_def] "dlist_when`f1`f2`UU=UU" (fn prems => [ (asm_simp_tac (HOLCF_ss addsimps [dlist_iso_strict]) 1) @@ -117,7 +117,7 @@ val dlist_when = [temp]; val temp = prove_goalw Dlist.thy [dlist_when_def,dnil_def] - "dlist_when[f1][f2][dnil]= f1" + "dlist_when`f1`f2`dnil= f1" (fn prems => [ (asm_simp_tac (HOLCF_ss addsimps [dlist_abs_iso]) 1) @@ -126,7 +126,7 @@ 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]" + "[|x~=UU;xl~=UU|] ==> dlist_when`f1`f2`(dcons`x`xl)= f2`x`xl" (fn prems => [ (cut_facts_tac prems 1), @@ -148,10 +148,10 @@ ]); 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 @@ -163,21 +163,21 @@ val dlist_discsel = [ prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "is_dnil[dnil]=TT", + "is_dnil`dnil=TT", prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "[|x~=UU;xl~=UU|] ==> is_dnil[dcons[x][xl]]=FF", + "[|x~=UU;xl~=UU|] ==> is_dnil`(dcons`x`xl)=FF", prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "is_dcons[dnil]=FF", + "is_dcons`dnil=FF", prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "[|x~=UU;xl~=UU|] ==> is_dcons[dcons[x][xl]]=TT", + "[|x~=UU;xl~=UU|] ==> is_dcons`(dcons`x`xl)=TT", prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "dhd[dnil]=UU", + "dhd`dnil=UU", prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "[|x~=UU;xl~=UU|] ==> dhd[dcons[x][xl]]=x", + "[|x~=UU;xl~=UU|] ==> dhd`(dcons`x`xl)=x", prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "dtl[dnil]=UU", + "dtl`dnil=UU", prover [is_dnil_def,is_dcons_def,dhd_def,dtl_def] - "[|x~=UU;xl~=UU|] ==> dtl[dcons[x][xl]]=xl"] @ dlist_discsel; + "[|x~=UU;xl~=UU|] ==> dtl`(dcons`x`xl)=xl"] @ dlist_discsel; val dlist_rews = dlist_discsel @ dlist_rews; @@ -197,8 +197,9 @@ 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" +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" ]; @@ -209,8 +210,8 @@ ]); val dlist_constrdef = [ - prover [dcons_def] "dcons[UU][xl]=UU", - prover [dcons_def] "dcons[x][UU]=UU" + prover [dcons_def] "dcons`UU`xl=UU", + prover [dcons_def] "dcons`x`UU=UU" ] @ dlist_constrdef; val dlist_rews = dlist_constrdef @ dlist_rews; @@ -220,7 +221,7 @@ (* Distinctness wrt. << and = *) (* ------------------------------------------------------------------------*) -val temp = prove_goal Dlist.thy "~dnil << dcons[x::'a][xl]" +val temp = prove_goal Dlist.thy "~dnil << dcons`(x::'a)`xl" (fn prems => [ (res_inst_tac [("P1","TT << FF")] classical3 1), @@ -237,7 +238,7 @@ val dlist_dist_less = [temp]; -val temp = prove_goal Dlist.thy "[|x~=UU;xl~=UU|]==>~dcons[x][xl] << dnil" +val temp = prove_goal Dlist.thy "[|x~=UU;xl~=UU|]==>~ dcons`x`xl << dnil" (fn prems => [ (cut_facts_tac prems 1), @@ -251,7 +252,7 @@ val dlist_dist_less = temp::dlist_dist_less; -val temp = prove_goal Dlist.thy "dnil ~= dcons[x][xl]" +val temp = prove_goal Dlist.thy "dnil ~= dcons`x`xl" (fn prems => [ (res_inst_tac [("Q","x=UU")] classical2 1), @@ -275,16 +276,16 @@ (* ------------------------------------------------------------------------*) val temp = prove_goal Dlist.thy "[|x1~=UU; y1~=UU;x2~=UU; y2~=UU;\ -\ dcons[x1][x2] << dcons[y1][y2]|] ==> x1<< y1 & x2 << y2" +\ 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), + (dres_inst_tac [("fo5","dlist_when`UU`(LAM x l.x)")] monofun_cfun_arg 1), (etac box_less 1), (asm_simp_tac (HOLCF_ss addsimps dlist_when) 1), (asm_simp_tac (HOLCF_ss addsimps dlist_when) 1), - (dres_inst_tac [("fo5","dlist_when[UU][LAM x l.l]")] monofun_cfun_arg 1), + (dres_inst_tac [("fo5","dlist_when`UU`(LAM x l.l)")] monofun_cfun_arg 1), (etac box_less 1), (asm_simp_tac (HOLCF_ss addsimps dlist_when) 1), (asm_simp_tac (HOLCF_ss addsimps dlist_when) 1) @@ -297,16 +298,16 @@ (* ------------------------------------------------------------------------*) val temp = prove_goal Dlist.thy "[|x1~=UU; y1~=UU;x2~=UU; y2~=UU;\ -\ dcons[x1][x2] = dcons[y1][y2]|] ==> x1= y1 & x2 = y2" +\ 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), + (dres_inst_tac [("f","dlist_when`UU`(LAM x l.x)")] cfun_arg_cong 1), (etac box_equals 1), (asm_simp_tac (HOLCF_ss addsimps dlist_when) 1), (asm_simp_tac (HOLCF_ss addsimps dlist_when) 1), - (dres_inst_tac [("f","dlist_when[UU][LAM x l.l]")] cfun_arg_cong 1), + (dres_inst_tac [("f","dlist_when`UU`(LAM x l.l)")] cfun_arg_cong 1), (etac box_equals 1), (asm_simp_tac (HOLCF_ss addsimps dlist_when) 1), (asm_simp_tac (HOLCF_ss addsimps dlist_when) 1) @@ -330,8 +331,8 @@ 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; @@ -340,7 +341,7 @@ (* enhance the simplifier *) (* ------------------------------------------------------------------------*) -qed_goal "dhd2" Dlist.thy "xl~=UU ==> dhd[dcons[x][xl]]=x" +qed_goal "dhd2" Dlist.thy "xl~=UU ==> dhd`(dcons`x`xl)=x" (fn prems => [ (cut_facts_tac prems 1), @@ -349,7 +350,7 @@ (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1) ]); -qed_goal "dtl2" Dlist.thy "x~=UU ==> dtl[dcons[x][xl]]=xl" +qed_goal "dtl2" Dlist.thy "x~=UU ==> dtl`(dcons`x`xl)=xl" (fn prems => [ (cut_facts_tac prems 1), @@ -364,7 +365,7 @@ (* Properties dlist_take *) (* ------------------------------------------------------------------------*) -val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take(n)[UU]=UU" +val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take n`UU=UU" (fn prems => [ (res_inst_tac [("n","n")] natE 1), @@ -375,7 +376,7 @@ val dlist_take = [temp]; -val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take(0)[xs]=UU" +val temp = prove_goalw Dlist.thy [dlist_take_def] "dlist_take 0`xs=UU" (fn prems => [ (asm_simp_tac iterate_ss 1) @@ -384,7 +385,7 @@ 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 iterate_ss 1), @@ -394,7 +395,7 @@ 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]]" + "dlist_take (Suc n)`(dcons`x`xl)= dcons`x`(dlist_take n`xl)" (fn prems => [ (res_inst_tac [("Q","x=UU")] classical2 1), @@ -439,7 +440,7 @@ ]); 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"; (* ------------------------------------------------------------------------*) @@ -447,7 +448,7 @@ (* ------------------------------------------------------------------------*) 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]" +"dlist_bisim R ==> ! p q. R p q --> dlist_take n`p = dlist_take n`q" (fn prems => [ (cut_facts_tac prems 1), @@ -468,7 +469,7 @@ (fast_tac HOL_cs 1) ]); -qed_goal "dlist_coind" Dlist.thy "[|dlist_bisim(R);R(p,q)|] ==> p = q" +qed_goal "dlist_coind" Dlist.thy "[|dlist_bisim R ; R p q |] ==> p = q" (fn prems => [ (rtac dlist_take_lemma 1), @@ -483,8 +484,8 @@ qed_goal "dlist_finite_ind" Dlist.thy "[|P(UU);P(dnil);\ -\ !! x l1.[|x~=UU;l1~=UU;P(l1)|] ==> P(dcons[x][l1])\ -\ |] ==> !l.P(dlist_take(n)[l])" +\ !! 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), @@ -497,7 +498,7 @@ (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1), (resolve_tac prems 1), (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1), - (res_inst_tac [("Q","dlist_take(n1)[xl]=UU")] classical2 1), + (res_inst_tac [("Q","dlist_take n1`xl=UU")] classical2 1), (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1), (resolve_tac prems 1), (resolve_tac prems 1), @@ -507,7 +508,7 @@ ]); qed_goal "dlist_all_finite_lemma1" Dlist.thy -"!l.dlist_take(n)[l]=UU |dlist_take(n)[l]=l" +"!l.dlist_take n`l=UU |dlist_take n`l=l" (fn prems => [ (nat_ind_tac "n" 1), @@ -523,18 +524,18 @@ (asm_simp_tac (HOLCF_ss addsimps dlist_rews) 1) ]); -qed_goal "dlist_all_finite_lemma2" Dlist.thy "? n.dlist_take(n)[l]=l" +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 (HOLCF_ss addsimps dlist_rews) 1), - (subgoal_tac "(!n.dlist_take(n)[l]=UU) |(? n.dlist_take(n)[l]=l)" 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 (HOLCF_ss addsimps dlist_rews) 1), (atac 1), - (subgoal_tac "!n.!l.dlist_take(n)[l]=UU |dlist_take(n)[l]=l" 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) @@ -548,7 +549,7 @@ qed_goal "dlist_ind" Dlist.thy "[|P(UU);P(dnil);\ -\ !! x l1.[|x~=UU;l1~=UU;P(l1)|] ==> P(dcons[x][l1])|] ==> P(l)" +\ !! x l1.[|x~=UU;l1~=UU;P(l1)|] ==> P(dcons`x`l1)|] ==> P(l)" (fn prems => [ (rtac (dlist_all_finite_lemma2 RS exE) 1), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Dlist.thy --- a/src/HOLCF/Dlist.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Dlist.thy Thu Jun 29 16:28:40 1995 +0200 @@ -66,44 +66,46 @@ (* dlist_abs is an isomorphism with inverse dlist_rep *) (* identity is the least endomorphism on 'a dlist *) -dlist_abs_iso "dlist_rep[dlist_abs[x]] = x" -dlist_rep_iso "dlist_abs[dlist_rep[x]] = x" -dlist_copy_def "dlist_copy == (LAM f. dlist_abs oo - (when[sinl][sinr oo (ssplit[LAM x y. x ## f[y]])]) - oo dlist_rep)" -dlist_reach "(fix[dlist_copy])[x]=x" +dlist_abs_iso "dlist_rep`(dlist_abs`x) = x" +dlist_rep_iso "dlist_abs`(dlist_rep`x) = x" +dlist_copy_def "dlist_copy == (LAM f. dlist_abs oo \ +\ (sswhen`sinl`(sinr oo (ssplit`(LAM x y. (|x,f`y|) ))))\ +\ oo dlist_rep)" +dlist_reach "(fix`dlist_copy)`x=x" + +defs (* ----------------------------------------------------------------------- *) (* properties of additional constants *) (* ----------------------------------------------------------------------- *) (* constructors *) -dnil_def "dnil == dlist_abs[sinl[one]]" -dcons_def "dcons == (LAM x l. dlist_abs[sinr[x##l]])" +dnil_def "dnil == dlist_abs`(sinl`one)" +dcons_def "dcons == (LAM x l. dlist_abs`(sinr`(|x,l|) ))" (* ----------------------------------------------------------------------- *) (* discriminator functional *) dlist_when_def -"dlist_when == (LAM f1 f2 l. - when[LAM x.f1][ssplit[LAM x l.f2[x][l]]][dlist_rep[l]])" +"dlist_when == (LAM f1 f2 l.\ +\ sswhen`(LAM x.f1) `(ssplit`(LAM x l.f2`x`l)) `(dlist_rep`l))" (* ----------------------------------------------------------------------- *) (* discriminators and selectors *) -is_dnil_def "is_dnil == dlist_when[TT][LAM x l.FF]" -is_dcons_def "is_dcons == dlist_when[FF][LAM x l.TT]" -dhd_def "dhd == dlist_when[UU][LAM x l.x]" -dtl_def "dtl == dlist_when[UU][LAM x l.l]" +is_dnil_def "is_dnil == dlist_when`TT`(LAM x l.FF)" +is_dcons_def "is_dcons == dlist_when`FF`(LAM x l.TT)" +dhd_def "dhd == dlist_when`UU`(LAM x l.x)" +dtl_def "dtl == dlist_when`UU`(LAM x l.l)" (* ----------------------------------------------------------------------- *) (* the taker for dlists *) -dlist_take_def "dlist_take == (%n.iterate(n,dlist_copy,UU))" +dlist_take_def "dlist_take == (%n.iterate n dlist_copy UU)" (* ----------------------------------------------------------------------- *) -dlist_finite_def "dlist_finite == (%s.? n.dlist_take(n)[s]=s)" +dlist_finite_def "dlist_finite == (%s.? n.dlist_take n`s=s)" (* ----------------------------------------------------------------------- *) (* definition of bisimulation is determined by domain equation *) @@ -111,11 +113,11 @@ dlist_bisim_def "dlist_bisim == ( %R.!l1 l2. - R(l1,l2) --> + R l1 l2 --> ((l1=UU & l2=UU) | (l1=dnil & l2=dnil) | (? x l11 l21. x~=UU & l11~=UU & l21~=UU & - l1=dcons[x][l11] & l2 = dcons[x][l21] & R(l11,l21))))" + l1=dcons`x`l11 & l2 = dcons`x`l21 & R l11 l21)))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Dnat.ML --- a/src/HOLCF/Dnat.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Dnat.ML Thu Jun 29 16:28:40 1995 +0200 @@ -32,10 +32,10 @@ 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] "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]]" + "n~=UU ==> dnat_copy`f`(dsucc`n) = dsucc`(f`n)" ]; val dnat_rews = dnat_copy @ dnat_rews; @@ -45,12 +45,12 @@ (* ------------------------------------------------------------------------*) 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 HOLCF_ss 1), (rtac (dnat_rep_iso RS subst) 1), - (res_inst_tac [("p","dnat_rep[n]")] ssumE 1), + (res_inst_tac [("p","dnat_rep`n")] ssumE 1), (rtac disjI1 1), (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1), (rtac (disjI1 RS disjI2) 1), @@ -64,7 +64,7 @@ ]); qed_goal "dnatE" Dnat.thy - "[| n=UU ==> Q; n=dzero ==> Q; !!x.[|n=dsucc[x];x~=UU|]==>Q|]==>Q" + "[| n=UU ==> Q; n=dzero ==> Q; !!x.[|n=dsucc`x;x~=UU|]==>Q|]==>Q" (fn prems => [ (rtac (Exh_dnat RS disjE) 1), @@ -91,10 +91,10 @@ 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] "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]" + "n~=UU ==> dnat_when`c`f`(dsucc`n)=f`n" ]; val dnat_rews = dnat_when @ dnat_rews; @@ -110,9 +110,9 @@ ]); 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" ]; @@ -124,12 +124,12 @@ ]); 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" + 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; @@ -149,8 +149,8 @@ ]); 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" ]; @@ -161,7 +161,7 @@ ]); val dnat_constrdef = [ - prover [dsucc_def] "dsucc[UU]=UU" + prover [dsucc_def] "dsucc`UU=UU" ] @ dnat_constrdef; val dnat_rews = dnat_constrdef @ dnat_rews; @@ -171,7 +171,7 @@ (* Distinctness wrt. << and = *) (* ------------------------------------------------------------------------*) -val temp = prove_goal Dnat.thy "~dzero << dsucc[n]" +val temp = prove_goal Dnat.thy "~dzero << dsucc`n" (fn prems => [ (res_inst_tac [("P1","TT << FF")] classical3 1), @@ -186,7 +186,7 @@ val dnat_dist_less = [temp]; -val temp = prove_goal Dnat.thy "n~=UU ==> ~dsucc[n] << dzero" +val temp = prove_goal Dnat.thy "n~=UU ==> ~dsucc`n << dzero" (fn prems => [ (cut_facts_tac prems 1), @@ -200,7 +200,7 @@ val dnat_dist_less = temp::dnat_dist_less; -val temp = prove_goal Dnat.thy "dzero ~= dsucc[n]" +val temp = prove_goal Dnat.thy "dzero ~= dsucc`n" (fn prems => [ (res_inst_tac [("Q","n=UU")] classical2 1), @@ -224,11 +224,11 @@ val dnat_invert = [ prove_goal Dnat.thy -"[|x1~=UU; y1~=UU; dsucc[x1] << dsucc[y1] |] ==> x1<< y1" +"[|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), + (dres_inst_tac [("fo5","dnat_when`c`(LAM x.x)")] monofun_cfun_arg 1), (etac box_less 1), (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1), (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1) @@ -242,11 +242,11 @@ val dnat_inject = [ prove_goal Dnat.thy -"[|x1~=UU; y1~=UU; dsucc[x1] = dsucc[y1] |] ==> x1= y1" +"[|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), + (dres_inst_tac [("f","dnat_when`c`(LAM x.x)")] cfun_arg_cong 1), (etac box_equals 1), (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1), (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1) @@ -269,8 +269,8 @@ 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; @@ -279,7 +279,7 @@ (* ------------------------------------------------------------------------*) (* Properties dnat_take *) (* ------------------------------------------------------------------------*) -val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take(n)[UU]=UU" +val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take n`UU = UU" (fn prems => [ (res_inst_tac [("n","n")] natE 1), @@ -290,7 +290,7 @@ val dnat_take = [temp]; -val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take(0)[xs]=UU" +val temp = prove_goalw Dnat.thy [dnat_take_def] "dnat_take 0`xs = UU" (fn prems => [ (asm_simp_tac iterate_ss 1) @@ -299,7 +299,7 @@ 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 iterate_ss 1), @@ -309,7 +309,7 @@ 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]]" + "dnat_take (Suc n)`(dsucc`xs)=dsucc`(dnat_take n ` xs)" (fn prems => [ (res_inst_tac [("Q","xs=UU")] classical2 1), @@ -352,7 +352,7 @@ ]); 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"; (* ------------------------------------------------------------------------*) @@ -360,7 +360,7 @@ (* ------------------------------------------------------------------------*) 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]" +"dnat_bisim R ==> ! p q. R p q --> dnat_take n`p = dnat_take n`q" (fn prems => [ (cut_facts_tac prems 1), @@ -380,7 +380,7 @@ (fast_tac HOL_cs 1) ]); -qed_goal "dnat_coind" Dnat.thy "[|dnat_bisim(R);R(p,q)|] ==> p = q" +qed_goal "dnat_coind" Dnat.thy "[|dnat_bisim R;R p q|] ==> p = q" (fn prems => [ (rtac dnat_take_lemma 1), @@ -399,7 +399,7 @@ "[| adm(P);\ \ P(UU);\ \ P(dzero);\ -\ !! s1.[|s1~=UU ; P(s1)|] ==> P(dsucc[s1])|] ==> P(s)" +\ !! s1.[|s1~=UU ; P(s1)|] ==> P(dsucc`s1)|] ==> P(s)" (fn prems => [ (rtac (dnat_reach RS subst) 1), @@ -407,7 +407,7 @@ (rtac fix_ind 1), (rtac adm_all2 1), (rtac adm_subst 1), - (contX_tacR 1), + (cont_tacR 1), (resolve_tac prems 1), (simp_tac HOLCF_ss 1), (resolve_tac prems 1), @@ -418,7 +418,7 @@ (asm_simp_tac (HOLCF_ss addsimps dnat_copy) 1), (resolve_tac prems 1), (asm_simp_tac (HOLCF_ss addsimps dnat_copy) 1), - (res_inst_tac [("Q","x[xb]=UU")] classical2 1), + (res_inst_tac [("Q","x`xb=UU")] classical2 1), (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1), (resolve_tac prems 1), (eresolve_tac prems 1), @@ -428,8 +428,8 @@ qed_goal "dnat_finite_ind" Dnat.thy "[|P(UU);P(dzero);\ -\ !! s1.[|s1~=UU;P(s1)|] ==> P(dsucc[s1])\ -\ |] ==> !s.P(dnat_take(n)[s])" +\ !! s1.[|s1~=UU;P(s1)|] ==> P(dsucc`s1)\ +\ |] ==> !s.P(dnat_take n`s)" (fn prems => [ (nat_ind_tac "n" 1), @@ -442,7 +442,7 @@ (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1), (resolve_tac prems 1), (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1), - (res_inst_tac [("Q","dnat_take(n1)[x]=UU")] classical2 1), + (res_inst_tac [("Q","dnat_take n1`x=UU")] classical2 1), (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1), (resolve_tac prems 1), (resolve_tac prems 1), @@ -451,7 +451,7 @@ ]); qed_goal "dnat_all_finite_lemma1" Dnat.thy -"!s.dnat_take(n)[s]=UU |dnat_take(n)[s]=s" +"!s.dnat_take n`s=UU |dnat_take n`s=s" (fn prems => [ (nat_ind_tac "n" 1), @@ -467,18 +467,18 @@ (asm_simp_tac (HOLCF_ss addsimps dnat_rews) 1) ]); -qed_goal "dnat_all_finite_lemma2" Dnat.thy "? n.dnat_take(n)[s]=s" +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 (HOLCF_ss addsimps dnat_rews) 1), - (subgoal_tac "(!n.dnat_take(n)[s]=UU) |(? n.dnat_take(n)[s]=s)" 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 (HOLCF_ss addsimps dnat_rews) 1), (atac 1), - (subgoal_tac "!n.!s.dnat_take(n)[s]=UU |dnat_take(n)[s]=s" 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) @@ -487,7 +487,7 @@ qed_goal "dnat_ind" Dnat.thy "[|P(UU);P(dzero);\ -\ !! s1.[|s1~=UU;P(s1)|] ==> P(dsucc[s1])\ +\ !! s1.[|s1~=UU;P(s1)|] ==> P(dsucc`s1)\ \ |] ==> P(s)" (fn prems => [ diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Dnat.thy --- a/src/HOLCF/Dnat.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Dnat.thy Thu Jun 29 16:28:40 1995 +0200 @@ -61,38 +61,40 @@ (* dnat_abs is an isomorphism with inverse dnat_rep *) (* identity is the least endomorphism on dnat *) -dnat_abs_iso "dnat_rep[dnat_abs[x]] = x" -dnat_rep_iso "dnat_abs[dnat_rep[x]] = x" +dnat_abs_iso "dnat_rep`(dnat_abs`x) = x" +dnat_rep_iso "dnat_abs`(dnat_rep`x) = x" dnat_copy_def "dnat_copy == (LAM f. dnat_abs oo - (when[sinl][sinr oo f]) oo dnat_rep )" -dnat_reach "(fix[dnat_copy])[x]=x" + (sswhen`sinl`(sinr oo f)) oo dnat_rep )" +dnat_reach "(fix`dnat_copy)`x=x" + +defs (* ----------------------------------------------------------------------- *) (* properties of additional constants *) (* ----------------------------------------------------------------------- *) (* constructors *) -dzero_def "dzero == dnat_abs[sinl[one]]" -dsucc_def "dsucc == (LAM n. dnat_abs[sinr[n]])" +dzero_def "dzero == dnat_abs`(sinl`one)" +dsucc_def "dsucc == (LAM n. dnat_abs`(sinr`n))" (* ----------------------------------------------------------------------- *) (* discriminator functional *) -dnat_when_def "dnat_when == (LAM f1 f2 n.when[LAM x.f1][f2][dnat_rep[n]])" +dnat_when_def "dnat_when == (LAM f1 f2 n.sswhen`(LAM x.f1)`f2`(dnat_rep`n))" (* ----------------------------------------------------------------------- *) (* discriminators and selectors *) -is_dzero_def "is_dzero == dnat_when[TT][LAM x.FF]" -is_dsucc_def "is_dsucc == dnat_when[FF][LAM x.TT]" -dpred_def "dpred == dnat_when[UU][LAM x.x]" +is_dzero_def "is_dzero == dnat_when`TT`(LAM x.FF)" +is_dsucc_def "is_dsucc == dnat_when`FF`(LAM x.TT)" +dpred_def "dpred == dnat_when`UU`(LAM x.x)" (* ----------------------------------------------------------------------- *) (* the taker for dnats *) -dnat_take_def "dnat_take == (%n.iterate(n,dnat_copy,UU))" +dnat_take_def "dnat_take == (%n.iterate n dnat_copy UU)" (* ----------------------------------------------------------------------- *) (* definition of bisimulation is determined by domain equation *) @@ -100,13 +102,9 @@ dnat_bisim_def "dnat_bisim == (%R.!s1 s2. - R(s1,s2) --> + R s1 s2 --> ((s1=UU & s2=UU) |(s1=dzero & s2=dzero) | - (? s11 s21. s11~=UU & s21~=UU & s1=dsucc[s11] & - s2 = dsucc[s21] & R(s11,s21))))" + (? s11 s21. s11~=UU & s21~=UU & s1=dsucc`s11 & + s2 = dsucc`s21 & R s11 s21)))" end - - - - diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Dnat2.ML --- a/src/HOLCF/Dnat2.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Dnat2.ML Thu Jun 29 16:28:40 1995 +0200 @@ -13,21 +13,21 @@ (* expand fixed point properties *) (* ------------------------------------------------------------------------- *) -val iterator_def2 = fix_prover Dnat2.thy iterator_def - "iterator = (LAM n f x. dnat_when[x][LAM m.f[iterator[m][f][x]]][n])"; +val iterator_def2 = fix_prover2 Dnat2.thy iterator_def + "iterator = (LAM n f x. dnat_when`x`(LAM m.f`(iterator`m`f`x)) `n)"; (* ------------------------------------------------------------------------- *) (* recursive properties *) (* ------------------------------------------------------------------------- *) -qed_goal "iterator1" Dnat2.thy "iterator[UU][f][x] = UU" +qed_goal "iterator1" Dnat2.thy "iterator`UU`f`x = UU" (fn prems => [ (rtac (iterator_def2 RS ssubst) 1), (simp_tac (HOLCF_ss addsimps dnat_when) 1) ]); -qed_goal "iterator2" Dnat2.thy "iterator[dzero][f][x] = x" +qed_goal "iterator2" Dnat2.thy "iterator`dzero`f`x = x" (fn prems => [ (rtac (iterator_def2 RS ssubst) 1), @@ -35,7 +35,7 @@ ]); qed_goal "iterator3" Dnat2.thy -"n~=UU ==> iterator[dsucc[n]][f][x] = f[iterator[n][f][x]]" +"n~=UU ==> iterator`(dsucc`n)`f`x = f`(iterator`n`f`x)" (fn prems => [ (cut_facts_tac prems 1), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Dnat2.thy --- a/src/HOLCF/Dnat2.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Dnat2.thy Thu Jun 29 16:28:40 1995 +0200 @@ -14,19 +14,16 @@ iterator :: "dnat -> ('a -> 'a) -> 'a -> 'a" -rules +defs -iterator_def "iterator = fix[LAM h n f x. - dnat_when[x][LAM m.f[h[m][f][x]]][n]]" - - +iterator_def "iterator == fix`(LAM h n f x. + dnat_when `x `(LAM m.f`(h`m`f`x)) `n)" end - (* - iterator[UU][f][x] = UU - iterator[dzero][f][x] = x - n~=UU --> iterator[dsucc[n]][f][x] = f[iterator[n][f][x]] + iterator`UU`f`x = UU + iterator`dzero`f`x = x + n~=UU --> iterator`(dsucc`n)`f`x = f`(iterator`n`f`x) *) diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Fix.ML --- a/src/HOLCF/Fix.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Fix.ML Thu Jun 29 16:28:40 1995 +0200 @@ -12,13 +12,13 @@ (* derive inductive properties of iterate from primitive recursion *) (* ------------------------------------------------------------------------ *) -qed_goal "iterate_0" Fix.thy "iterate(0,F,x) = x" +qed_goal "iterate_0" Fix.thy "iterate 0 F x = x" (fn prems => [ (resolve_tac (nat_recs iterate_def) 1) ]); -qed_goal "iterate_Suc" Fix.thy "iterate(Suc(n),F,x) = F[iterate(n,F,x)]" +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) @@ -26,7 +26,7 @@ val iterate_ss = Cfun_ss addsimps [iterate_0,iterate_Suc]; -qed_goal "iterate_Suc2" Fix.thy "iterate(Suc(n),F,x) = iterate(n,F,F[x])" +qed_goal "iterate_Suc2" Fix.thy "iterate (Suc n) F x = iterate n F (F`x)" (fn prems => [ (nat_ind_tac "n" 1), @@ -40,7 +40,7 @@ (* ------------------------------------------------------------------------ *) 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), @@ -54,7 +54,7 @@ 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), @@ -68,7 +68,7 @@ (* ------------------------------------------------------------------------ *) -qed_goalw "Ifix_eq" Fix.thy [Ifix_def] "Ifix(F)=F[Ifix(F)]" +qed_goalw "Ifix_eq" Fix.thy [Ifix_def] "Ifix F =F`(Ifix F)" (fn prems => [ (rtac (contlub_cfun_arg RS ssubst) 1), @@ -92,7 +92,7 @@ ]); -qed_goalw "Ifix_least" Fix.thy [Ifix_def] "F[x]=x ==> Ifix(F) << x" +qed_goalw "Ifix_least" Fix.thy [Ifix_def] "F`x=x ==> Ifix(F) << x" (fn prems => [ (cut_facts_tac prems 1), @@ -165,10 +165,10 @@ ]); -qed_goal "contX_iterate" Fix.thy "contX(iterate(i))" +qed_goal "cont_iterate" Fix.thy "cont(iterate(i))" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_iterate 1), (rtac contlub_iterate 1) ]); @@ -177,7 +177,7 @@ (* a lemma about continuity of iterate in its third argument *) (* ------------------------------------------------------------------------ *) -qed_goal "monofun_iterate2" Fix.thy "monofun(iterate(n,F))" +qed_goal "monofun_iterate2" Fix.thy "monofun(iterate n F)" (fn prems => [ (rtac monofunI 1), @@ -188,7 +188,7 @@ (etac monofun_cfun_arg 1) ]); -qed_goal "contlub_iterate2" Fix.thy "contlub(iterate(n,F))" +qed_goal "contlub_iterate2" Fix.thy "contlub(iterate n F)" (fn prems => [ (rtac contlubI 1), @@ -196,17 +196,17 @@ (nat_ind_tac "n" 1), (simp_tac iterate_ss 1), (simp_tac iterate_ss 1), - (res_inst_tac [("t","iterate(n1, F, lub(range(%u. Y(u))))"), - ("s","lub(range(%i. iterate(n1, F, Y(i))))")] ssubst 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 "contX_iterate2" Fix.thy "contX(iterate(n,F))" +qed_goal "cont_iterate2" Fix.thy "cont (iterate n F)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_iterate2 1), (rtac contlub_iterate2 1) ]); @@ -234,7 +234,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "is_chain_iterate_lub" Fix.thy -"is_chain(Y) ==> is_chain(%i. lub(range(%ia. iterate(ia,Y(i),UU))))" +"is_chain(Y) ==> is_chain(%i. lub(range(%ia. iterate ia (Y i) UU)))" (fn prems => [ (cut_facts_tac prems 1), @@ -255,7 +255,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "contlub_Ifix_lemma1" Fix.thy -"is_chain(Y) ==> iterate(n,lub(range(Y)),y) = lub(range(%i. iterate(n,Y(i),y)))" +"is_chain(Y) ==>iterate n (lub(range Y)) y = lub(range(%i. iterate n (Y i) y))" (fn prems => [ (cut_facts_tac prems 1), @@ -270,8 +270,8 @@ 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)))))" +\ 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), @@ -313,10 +313,10 @@ ]); -qed_goal "contX_Ifix" Fix.thy "contX(Ifix)" +qed_goal "cont_Ifix" Fix.thy "cont(Ifix)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Ifix 1), (rtac contlub_Ifix 1) ]); @@ -325,30 +325,30 @@ (* propagate properties of Ifix to its continuous counterpart *) (* ------------------------------------------------------------------------ *) -qed_goalw "fix_eq" Fix.thy [fix_def] "fix[F]=F[fix[F]]" +qed_goalw "fix_eq" Fix.thy [fix_def] "fix`F = F`(fix`F)" (fn prems => [ - (asm_simp_tac (Cfun_ss addsimps [contX_Ifix]) 1), + (asm_simp_tac (Cfun_ss addsimps [cont_Ifix]) 1), (rtac Ifix_eq 1) ]); -qed_goalw "fix_least" Fix.thy [fix_def] "F[x]=x ==> fix[F] << x" +qed_goalw "fix_least" Fix.thy [fix_def] "F`x = x ==> fix`F << x" (fn prems => [ (cut_facts_tac prems 1), - (asm_simp_tac (Cfun_ss addsimps [contX_Ifix]) 1), + (asm_simp_tac (Cfun_ss addsimps [cont_Ifix]) 1), (etac Ifix_least 1) ]); -qed_goal "fix_eq2" Fix.thy "f == fix[F] ==> f = F[f]" +qed_goal "fix_eq2" Fix.thy "f == fix`F ==> f = F`f" (fn prems => [ (rewrite_goals_tac prems), (rtac fix_eq 1) ]); -qed_goal "fix_eq3" Fix.thy "f == fix[F] ==> f[x] = F[f][x]" +qed_goal "fix_eq3" Fix.thy "f == fix`F ==> f`x = F`f`x" (fn prems => [ (rtac trans 1), @@ -358,7 +358,7 @@ 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]" +qed_goal "fix_eq4" Fix.thy "f = fix`F ==> f = F`f" (fn prems => [ (cut_facts_tac prems 1), @@ -366,7 +366,7 @@ (rtac fix_eq 1) ]); -qed_goal "fix_eq5" Fix.thy "f = fix[F] ==> f[x] = F[f][x]" +qed_goal "fix_eq5" Fix.thy "f = fix`F ==> f`x = F`f`x" (fn prems => [ (rtac trans 1), @@ -383,30 +383,28 @@ (rtac (fixdef RS fix_eq4) 1), (rtac trans 1), (rtac beta_cfun 1), - (contX_tacR 1), + (cont_tacR 1), (rtac refl 1) ]); -(* ------------------------------------------------------------------------ - -given the definition - -smap_def - "smap = fix[LAM h f s. stream_when[LAM x l.scons[f[x]][h[f][l]]][s]]" +(* use this one for definitions! *) -use fix_prover for - -val smap_def2 = fix_prover Stream2.thy smap_def - "smap = (LAM f s. stream_when[LAM x l.scons[f[x]][smap[f][l]]][s])"; - - ------------------------------------------------------------------------ *) +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) + ]); (* ------------------------------------------------------------------------ *) (* better access to definitions *) (* ------------------------------------------------------------------------ *) -qed_goal "Ifix_def2" Fix.thy "Ifix=(%x. lub(range(%i. iterate(i,x,UU))))" +qed_goal "Ifix_def2" Fix.thy "Ifix=(%x. lub(range(%i. iterate i x UU)))" (fn prems => [ (rtac ext 1), @@ -419,11 +417,11 @@ (* ------------------------------------------------------------------------ *) qed_goalw "fix_def2" Fix.thy [fix_def] - "fix[F] = lub(range(%i. iterate(i,F,UU)))" + "fix`F = lub(range(%i. iterate i F UU))" (fn prems => [ (fold_goals_tac [Ifix_def]), - (asm_simp_tac (Cfun_ss addsimps [contX_Ifix]) 1) + (asm_simp_tac (Cfun_ss addsimps [cont_Ifix]) 1) ]); @@ -443,8 +441,8 @@ ]); 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) @@ -470,7 +468,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "fix_ind" Fix.thy -"[| adm(P);P(UU);!!x. P(x) ==> P(F[x])|] ==> P(fix[F])" +"[| adm(P);P(UU);!!x. P(x) ==> P(F`x)|] ==> P(fix`F)" (fn prems => [ (cut_facts_tac prems 1), @@ -492,7 +490,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "wfix_ind" Fix.thy -"[| admw(P); !n. P(iterate(n,F,UU))|] ==> P(fix[F])" +"[| admw(P); !n. P(iterate n F UU)|] ==> P(fix`F)" (fn prems => [ (cut_facts_tac prems 1), @@ -508,7 +506,7 @@ (* ------------------------------------------------------------------------ *) 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)" +"!Y. is_chain(Y::nat=>'a) --> (? n.max_in_chain n Y) ==> adm(P::'a=>bool)" (fn prems => [ (cut_facts_tac prems 1), @@ -590,24 +588,24 @@ (* ------------------------------------------------------------------------ *) qed_goal "iso_strict" Fix.thy -"!!f g.[|!y.f[g[y]]=(y::'b) ; !x.g[f[x]]=(x::'a) |] \ -\ ==> f[UU]=UU & g[UU]=UU" +"!!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), + (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), + (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), @@ -621,7 +619,7 @@ ]); 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), @@ -639,21 +637,21 @@ (* ------------------------------------------------------------------------ *) qed_goalw "chfin2chfin" Fix.thy [chain_finite_def] -"!!f g.[|chain_finite(x::'a); !y.f[g[y]]=(y::'b) ; !x.g[f[x]]=(x::'a) |] \ +"!!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), + (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), + (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), + (res_inst_tac [("s","f`(g`(Y j))"),("t","Y(j)")] subst 1), (etac spec 1), (rtac cfun_arg_cong 1), (rtac mp 1), @@ -662,28 +660,28 @@ ]); qed_goalw "flat2flat" Fix.thy [flat_def] -"!!f g.[|flat(x::'a); !y.f[g[y]]=(y::'b) ; !x.g[f[x]]=(x::'a) |] \ +"!!f g.[|flat(x::'a); !y.f`(g`y)=(y::'b) ; !x.g`(f`x)=(x::'a) |] \ \ ==> 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)" +"[|flat(y::'b);f`(x::'a)=(c::'b)|] ==> 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), + (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), + (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), + (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), + (res_inst_tac [("s","f`x"),("t","c")] subst 1), (atac 1), - (res_inst_tac [("a","f[UU::'a]")] (refl RS box_equals) 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), @@ -730,27 +728,27 @@ (* ------------------------------------------------------------------------ *) qed_goalw "adm_less" Fix.thy [adm_def] - "[|contX(u);contX(v)|]==> adm(%x.u(x)< adm(%x.u x << v x)" (fn prems => [ (cut_facts_tac prems 1), (strip_tac 1), - (etac (contX2contlub RS contlubE RS spec RS mp RS ssubst) 1), + (etac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1), (atac 1), - (etac (contX2contlub RS contlubE RS spec RS mp RS ssubst) 1), + (etac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1), (atac 1), (rtac lub_mono 1), (cut_facts_tac prems 1), - (etac (contX2mono RS ch2ch_monofun) 1), + (etac (cont2mono RS ch2ch_monofun) 1), (atac 1), (cut_facts_tac prems 1), - (etac (contX2mono RS ch2ch_monofun) 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), @@ -768,7 +766,7 @@ ]); 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), @@ -785,7 +783,7 @@ ]); qed_goalw "adm_not_less" Fix.thy [adm_def] - "contX(t) ==> adm(%x.~ t(x) << u)" + "cont t ==> adm(%x.~ (t x) << u)" (fn prems => [ (cut_facts_tac prems 1), @@ -794,13 +792,13 @@ (etac spec 1), (rtac trans_less 1), (atac 2), - (etac (contX2mono RS monofun_fun_arg) 1), + (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), @@ -817,18 +815,18 @@ val adm_all2 = (allI RS adm_all); qed_goal "adm_subst" Fix.thy - "[|contX(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 (contX2contlub RS contlubE RS spec RS mp RS ssubst) 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 (contX2mono RS ch2ch_monofun) 1), + (rtac (cont2mono RS ch2ch_monofun) 1), (atac 1), (atac 1), (atac 1) @@ -843,7 +841,7 @@ ]); qed_goalw "adm_not_UU" Fix.thy [adm_def] - "contX(t)==> adm(%x.~ t(x) = UU)" + "cont(t)==> adm(%x.~ (t x) = UU)" (fn prems => [ (cut_facts_tac prems 1), @@ -851,17 +849,17 @@ (rtac contrapos 1), (etac spec 1), (rtac (chain_UU_I RS spec) 1), - (rtac (contX2mono RS ch2ch_monofun) 1), + (rtac (cont2mono RS ch2ch_monofun) 1), (atac 1), (atac 1), - (rtac (contX2contlub RS contlubE RS spec RS mp RS subst) 1), + (rtac (cont2contlub RS contlubE RS spec RS mp RS subst) 1), (atac 1), (atac 1), (atac 1) ]); qed_goal "adm_eq" Fix.thy - "[|contX(u);contX(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), @@ -887,7 +885,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "adm_disj_lemma1" Pcpo.thy -"[| is_chain(Y); !n.P(Y(n))|Q(Y(n))|]\ +"[| is_chain Y; !n.P (Y n) | Q(Y n)|]\ \ ==> (? i.!j. i Q(Y(j))) | (!i.? j.i [ @@ -914,7 +912,7 @@ qed_goal "adm_disj_lemma3" Fix.thy "[| is_chain(Y); ! j. i < j --> Q(Y(j)) |] ==>\ -\ is_chain(%m. if(m < Suc(i),Y(Suc(i)),Y(m)))" +\ is_chain(%m. if m < Suc i then Y(Suc i) else Y m)" (fn prems => [ (cut_facts_tac prems 1), @@ -946,29 +944,26 @@ qed_goal "adm_disj_lemma4" Fix.thy "[| ! j. i < j --> Q(Y(j)) |] ==>\ -\ ! n. Q(if(n < Suc(i),Y(Suc(i)),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),Y(Suc(i)),Y(m))))" +\ lub(range(Y)) = lub(range(%m. if m< Suc(i) then Y(Suc(i)) else Y m))" (fn prems => [ (cut_facts_tac prems 1), @@ -1013,7 +1008,7 @@ [ (cut_facts_tac prems 1), (etac exE 1), - (res_inst_tac [("x","%m.if(m< Suc(i),Y(Suc(i)),Y(m))")] exI 1), + (res_inst_tac [("x","%m.if m adm(%x.P(x)|Q(x))" + "[| adm P; adm Q |] ==> adm(%x.P x | Q x)" (fn prems => [ (cut_facts_tac prems 1), @@ -1154,11 +1149,11 @@ 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), + (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), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Fix.thy --- a/src/HOLCF/Fix.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Fix.thy Thu Jun 29 16:28:40 1995 +0200 @@ -20,22 +20,22 @@ chain_finite :: "'a=>bool" flat :: "'a=>bool" -rules +defs -iterate_def "iterate(n,F,c) == nat_rec(n,c,%n x.F[x])" -Ifix_def "Ifix(F) == lub(range(%i.iterate(i,F,UU)))" -fix_def "fix == (LAM f. Ifix(f))" +iterate_def "iterate n F c == nat_rec n c (%n x.F`x)" +Ifix_def "Ifix F == lub(range(%i.iterate i F UU))" +fix_def "fix == (LAM f. Ifix f)" -adm_def "adm(P) == !Y. is_chain(Y) --> - (!i.P(Y(i))) --> P(lub(range(Y)))" +adm_def "adm P == !Y. is_chain(Y) --> + (!i.P(Y i)) --> P(lub(range Y))" -admw_def "admw(P)== (!F.((!n.P(iterate(n,F,UU)))--> - P(lub(range(%i.iterate(i,F,UU))))))" +admw_def "admw P == !F. (!n.P (iterate n F UU)) --> + P (lub(range (%i. iterate i F UU)))" -chain_finite_def "chain_finite(x::'a)== - !Y. is_chain(Y::nat=>'a) --> (? n.max_in_chain(n,Y))" +chain_finite_def "chain_finite (x::'a)== + !Y. is_chain (Y::nat=>'a) --> (? n.max_in_chain n Y)" -flat_def "flat(x::'a) == +flat_def "flat (x::'a) == ! x y. (x::'a) << y --> (x = UU) | (x=y)" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Fun1.ML --- a/src/HOLCF/Fun1.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Fun1.ML Thu Jun 29 16:28:40 1995 +0200 @@ -12,14 +12,14 @@ (* less_fun is a partial order on 'a => 'b *) (* ------------------------------------------------------------------------ *) -qed_goalw "refl_less_fun" Fun1.thy [less_fun_def] "less_fun(f,f)" +qed_goalw "refl_less_fun" Fun1.thy [less_fun_def] "less_fun f f" (fn prems => [ (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), @@ -28,7 +28,7 @@ ]); 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), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Fun1.thy --- a/src/HOLCF/Fun1.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Fun1.thy Thu Jun 29 16:28:40 1995 +0200 @@ -17,11 +17,11 @@ consts less_fun :: "['a=>'b::po,'a=>'b] => bool" -rules +defs (* definition of the ordering less_fun *) (* in fun1.ML it is proved that less_fun is a po *) - less_fun_def "less_fun(f1,f2) == ! x. f1(x) << f2(x)" + less_fun_def "less_fun f1 f2 == ! x. f1(x) << f2(x)" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Fun2.ML --- a/src/HOLCF/Fun2.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Fun2.ML Thu Jun 29 16:28:40 1995 +0200 @@ -54,7 +54,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "ub2ub_fun" Fun2.thy - " range(S::nat=>('a::term => 'b::po)) <| u ==> range(%i. S(i,x)) <| u(x)" + " range(S::nat=>('a::term => 'b::po)) <| u ==> range(%i. S i x) <| u(x)" (fn prems => [ (cut_facts_tac prems 1), @@ -93,7 +93,7 @@ ]); val thelub_fun = (lub_fun RS thelubI); -(* is_chain(?S1) ==> lub(range(?S1)) = (%x. lub(range(%i. ?S1(i,x)))) *) +(* 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" diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Fun2.thy --- a/src/HOLCF/Fun2.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Fun2.thy Thu Jun 29 16:28:40 1995 +0200 @@ -24,7 +24,8 @@ inst_fun_po "((op <<)::['a=>'b::po,'a=>'b::po ]=>bool) = less_fun" -(* definitions *) +defs + (* The least element in type 'a::term => 'b::pcpo *) UU_fun_def "UU_fun == (% x.UU)" diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Holcfb.thy --- a/src/HOLCF/Holcfb.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Holcfb.thy Thu Jun 29 16:28:40 1995 +0200 @@ -13,9 +13,9 @@ theleast :: "(nat=>bool)=>nat" -rules +defs -theleast_def "theleast(P) == (@z.(P(z) & (!n.P(n)-->z<=n)))" +theleast_def "theleast P == (@z.(P z & (!n.P n --> z<=n)))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Lift1.ML --- a/src/HOLCF/Lift1.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Lift1.ML Thu Jun 29 16:28:40 1995 +0200 @@ -45,7 +45,7 @@ (atac 1) ]); -qed_goalw "defined_Iup" Lift1.thy [Iup_def,UU_lift_def] "~ Iup(x)=UU_lift" +qed_goalw "defined_Iup" Lift1.thy [Iup_def,UU_lift_def] "Iup(x)~=UU_lift" (fn prems => [ (rtac notI 1), @@ -76,7 +76,7 @@ ]); 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), @@ -96,7 +96,7 @@ ]); 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), @@ -110,7 +110,7 @@ ]); 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), @@ -121,7 +121,7 @@ ]); -qed_goal "refl_less_lift" Lift1.thy "less_lift(p,p)" +qed_goal "refl_less_lift" Lift1.thy "less_lift p p" (fn prems => [ (res_inst_tac [("p","p")] liftE 1), @@ -133,7 +133,7 @@ ]); 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), @@ -143,13 +143,13 @@ (hyp_subst_tac 1), (rtac refl 1), (hyp_subst_tac 1), - (res_inst_tac [("P","less_lift(Iup(x),UU_lift)")] notE 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), + (res_inst_tac [("P","less_lift (Iup x) UU_lift")] notE 1), (rtac less_lift1b 1), (atac 1), (hyp_subst_tac 1), @@ -160,7 +160,7 @@ ]); 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), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Lift1.thy --- a/src/HOLCF/Lift1.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Lift1.thy Thu Jun 29 16:28:40 1995 +0200 @@ -36,11 +36,12 @@ (*defining the abstract constants*) +defs UU_lift_def "UU_lift == Abs_Lift(Inl(UU))" Iup_def "Iup(x) == Abs_Lift(Inr(x))" Ilift_def "Ilift(f)(x)== - case Rep_Lift(x) of Inl(y) => UU | Inr(z) => f[z]" + case Rep_Lift(x) of Inl(y) => UU | Inr(z) => f`z" less_lift_def "less_lift(x1)(x2) == (case Rep_Lift(x1) of diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Lift2.ML --- a/src/HOLCF/Lift2.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Lift2.ML Thu Jun 29 16:28:40 1995 +0200 @@ -132,7 +132,7 @@ ]); qed_goal "lub_lift1b" Lift2.thy -"[|is_chain(Y);!i x.~Y(i)=Iup(x)|] ==>\ +"[|is_chain(Y);!i x. Y(i)~=Iup(x)|] ==>\ \ range(Y) <<| UU_lift" (fn prems => [ @@ -155,13 +155,16 @@ ]); val thelub_lift1a = lub_lift1a RS thelubI; -(* [| is_chain(?Y1); ? i x. ?Y1(i) = Iup(x) |] ==> *) -(* lub(range(?Y1)) = Iup(lub(range(%i. Ilift(LAM x. x,?Y1(i))))) *) +(* +[| is_chain ?Y1; ? i x. ?Y1 i = Iup x |] ==> + lub (range ?Y1) = Iup (lub (range (%i. Ilift (LAM x. x) (?Y1 i)))) +*) val thelub_lift1b = lub_lift1b RS thelubI; -(* [| is_chain(?Y1); ! i x. ~ ?Y1(i) = Iup(x) |] ==> *) -(* lub(range(?Y1)) = UU_lift *) - +(* +[| is_chain ?Y1; ! i x. ?Y1 i ~= Iup x |] ==> + lub (range ?Y1) = UU_lift +*) qed_goal "cpo_lift" Lift2.thy "is_chain(Y::nat=>('a)u) ==> ? x.range(Y) <<|x" diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Lift3.ML --- a/src/HOLCF/Lift3.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Lift3.ML Thu Jun 29 16:28:40 1995 +0200 @@ -19,7 +19,7 @@ (rtac less_lift2b 1) ]); -qed_goal "defined_Iup2" Lift3.thy "~ Iup(x) = UU" +qed_goal "defined_Iup2" Lift3.thy "Iup(x) ~= UU" (fn prems => [ (rtac (inst_lift_pcpo RS ssubst) 1), @@ -47,10 +47,10 @@ (asm_simp_tac Lift_ss 1) ]); -qed_goal "contX_Iup" Lift3.thy "contX(Iup)" +qed_goal "cont_Iup" Lift3.thy "cont(Iup)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Iup 1), (rtac contlub_Iup 1) ]); @@ -124,18 +124,18 @@ (atac 1) ]); -qed_goal "contX_Ilift1" Lift3.thy "contX(Ilift)" +qed_goal "cont_Ilift1" Lift3.thy "cont(Ilift)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Ilift1 1), (rtac contlub_Ilift1 1) ]); -qed_goal "contX_Ilift2" Lift3.thy "contX(Ilift(f))" +qed_goal "cont_Ilift2" Lift3.thy "cont(Ilift(f))" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Ilift2 1), (rtac contlub_Ilift2 1) ]); @@ -145,100 +145,100 @@ (* continuous versions of lemmas for ('a)u *) (* ------------------------------------------------------------------------ *) -qed_goalw "Exh_Lift1" Lift3.thy [up_def] "z = UU | (? x. z = up[x])" +qed_goalw "Exh_Lift1" Lift3.thy [up_def] "z = UU | (? x. z = up`x)" (fn prems => [ - (simp_tac (Lift_ss addsimps [contX_Iup]) 1), + (simp_tac (Lift_ss addsimps [cont_Iup]) 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" +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), - (simp_tac (Lift_ss addsimps [contX_Iup]) 1), - (simp_tac (Lift_ss addsimps [contX_Iup]) 1) + (simp_tac (Lift_ss addsimps [cont_Iup]) 1), + (simp_tac (Lift_ss addsimps [cont_Iup]) 1) ]); -qed_goalw "defined_up" Lift3.thy [up_def] "~ up[x]=UU" +qed_goalw "defined_up" Lift3.thy [up_def] " up`x ~= UU" (fn prems => [ - (simp_tac (Lift_ss addsimps [contX_Iup]) 1), + (simp_tac (Lift_ss addsimps [cont_Iup]) 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), - (asm_simp_tac (Lift_ss addsimps [contX_Iup]) 1) + (asm_simp_tac (Lift_ss addsimps [cont_Iup]) 1) ]); -qed_goalw "lift1" Lift3.thy [up_def,lift_def] "lift[f][UU]=UU" +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 (contX_lemmas @ [contX_Iup,contX_Ilift1, - contX_Ilift2,contX2contX_CF1L]) 1)), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (contX_lemmas @ [contX_Iup,contX_Ilift1, - contX_Ilift2,contX2contX_CF1L]) 1)), - (simp_tac (Lift_ss addsimps [contX_Iup,contX_Ilift1,contX_Ilift2]) 1) + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), + (simp_tac (Lift_ss addsimps [cont_Iup,cont_Ilift1,cont_Ilift2]) 1) ]); -qed_goalw "lift2" Lift3.thy [up_def,lift_def] "lift[f][up[x]]=f[x]" +qed_goalw "lift2" Lift3.thy [up_def,lift_def] "lift`f`(up`x)=f`x" (fn prems => [ (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Iup 1), + (rtac cont_Iup 1), (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (contX_lemmas @ [contX_Iup,contX_Ilift1, - contX_Ilift2,contX2contX_CF1L]) 1)), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Ilift2 1), - (simp_tac (Lift_ss addsimps [contX_Iup,contX_Ilift1,contX_Ilift2]) 1) + (rtac cont_Ilift2 1), + (simp_tac (Lift_ss addsimps [cont_Iup,cont_Ilift1,cont_Ilift2]) 1) ]); -qed_goalw "less_lift4b" Lift3.thy [up_def,lift_def] "~ up[x] << UU" +qed_goalw "less_lift4b" Lift3.thy [up_def,lift_def] "~ up`x << UU" (fn prems => [ - (simp_tac (Lift_ss addsimps [contX_Iup]) 1), + (simp_tac (Lift_ss addsimps [cont_Iup]) 1), (rtac less_lift3b 1) ]); qed_goalw "less_lift4c" Lift3.thy [up_def,lift_def] - "(up[x]< [ - (simp_tac (Lift_ss addsimps [contX_Iup]) 1), + (simp_tac (Lift_ss addsimps [cont_Iup]) 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)]))]" +"[| 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 (contX_lemmas @ [contX_Iup,contX_Ilift1, - contX_Ilift2,contX2contX_CF1L]) 1)), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (contX_lemmas @ [contX_Iup,contX_Ilift1, - contX_Ilift2,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iup,contX_Ilift1, - contX_Ilift2,contX2contX_CF1L]) 1)), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iup,cont_Ilift1, + cont_Ilift2,cont2cont_CF1L]) 1)), (rtac thelub_lift1a 1), (atac 1), (etac exE 1), @@ -247,13 +247,13 @@ (rtac exI 1), (etac box_equals 1), (rtac refl 1), - (simp_tac (Lift_ss addsimps [contX_Iup]) 1) + (simp_tac (Lift_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" +"[| is_chain(Y); ! i x. Y(i) ~= up`x |] ==> lub(range(Y)) = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -268,11 +268,11 @@ (dtac notnotD 1), (etac box_equals 1), (rtac refl 1), - (simp_tac (Lift_ss addsimps [contX_Iup]) 1) + (simp_tac (Lift_ss addsimps [cont_Iup]) 1) ]); -qed_goal "lift_lemma2" Lift3.thy " (? x.z = up[x]) = (~z=UU)" +qed_goal "lift_lemma2" Lift3.thy " (? x.z = up`x) = (z~=UU)" (fn prems => [ (rtac iffI 1), @@ -287,7 +287,7 @@ qed_goal "thelub_lift2a_rev" Lift3.thy -"[| is_chain(Y); lub(range(Y)) = up[x] |] ==> ? i x. Y(i) = up[x]" +"[| is_chain(Y); lub(range(Y)) = up`x |] ==> ? i x. Y(i) = up`x" (fn prems => [ (cut_facts_tac prems 1), @@ -301,7 +301,7 @@ ]); qed_goal "thelub_lift2b_rev" Lift3.thy -"[| is_chain(Y); lub(range(Y)) = UU |] ==> ! i x. ~ Y(i) = up[x]" +"[| is_chain(Y); lub(range(Y)) = UU |] ==> ! i x. Y(i) ~= up`x" (fn prems => [ (cut_facts_tac prems 1), @@ -318,7 +318,7 @@ 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)]))]" +\ lub(range(Y)) = up`(lub(range(%i. lift`(LAM x.x)`(Y i))))" (fn prems => [ (cut_facts_tac prems 1), @@ -334,7 +334,7 @@ (fast_tac HOL_cs 1) ]); -qed_goal "lift3" Lift3.thy "lift[up][x]=x" +qed_goal "lift3" Lift3.thy "lift`up`x=x" (fn prems => [ (res_inst_tac [("p","x")] liftE1 1), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Lift3.thy --- a/src/HOLCF/Lift3.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Lift3.thy Thu Jun 29 16:28:40 1995 +0200 @@ -18,10 +18,11 @@ rules -inst_lift_pcpo "(UU::('a)u) = UU_lift" + inst_lift_pcpo "(UU::('a)u) = UU_lift" -up_def "up == (LAM x.Iup(x))" -lift_def "lift == (LAM f p.Ilift(f)(p))" +defs + up_def "up == (LAM x.Iup(x))" + lift_def "lift == (LAM f p.Ilift(f)(p))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/One.ML --- a/src/HOLCF/One.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/One.ML Thu Jun 29 16:28:40 1995 +0200 @@ -15,7 +15,7 @@ qed_goalw "Exh_one" One.thy [one_def] "z=UU | z = one" (fn prems => [ - (res_inst_tac [("p","rep_one[z]")] liftE1 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), @@ -55,7 +55,7 @@ ]) ]; -val dist_eq_one = [prove_goal One.thy "~one=UU" +val dist_eq_one = [prove_goal One.thy "one~=UU" (fn prems => [ (rtac not_less2not_eq 1), @@ -94,5 +94,5 @@ RS iso_strict) RS conjunct1] )1) ]); -val one_when = map prover ["one_when[x][UU] = UU","one_when[x][one] = x"]; +val one_when = map prover ["one_when`x`UU = UU","one_when`x`one = x"]; diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/One.thy --- a/src/HOLCF/One.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/One.thy Thu Jun 29 16:28:40 1995 +0200 @@ -29,12 +29,12 @@ one_when :: "'c -> one -> 'c" rules - abs_one_iso "abs_one[rep_one[u]] = u" - rep_one_iso "rep_one[abs_one[x]] = x" + abs_one_iso "abs_one`(rep_one`u) = u" + rep_one_iso "rep_one`(abs_one`x) = x" - one_def "one == abs_one[up[UU]]" - one_when_def "one_when == (LAM c u.lift[LAM x.c][rep_one[u]])" - +defs + one_def "one == abs_one`(up`UU)" + one_when_def "one_when == (LAM c u.lift`(LAM x.c)`(rep_one`u))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Porder.ML --- a/src/HOLCF/Porder.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Porder.ML Thu Jun 29 16:28:40 1995 +0200 @@ -301,7 +301,7 @@ (* ------------------------------------------------------------------------ *) 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), @@ -323,7 +323,7 @@ ]); 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), @@ -334,7 +334,7 @@ ]); -qed_goal "bin_chain" Porder.thy "x< is_chain(%i. if(i=0,x,y))" +qed_goal "bin_chain" Porder.thy "x< is_chain (%i. if i=0 then x else y)" (fn prems => [ (cut_facts_tac prems 1), @@ -347,7 +347,7 @@ ]); qed_goalw "bin_chainmax" Porder.thy [max_in_chain_def,le_def] - "x< max_in_chain(Suc(0),%i. if(i=0,x,y))" + "x< max_in_chain (Suc 0) (%i. if (i=0) then x else y)" (fn prems => [ (cut_facts_tac prems 1), @@ -358,10 +358,10 @@ ]); qed_goal "lub_bin_chain" Porder.thy - "x << y ==> range(%i. if(i = 0,x,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,x,y)")] subst 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), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Porder.thy --- a/src/HOLCF/Porder.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Porder.thy Thu Jun 29 16:28:40 1995 +0200 @@ -18,14 +18,13 @@ max_in_chain :: "[nat,nat=>'a::po]=>bool" finite_chain :: "(nat=>'a::po)=>bool" -rules +defs (* class definitions *) is_ub "S <| x == ! y.y:S --> y< x << u)" -lub "lub(S) = (@x. S <<| x)" (* Arbitrary chains are total orders *) is_tord "is_tord(S) == ! x y. x:S & y:S --> (x< C(i) = C(j)" +max_in_chain_def "max_in_chain i C == ! j. i <= j --> C(i) = C(j)" + +finite_chain_def "finite_chain(C) == is_chain(C) & (? i. max_in_chain i C)" -finite_chain_def "finite_chain(C) == is_chain(C) & (? i. max_in_chain(i,C))" +rules + +lub "lub(S) = (@x. S <<| x)" end + + diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/README --- a/src/HOLCF/README Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/README Thu Jun 29 16:28:40 1995 +0200 @@ -13,4 +13,9 @@ Dissertation, Technische Universit"at M"unchen, 1994 Changes: -14.10. New translation mechanism for continuous infixes +14.10.94 New translation mechanism for continuous infixes +18.05.95 Conversion to curried version of HOL. + +28.06.95 The old uncurried version of HOLCF is no longer supported + in the distribution. + diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ROOT.ML --- a/src/HOLCF/ROOT.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ROOT.ML Thu Jun 29 16:28:40 1995 +0200 @@ -7,12 +7,13 @@ Should be executed in subdirectory HOLCF. *) -val banner = "Higher-order Logic of Computable Functions"; +val banner = "Higher-order Logic of Computable Functions; curried version"; writeln banner; print_depth 1; init_thy_reader(); + use_thy "Holcfb"; use_thy "Void"; @@ -63,6 +64,7 @@ use_thy "Stream"; use_thy "Stream2"; + use_thy "Dlist"; use "../Pure/install_pp.ML"; diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Sprod0.ML --- a/src/HOLCF/Sprod0.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Sprod0.ML Thu Jun 29 16:28:40 1995 +0200 @@ -13,7 +13,7 @@ (* ------------------------------------------------------------------------ *) 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]) @@ -21,7 +21,7 @@ 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), @@ -35,7 +35,7 @@ qed_goalw "strict_Spair_Rep" Sprod0.thy [Spair_Rep_def] - "(a=UU | b=UU) ==> (Spair_Rep(a,b) = Spair_Rep(UU,UU))" + "(a=UU | b=UU) ==> (Spair_Rep a b) = (Spair_Rep UU UU)" (fn prems => [ (cut_facts_tac prems 1), @@ -47,7 +47,7 @@ ]); qed_goalw "defined_Spair_Rep_rev" Sprod0.thy [Spair_Rep_def] - "(Spair_Rep(a,b) = Spair_Rep(UU,UU)) ==> (a=UU | b=UU)" + "(Spair_Rep a b) = (Spair_Rep UU UU) ==> (a=UU | b=UU)" (fn prems => [ (res_inst_tac [("Q","a=UU|b=UU")] classical2 1), @@ -65,7 +65,7 @@ (* ------------------------------------------------------------------------ *) 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" +"[|~aa=UU ; ~ba=UU ; Spair_Rep a b = Spair_Rep aa ba |] ==> a=aa & b=ba" (fn prems => [ (cut_facts_tac prems 1), @@ -77,7 +77,7 @@ 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), @@ -94,7 +94,7 @@ (* ------------------------------------------------------------------------ *) qed_goalw "strict_Ispair" Sprod0.thy [Ispair_def] - "(a=UU | b=UU) ==> Ispair(a,b)=Ispair(UU,UU)" + "(a=UU | b=UU) ==> Ispair a b = Ispair UU UU" (fn prems => [ (cut_facts_tac prems 1), @@ -102,7 +102,7 @@ ]); 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), @@ -111,7 +111,7 @@ ]); 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), @@ -120,7 +120,7 @@ ]); 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), @@ -130,7 +130,7 @@ ]); 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), @@ -142,7 +142,7 @@ ]); qed_goal "defined_Ispair" Sprod0.thy -"[|~a=UU; ~b=UU|] ==> ~(Ispair(a,b) = Ispair(UU,UU))" +"[|a~=UU; b~=UU|] ==> (Ispair a b) ~= (Ispair UU UU)" (fn prems => [ (cut_facts_tac prems 1), @@ -159,7 +159,7 @@ (* ------------------------------------------------------------------------ *) 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), @@ -186,7 +186,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "IsprodE" Sprod0.thy -"[|p=Ispair(UU,UU) ==> Q ;!!x y. [|p=Ispair(x,y); ~x=UU ; ~y=UU|] ==> Q|] ==> Q" +"[|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), @@ -206,7 +206,7 @@ (* ------------------------------------------------------------------------ *) 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), @@ -214,7 +214,7 @@ (rtac conjI 1), (fast_tac HOL_cs 1), (strip_tac 1), - (res_inst_tac [("P","Ispair(UU,UU) = Ispair(a,b)")] notE 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)) @@ -222,7 +222,7 @@ qed_goal "strict_Isfst1" Sprod0.thy - "Isfst(Ispair(UU,y)) = UU" + "Isfst(Ispair UU y) = UU" (fn prems => [ (rtac (strict_Ispair1 RS ssubst) 1), @@ -231,7 +231,7 @@ ]); qed_goal "strict_Isfst2" Sprod0.thy - "Isfst(Ispair(x,UU)) = UU" + "Isfst(Ispair x UU) = UU" (fn prems => [ (rtac (strict_Ispair2 RS ssubst) 1), @@ -241,7 +241,7 @@ 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), @@ -249,14 +249,14 @@ (rtac conjI 1), (fast_tac HOL_cs 1), (strip_tac 1), - (res_inst_tac [("P","Ispair(UU,UU) = Ispair(a,b)")] notE 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), @@ -265,7 +265,7 @@ ]); qed_goal "strict_Issnd2" Sprod0.thy - "Issnd(Ispair(x,UU)) = UU" + "Issnd(Ispair x UU) = UU" (fn prems => [ (rtac (strict_Ispair2 RS ssubst) 1), @@ -274,14 +274,14 @@ ]); 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), + (res_inst_tac [("P","Ispair x y = Ispair UU UU")] notE 1), (etac defined_Ispair 1), (atac 1), (atac 1), @@ -294,14 +294,14 @@ ]); 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), + (res_inst_tac [("P","Ispair x y = Ispair UU UU")] notE 1), (etac defined_Ispair 1), (atac 1), (atac 1), @@ -313,7 +313,7 @@ (fast_tac HOL_cs 1) ]); -qed_goal "Isfst2" Sprod0.thy "~y=UU ==>Isfst(Ispair(x,y))=x" +qed_goal "Isfst2" Sprod0.thy "y~=UU ==>Isfst(Ispair x y)=x" (fn prems => [ (cut_facts_tac prems 1), @@ -324,7 +324,7 @@ (rtac strict_Isfst1 1) ]); -qed_goal "Issnd2" Sprod0.thy "~x=UU ==>Issnd(Ispair(x,y))=y" +qed_goal "Issnd2" Sprod0.thy "~x=UU ==>Issnd(Ispair x y)=y" (fn prems => [ (cut_facts_tac prems 1), @@ -347,7 +347,7 @@ 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), @@ -365,7 +365,7 @@ (* ------------------------------------------------------------------------ *) 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), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Sprod0.thy --- a/src/HOLCF/Sprod0.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Sprod0.thy Thu Jun 29 16:28:40 1995 +0200 @@ -23,13 +23,13 @@ Isfst :: "('a ** 'b) => 'a" Issnd :: "('a ** 'b) => 'b" -rules - +defs Spair_Rep_def "Spair_Rep == (%a b. %x y. (~a=UU & ~b=UU --> x=a & y=b ))" - Sprod_def "Sprod == {f. ? a b. f = Spair_Rep(a,b)}" + Sprod_def "Sprod == {f. ? a b. f = Spair_Rep a b}" +rules (*faking a type definition... *) (* "**" is isomorphic to Sprod *) @@ -37,17 +37,18 @@ Rep_Sprod_inverse "Abs_Sprod(Rep_Sprod(p)) = p" Abs_Sprod_inverse "f:Sprod ==> Rep_Sprod(Abs_Sprod(f)) = f" +defs (*defining the abstract constants*) - Ispair_def "Ispair(a,b) == Abs_Sprod(Spair_Rep(a,b))" + Ispair_def "Ispair a b == Abs_Sprod(Spair_Rep a b)" Isfst_def "Isfst(p) == @z. - (p=Ispair(UU,UU) --> z=UU) - &(! a b. ~a=UU & ~b=UU & p=Ispair(a,b) --> z=a)" + (p=Ispair UU UU --> z=UU) + &(! a b. ~a=UU & ~b=UU & p=Ispair a b --> z=a)" Issnd_def "Issnd(p) == @z. - (p=Ispair(UU,UU) --> z=UU) - &(! a b. ~a=UU & ~b=UU & p=Ispair(a,b) --> z=b)" + (p=Ispair UU UU --> z=UU) + &(! a b. ~a=UU & ~b=UU & p=Ispair a b --> z=b)" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Sprod1.ML --- a/src/HOLCF/Sprod1.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Sprod1.ML Thu Jun 29 16:28:40 1995 +0200 @@ -14,39 +14,24 @@ qed_goalw "less_sprod1a" Sprod1.thy [less_sprod_def] - "p1=Ispair(UU,UU) ==> less_sprod(p1,p2)" -(fn prems => + "p1=Ispair UU UU ==> less_sprod p1 p2" + (fn prems => [ (cut_facts_tac prems 1), - (rtac eqTrueE 1), - (rtac select_equality 1), - (rtac conjI 1), - (fast_tac HOL_cs 1), - (strip_tac 1), - (contr_tac 1), - (dtac conjunct1 1), - (etac rev_mp 1), - (atac 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 => + "p1~=Ispair UU UU ==> \ +\ less_sprod p1 p2 = ( Isfst p1 << Isfst p2 & Issnd p1 << Issnd p2)" + (fn prems => [ (cut_facts_tac prems 1), - (rtac select_equality 1), - (rtac conjI 1), - (strip_tac 1), - (contr_tac 1), - (fast_tac HOL_cs 1), - (dtac conjunct2 1), - (etac rev_mp 1), - (atac 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), @@ -55,18 +40,18 @@ (rtac disjI1 1), (rtac antisym_less 1), (rtac minimal 2), - (res_inst_tac [("s","Isfst(Ispair(x,y))"),("t","x")] subst 1), + (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), + (res_inst_tac [("s","Isfst(Ispair UU UU)"),("t","UU")] subst 1), (simp_tac Sprod_ss 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)" + "less_sprod p (Ispair UU UU) ==> p = Ispair UU UU" (fn prems => [ (cut_facts_tac prems 1), @@ -78,22 +63,22 @@ ]); 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" + "[|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), + (res_inst_tac [("s","Isfst(Ispair xa ya)"),("t","xa")] subst 1), (simp_tac (Sprod_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 (Sprod_ss addsimps prems)1), (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct1) 1), (resolve_tac prems 1), (resolve_tac prems 1), (simp_tac (Sprod_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 (Sprod_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 (Sprod_ss addsimps prems)1), (rtac (defined_Ispair RS less_sprod1b RS iffD1 RS conjunct2) 1), (resolve_tac prems 1), @@ -105,7 +90,7 @@ (* less_sprod is a partial order on Sprod *) (* ------------------------------------------------------------------------ *) -qed_goal "refl_less_sprod" Sprod1.thy "less_sprod(p,p)" +qed_goal "refl_less_sprod" Sprod1.thy "less_sprod p p" (fn prems => [ (res_inst_tac [("p","p")] IsprodE 1), @@ -118,7 +103,7 @@ qed_goal "antisym_less_sprod" Sprod1.thy - "[|less_sprod(p1,p2);less_sprod(p2,p1)|] ==> p1=p2" + "[|less_sprod p1 p2;less_sprod p2 p1|] ==> p1=p2" (fn prems => [ (cut_facts_tac prems 1), @@ -146,7 +131,7 @@ ]); qed_goal "trans_less_sprod" Sprod1.thy - "[|less_sprod(p1::'a**'b,p2);less_sprod(p2,p3)|] ==> less_sprod(p1,p3)" + "[|less_sprod (p1::'a**'b) p2;less_sprod p2 p3|] ==> less_sprod p1 p3" (fn prems => [ (cut_facts_tac prems 1), @@ -155,11 +140,11 @@ (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), + (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)")] + (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)), @@ -181,7 +166,7 @@ (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)")] + (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 cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Sprod1.thy --- a/src/HOLCF/Sprod1.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Sprod1.thy Thu Jun 29 16:28:40 1995 +0200 @@ -11,12 +11,10 @@ consts less_sprod :: "[('a ** 'b),('a ** 'b)] => bool" -rules - - less_sprod_def "less_sprod(p1,p2) == @z. - ( p1=Ispair(UU,UU) --> z = True) - &(~p1=Ispair(UU,UU) --> z = ( Isfst(p1) << Isfst(p2) & - Issnd(p1) << Issnd(p2)))" +defs + less_sprod_def "less_sprod p1 p2 == + if p1 = Ispair UU UU + then True + else Isfst p1 << Isfst p2 & Issnd p1 << Issnd p2" end - diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Sprod2.ML --- a/src/HOLCF/Sprod2.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Sprod2.ML Thu Jun 29 16:28:40 1995 +0200 @@ -14,7 +14,7 @@ (* ------------------------------------------------------------------------ *) 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), @@ -24,7 +24,7 @@ qed_goal "less_sprod3b" Sprod2.thy - "~p1=Ispair(UU,UU) ==>\ + "p1~=Ispair UU UU ==>\ \ (p1< [ @@ -34,7 +34,7 @@ ]); 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), @@ -43,10 +43,10 @@ ]); val less_sprod4a = (less_sprod4b RS defined_Ispair_rev); -(* Ispair(?a,?b) << Ispair(UU,UU) ==> ?a = UU | ?b = UU *) +(* Ispair ?a ?b << Ispair UU UU ==> ?a = UU | ?b = UU *) qed_goal "less_sprod4c" Sprod2.thy - "[|Ispair(xa,ya)<\ + "[|Ispair xa ya << Ispair x y; xa~=UU; ya~=UU; x~=UU; y~=UU|] ==>\ \ xa< [ @@ -60,7 +60,7 @@ (* type sprod is pointed *) (* ------------------------------------------------------------------------ *) -qed_goal "minimal_sprod" Sprod2.thy "Ispair(UU,UU)< [ (rtac less_sprod3a 1), @@ -78,9 +78,9 @@ (rtac (less_fun RS iffD2) 1), (strip_tac 1), (res_inst_tac [("Q", - " Ispair(y,xa) =Ispair(UU,UU)")] (excluded_middle RS disjE) 1), + " 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), + " Ispair x xa = Ispair UU UU")] (excluded_middle RS disjE) 1), (rtac (less_sprod3b RS iffD2) 1), (atac 1), (rtac conjI 1), @@ -100,9 +100,9 @@ (rtac refl_less 1), (etac less_sprod3a 1), (res_inst_tac [("Q", - " Ispair(x,xa) =Ispair(UU,UU)")] (excluded_middle RS disjE) 1), + " 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), + (res_inst_tac [("P","Ispair y xa = Ispair UU UU")] notE 1), (atac 2), (rtac defined_Ispair 1), (etac notUU_I 1), @@ -116,9 +116,9 @@ [ (strip_tac 1), (res_inst_tac [("Q", - " Ispair(x,y) =Ispair(UU,UU)")] (excluded_middle RS disjE) 1), + " 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), + " Ispair x xa = Ispair UU UU")] (excluded_middle RS disjE) 1), (rtac (less_sprod3b RS iffD2) 1), (atac 1), (rtac conjI 1), @@ -138,9 +138,9 @@ (atac 1), (etac less_sprod3a 1), (res_inst_tac [("Q", - " Ispair(x,xa) =Ispair(UU,UU)")] (excluded_middle RS disjE) 1), + " 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), + (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), @@ -149,7 +149,7 @@ ]); qed_goal " monofun_Ispair" Sprod2.thy - "[|x1< Ispair(x1,y1)< Ispair x1 y1 << Ispair x2 y2" (fn prems => [ (cut_facts_tac prems 1), @@ -179,7 +179,7 @@ (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), + (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), @@ -206,7 +206,7 @@ (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), + (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), @@ -227,7 +227,7 @@ qed_goal "lub_sprod" Sprod2.thy "[|is_chain(S)|] ==> range(S) <<| \ -\ Ispair(lub(range(%i.Isfst(S(i)))),lub(range(%i.Issnd(S(i)))))" +\ Ispair (lub(range(%i.Isfst(S i)))) (lub(range(%i.Issnd(S i))))" (fn prems => [ (cut_facts_tac prems 1), @@ -253,8 +253,7 @@ ]); val thelub_sprod = (lub_sprod RS thelubI); -(* is_chain(?S1) ==> lub(range(?S1)) = *) -(* Ispair(lub(range(%i. Isfst(?S1(i)))),lub(range(%i. Issnd(?S1(i))))) *) + qed_goal "cpo_sprod" Sprod2.thy "is_chain(S::nat=>'a**'b)==>? x.range(S)<<| x" diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Sprod3.ML --- a/src/HOLCF/Sprod3.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Sprod3.ML Thu Jun 29 16:28:40 1995 +0200 @@ -14,9 +14,9 @@ qed_goal "sprod3_lemma1" Sprod3.thy "[| is_chain(Y); x~= UU; lub(range(Y))~= UU |] ==>\ -\ Ispair(lub(range(Y)),x) =\ -\ Ispair(lub(range(%i. Isfst(Ispair(Y(i),x)))),\ -\ lub(range(%i. Issnd(Ispair(Y(i),x)))))" +\ Ispair (lub(range Y)) x =\ +\ Ispair (lub(range(%i. Isfst(Ispair(Y i) x)))) \ +\ (lub(range(%i. Issnd(Ispair(Y i) x))))" (fn prems => [ (cut_facts_tac prems 1), @@ -50,10 +50,10 @@ ]); qed_goal "sprod3_lemma2" Sprod3.thy -"[| is_chain(Y); ~ x = UU; lub(range(Y)) = UU |] ==>\ -\ Ispair(lub(range(Y)),x) =\ -\ Ispair(lub(range(%i. Isfst(Ispair(Y(i),x)))),\ -\ lub(range(%i. Issnd(Ispair(Y(i),x)))))" +"[| is_chain(Y); x ~= UU; lub(range(Y)) = UU |] ==>\ +\ Ispair (lub(range Y)) x =\ +\ Ispair (lub(range(%i. Isfst(Ispair(Y i) x))))\ +\ (lub(range(%i. Issnd(Ispair(Y i) x))))" (fn prems => [ (cut_facts_tac prems 1), @@ -73,9 +73,9 @@ qed_goal "sprod3_lemma3" Sprod3.thy "[| is_chain(Y); x = UU |] ==>\ -\ Ispair(lub(range(Y)),x) =\ -\ Ispair(lub(range(%i. Isfst(Ispair(Y(i),x)))),\ -\ lub(range(%i. Issnd(Ispair(Y(i),x)))))" +\ Ispair (lub(range Y)) x =\ +\ Ispair (lub(range(%i. Isfst(Ispair (Y i) x))))\ +\ (lub(range(%i. Issnd(Ispair (Y i) x))))" (fn prems => [ (cut_facts_tac prems 1), @@ -118,17 +118,17 @@ ]); qed_goal "sprod3_lemma4" Sprod3.thy -"[| is_chain(Y); ~ x = UU; ~ lub(range(Y)) = UU |] ==>\ -\ Ispair(x,lub(range(Y))) =\ -\ Ispair(lub(range(%i. Isfst(Ispair(x,Y(i))))),\ -\ lub(range(%i. Issnd(Ispair(x,Y(i))))))" +"[| is_chain(Y); x ~= UU; lub(range(Y)) ~= UU |] ==>\ +\ Ispair x (lub(range Y)) =\ +\ Ispair (lub(range(%i. Isfst (Ispair x (Y i)))))\ +\ (lub(range(%i. Issnd (Ispair x (Y i)))))" (fn prems => [ (cut_facts_tac prems 1), (res_inst_tac [("f1","Ispair")] (arg_cong RS cong) 1), (rtac sym 1), (rtac lub_chain_maxelem 1), - (res_inst_tac [("P","%j.~Y(j)=UU")] exE 1), + (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), @@ -154,10 +154,10 @@ ]); qed_goal "sprod3_lemma5" Sprod3.thy -"[| is_chain(Y); ~ x = UU; lub(range(Y)) = UU |] ==>\ -\ Ispair(x,lub(range(Y))) =\ -\ Ispair(lub(range(%i. Isfst(Ispair(x,Y(i))))),\ -\ lub(range(%i. Issnd(Ispair(x,Y(i))))))" +"[| is_chain(Y); x ~= UU; lub(range(Y)) = UU |] ==>\ +\ Ispair x (lub(range Y)) =\ +\ Ispair (lub(range(%i. Isfst(Ispair x (Y i)))))\ +\ (lub(range(%i. Issnd(Ispair x (Y i)))))" (fn prems => [ (cut_facts_tac prems 1), @@ -176,9 +176,9 @@ qed_goal "sprod3_lemma6" Sprod3.thy "[| is_chain(Y); x = UU |] ==>\ -\ Ispair(x,lub(range(Y))) =\ -\ Ispair(lub(range(%i. Isfst(Ispair(x,Y(i))))),\ -\ lub(range(%i. Issnd(Ispair(x,Y(i))))))" +\ Ispair x (lub(range Y)) =\ +\ Ispair (lub(range(%i. Isfst (Ispair x (Y i)))))\ +\ (lub(range(%i. Issnd (Ispair x (Y i)))))" (fn prems => [ (cut_facts_tac prems 1), @@ -215,19 +215,19 @@ ]); -qed_goal "contX_Ispair1" Sprod3.thy "contX(Ispair)" +qed_goal "cont_Ispair1" Sprod3.thy "cont(Ispair)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Ispair1 1), (rtac contlub_Ispair1 1) ]); -qed_goal "contX_Ispair2" Sprod3.thy "contX(Ispair(x))" +qed_goal "cont_Ispair2" Sprod3.thy "cont(Ispair(x))" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Ispair2 1), (rtac contlub_Ispair2 1) ]); @@ -289,18 +289,18 @@ ]); -qed_goal "contX_Isfst" Sprod3.thy "contX(Isfst)" +qed_goal "cont_Isfst" Sprod3.thy "cont(Isfst)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Isfst 1), (rtac contlub_Isfst 1) ]); -qed_goal "contX_Issnd" Sprod3.thy "contX(Issnd)" +qed_goal "cont_Issnd" Sprod3.thy "cont(Issnd)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Issnd 1), (rtac contlub_Issnd 1) ]); @@ -312,7 +312,7 @@ -------------------------------------------------------------------------- *) -qed_goal "spair_eq" Sprod3.thy "[|x1=x2;y1=y2|] ==> x1##y1 = x2##y2" +qed_goal "spair_eq" Sprod3.thy "[|x1=x2;y1=y2|] ==> (|x1,y1|) = (|x2,y2|)" (fn prems => [ (cut_facts_tac prems 1), @@ -324,21 +324,21 @@ (* ------------------------------------------------------------------------ *) 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), - (contX_tac 1), - (rtac contX_Ispair2 1), - (rtac contX2contX_CF1L 1), - (rtac contX_Ispair1 1), + (cont_tac 1), + (rtac cont_Ispair2 1), + (rtac cont2cont_CF1L 1), + (rtac cont_Ispair1 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Ispair2 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), @@ -349,7 +349,7 @@ (rtac beta_cfun_sprod 1) ]); -qed_goalw "inst_sprod_pcpo2" Sprod3.thy [spair_def] "UU = (UU##UU)" +qed_goalw "inst_sprod_pcpo2" Sprod3.thy [spair_def] "UU = (|UU,UU|)" (fn prems => [ (rtac sym 1), @@ -360,7 +360,7 @@ ]); 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), @@ -371,7 +371,7 @@ (etac strict_Ispair 1) ]); -qed_goalw "strict_spair1" Sprod3.thy [spair_def] "(UU##b) = UU" +qed_goalw "strict_spair1" Sprod3.thy [spair_def] "(|UU,b|) = UU" (fn prems => [ (rtac (beta_cfun_sprod RS ssubst) 1), @@ -380,7 +380,7 @@ (rtac strict_Ispair1 1) ]); -qed_goalw "strict_spair2" Sprod3.thy [spair_def] "(a##UU) = UU" +qed_goalw "strict_spair2" Sprod3.thy [spair_def] "(|a,UU|) = UU" (fn prems => [ (rtac (beta_cfun_sprod RS ssubst) 1), @@ -390,7 +390,7 @@ ]); 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), @@ -401,7 +401,7 @@ ]); qed_goalw "defined_spair_rev" Sprod3.thy [spair_def] - "(a##b) = UU ==> (a = UU | b = UU)" + "(|a,b|) = UU ==> (a = UU | b = UU)" (fn prems => [ (cut_facts_tac prems 1), @@ -412,7 +412,7 @@ ]); 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), @@ -423,7 +423,7 @@ ]); 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), @@ -443,7 +443,7 @@ qed_goalw "sprodE" Sprod3.thy [spair_def] -"[|p=UU ==> Q;!!x y. [|p=(x##y);~x=UU ; ~y=UU|] ==> Q|] ==> Q" +"[|p=UU ==> Q;!!x y. [|p=(|x,y|);x~=UU ; y~=UU|] ==> Q|] ==> Q" (fn prems => [ (rtac IsprodE 1), @@ -459,101 +459,101 @@ 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 contX_Isfst 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 contX_Isfst 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 contX_Isfst 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 contX_Issnd 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 contX_Issnd 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 contX_Issnd 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 contX_Isfst 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 contX_Issnd 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 contX_Issnd 1), + (rtac cont_Issnd 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Isfst 1), + (rtac cont_Isfst 1), (rtac defined_IsfstIssnd 1), (rtac (inst_sprod_pcpo RS subst) 1), (atac 1) @@ -561,31 +561,31 @@ 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 contX_Issnd 1), + (rtac cont_Issnd 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Isfst 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< (p1< [ (cut_facts_tac prems 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Issnd 1), + (rtac cont_Issnd 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Issnd 1), + (rtac cont_Issnd 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Isfst 1), + (rtac cont_Isfst 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Isfst 1), + (rtac cont_Isfst 1), (rtac less_sprod3b 1), (rtac (inst_sprod_pcpo RS subst) 1), (atac 1) @@ -593,7 +593,7 @@ qed_goalw "less_sprod5c" Sprod3.thy [sfst_def,ssnd_def,spair_def] - "[|xa##ya<xa<xa< [ (cut_facts_tac prems 1), @@ -606,48 +606,49 @@ 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)])))" +\ (| 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 contX_Issnd 1), + (rtac cont_Issnd 1), (rtac (beta_cfun RS ext RS ssubst) 1), - (rtac contX_Isfst 1), + (rtac cont_Isfst 1), (rtac lub_sprod 1), (resolve_tac prems 1) ]); val thelub_sprod2 = (lub_sprod2 RS thelubI); -(* is_chain(?S1) ==> lub(range(?S1)) = *) -(* (lub(range(%i. sfst[?S1(i)]))## lub(range(%i. ssnd[?S1(i)]))) *) - - +(* + "is_chain ?S1 ==> + lub (range ?S1) = + (|lub (range (%i. sfst`(?S1 i))), lub (range (%i. ssnd`(?S1 i)))|)" : thm +*) qed_goalw "ssplit1" Sprod3.thy [ssplit_def] - "ssplit[f][UU]=UU" + "ssplit`f`UU=UU" (fn prems => [ (rtac (beta_cfun RS ssubst) 1), - (contX_tacR 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), - (contX_tacR 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), - (contX_tacR 1), + (cont_tacR 1), (rtac (sfst2 RS ssubst) 1), (resolve_tac prems 1), (rtac (ssnd2 RS ssubst) 1), @@ -657,11 +658,11 @@ qed_goalw "ssplit3" Sprod3.thy [ssplit_def] - "ssplit[spair][z]=z" + "ssplit`spair`z=z" (fn prems => [ (rtac (beta_cfun RS ssubst) 1), - (contX_tacR 1), + (cont_tacR 1), (res_inst_tac [("Q","z=UU")] classical2 1), (hyp_subst_tac 1), (rtac strictify1 1), @@ -669,7 +670,7 @@ (rtac strictify2 1), (atac 1), (rtac (beta_cfun RS ssubst) 1), - (contX_tacR 1), + (cont_tacR 1), (rtac surjective_pairing_Sprod2 1) ]); diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Sprod3.thy --- a/src/HOLCF/Sprod3.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Sprod3.thy Thu Jun 29 16:28:40 1995 +0200 @@ -16,17 +16,22 @@ ssnd :: "('a**'b)->'b" ssplit :: "('a->'b->'c)->('a**'b)->'c" -syntax "@spair" :: "'a => 'b => ('a**'b)" ("_##_" [101,100] 100) +syntax + "@stuple" :: "['a, args] => 'a ** 'b" ("(1'(|_,/ _|'))") -translations "x##y" == "spair[x][y]" +translations + "(|x, y, z|)" == "(|x, (|y, z|)|)" + "(|x, y|)" == "spair`x`y" rules -inst_sprod_pcpo "(UU::'a**'b) = Ispair(UU,UU)" -spair_def "spair == (LAM x y.Ispair(x,y))" -sfst_def "sfst == (LAM p.Isfst(p))" -ssnd_def "ssnd == (LAM p.Issnd(p))" -ssplit_def "ssplit == (LAM f. strictify[LAM p.f[sfst[p]][ssnd[p]]])" +inst_sprod_pcpo "(UU::'a**'b) = Ispair UU UU" + +defs +spair_def "spair == (LAM x y.Ispair x y)" +sfst_def "sfst == (LAM p.Isfst p)" +ssnd_def "ssnd == (LAM p.Issnd p)" +ssplit_def "ssplit == (LAM f. strictify`(LAM p.f`(sfst`p)`(ssnd`p)))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Ssum0.ML --- a/src/HOLCF/Ssum0.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Ssum0.ML Thu Jun 29 16:28:40 1995 +0200 @@ -30,7 +30,7 @@ (rtac refl 1) ]); -qed_goal "inj_onto_Abs_Ssum" Ssum0.thy "inj_onto(Abs_Ssum,Ssum)" +qed_goal "inj_onto_Abs_Ssum" Ssum0.thy "inj_onto Abs_Ssum Ssum" (fn prems => [ (rtac inj_onto_inverseI 1), @@ -125,7 +125,7 @@ ]); qed_goalw "inject_Sinl_Rep2" Ssum0.thy [Sinl_Rep_def] -"[|~a1=UU ; ~a2=UU ; Sinl_Rep(a1)=Sinl_Rep(a2) |] ==> a1=a2" +"[| 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 @@ -135,7 +135,7 @@ ]); qed_goalw "inject_Sinr_Rep2" Ssum0.thy [Sinr_Rep_def] -"[|~b1=UU ; ~b2=UU ; Sinr_Rep(b1)=Sinr_Rep(b2) |] ==> b1=b2" +"[|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 @@ -201,7 +201,7 @@ ]); qed_goal "inject_Isinl_rev" Ssum0.thy -"~a1=a2 ==> ~Isinl(a1) = Isinl(a2)" +"a1~=a2 ==> Isinl(a1) ~= Isinl(a2)" (fn prems => [ (cut_facts_tac prems 1), @@ -211,7 +211,7 @@ ]); qed_goal "inject_Isinr_rev" Ssum0.thy -"~b1=b2 ==> ~Isinr(b1) = Isinr(b2)" +"b1~=b2 ==> Isinr(b1) ~= Isinr(b2)" (fn prems => [ (cut_facts_tac prems 1), @@ -226,7 +226,7 @@ (* ------------------------------------------------------------------------ *) 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), @@ -272,8 +272,8 @@ qed_goal "IssumE" Ssum0.thy "[|p=Isinl(UU) ==> Q ;\ -\ !!x.[|p=Isinl(x); ~x=UU |] ==> Q;\ -\ !!y.[|p=Isinr(y); ~y=UU |] ==> Q|] ==> Q" +\ !!x.[|p=Isinl(x); x~=UU |] ==> Q;\ +\ !!y.[|p=Isinr(y); y~=UU |] ==> Q|] ==> Q" (fn prems => [ (rtac (Exh_Ssum RS disjE) 1), @@ -307,7 +307,7 @@ (* ------------------------------------------------------------------------ *) 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), @@ -332,7 +332,7 @@ 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), @@ -358,7 +358,7 @@ ]); 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), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Ssum0.thy --- a/src/HOLCF/Ssum0.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Ssum0.thy Thu Jun 29 16:28:40 1995 +0200 @@ -24,16 +24,17 @@ Isinr :: "'b => ('a ++ 'b)" Iwhen :: "('a->'c)=>('b->'c)=>('a ++ 'b)=> 'c" -rules +defs Sinl_Rep_def "Sinl_Rep == (%a.%x y p. - (~a=UU --> x=a & p))" + (a~=UU --> x=a & p))" Sinr_Rep_def "Sinr_Rep == (%b.%x y p. - (~b=UU --> y=b & ~p))" + (b~=UU --> y=b & ~p))" Ssum_def "Ssum =={f.(? a.f=Sinl_Rep(a))|(? b.f=Sinr_Rep(b))}" +rules (*faking a type definition... *) (* "++" is isomorphic to Ssum *) @@ -41,14 +42,14 @@ Rep_Ssum_inverse "Abs_Ssum(Rep_Ssum(p)) = p" Abs_Ssum_inverse "f:Ssum ==> Rep_Ssum(Abs_Ssum(f)) = f" - (*defining the abstract constants*) +defs (*defining the abstract constants*) Isinl_def "Isinl(a) == Abs_Ssum(Sinl_Rep(a))" Isinr_def "Isinr(b) == Abs_Ssum(Sinr_Rep(b))" Iwhen_def "Iwhen(f)(g)(s) == @z. (s=Isinl(UU) --> z=UU) - &(!a. ~a=UU & s=Isinl(a) --> z=f[a]) - &(!b. ~b=UU & s=Isinr(b) --> z=g[b])" + &(!a. a~=UU & s=Isinl(a) --> z=f`a) + &(!b. b~=UU & s=Isinr(b) --> z=g`b)" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Ssum1.ML --- a/src/HOLCF/Ssum1.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Ssum1.ML Thu Jun 29 16:28:40 1995 +0200 @@ -41,7 +41,7 @@ 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)" +"[|s1=Isinl(x::'a); s2=Isinl(y::'a)|] ==> less_ssum s1 s2 = (x << y)" (fn prems => [ (cut_facts_tac prems 1), @@ -82,7 +82,7 @@ val less_ssum1b = prove_goalw Ssum1.thy [less_ssum_def] -"[|s1=Isinr(x::'b); s2=Isinr(y::'b)|] ==> less_ssum(s1,s2) = (x << y)" +"[|s1=Isinr(x::'b); s2=Isinr(y::'b)|] ==> less_ssum s1 s2 = (x << y)" (fn prems => [ (cut_facts_tac prems 1), @@ -124,7 +124,7 @@ 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)" +"[|s1=Isinl(x::'a); s2=Isinr(y::'b)|] ==> less_ssum s1 s2 = ((x::'a) = UU)" (fn prems => [ (cut_facts_tac prems 1), @@ -166,7 +166,7 @@ val less_ssum1d = prove_goalw Ssum1.thy [less_ssum_def] -"[|s1=Isinr(x); s2=Isinl(y)|] ==> less_ssum(s1,s2) = (x = UU)" +"[|s1=Isinr(x); s2=Isinl(y)|] ==> less_ssum s1 s2 = (x = UU)" (fn prems => [ (cut_facts_tac prems 1), @@ -213,7 +213,7 @@ (* ------------------------------------------------------------------------ *) 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), @@ -222,7 +222,7 @@ ]); 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), @@ -231,7 +231,7 @@ ]); 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), @@ -240,7 +240,7 @@ ]); 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), @@ -253,7 +253,7 @@ (* less_ssum is a partial order on ++ *) (* ------------------------------------------------------------------------ *) -qed_goal "refl_less_ssum" Ssum1.thy "less_ssum(p,p)" +qed_goal "refl_less_ssum" Ssum1.thy "less_ssum p p" (fn prems => [ (res_inst_tac [("p","p")] IssumE2 1), @@ -266,7 +266,7 @@ ]); qed_goal "antisym_less_ssum" Ssum1.thy - "[|less_ssum(p1,p2);less_ssum(p2,p1)|] ==> p1=p2" + "[|less_ssum p1 p2; less_ssum p2 p1|] ==> p1=p2" (fn prems => [ (cut_facts_tac prems 1), @@ -296,7 +296,7 @@ ]); qed_goal "trans_less_ssum" Ssum1.thy - "[|less_ssum(p1,p2);less_ssum(p2,p3)|] ==> less_ssum(p1,p3)" + "[|less_ssum p1 p2; less_ssum p2 p3|] ==> less_ssum p1 p3" (fn prems => [ (cut_facts_tac prems 1), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Ssum1.thy --- a/src/HOLCF/Ssum1.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Ssum1.thy Thu Jun 29 16:28:40 1995 +0200 @@ -14,7 +14,7 @@ rules - less_ssum_def "less_ssum(s1,s2) == (@z. + less_ssum_def "less_ssum s1 s2 == (@z. (! u x.s1=Isinl(u) & s2=Isinl(x) --> z = (u << x)) &(! v y.s1=Isinr(v) & s2=Isinr(y) --> z = (v << y)) &(! u y.s1=Isinl(u) & s2=Isinr(y) --> z = (u = UU)) diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Ssum2.ML --- a/src/HOLCF/Ssum2.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Ssum2.ML Thu Jun 29 16:28:40 1995 +0200 @@ -166,7 +166,7 @@ qed_goal "ssum_lemma1" Ssum2.thy -"[|~(!i.? x.Y(i::nat)=Isinl(x))|] ==> (? i.! x.~Y(i)=Isinl(x))" +"[|~(!i.? x.Y(i::nat)=Isinl(x))|] ==> (? i.! x.Y(i)~=Isinl(x))" (fn prems => [ (cut_facts_tac prems 1), @@ -174,8 +174,8 @@ ]); 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)" +"[|(? 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), @@ -271,7 +271,7 @@ (* ------------------------------------------------------------------------ *) qed_goal "ssum_lemma7" Ssum2.thy -"[|Isinl(x) << z; ~x=UU|] ==> ? y.z=Isinl(y) & ~y=UU" +"[|Isinl(x) << z; x~=UU|] ==> ? y.z=Isinl(y) & y~=UU" (fn prems => [ (cut_facts_tac prems 1), @@ -289,7 +289,7 @@ ]); qed_goal "ssum_lemma8" Ssum2.thy -"[|Isinr(x) << z; ~x=UU|] ==> ? y.z=Isinr(y) & ~y=UU" +"[|Isinr(x) << z; x~=UU|] ==> ? y.z=Isinr(y) & y~=UU" (fn prems => [ (cut_facts_tac prems 1), @@ -311,7 +311,7 @@ qed_goal "lub_ssum1a" Ssum2.thy "[|is_chain(Y);(!i.? x.Y(i)=Isinl(x))|] ==>\ \ range(Y) <<|\ -\ Isinl(lub(range(%i.(Iwhen (LAM x.x) (LAM y.UU))(Y(i)))))" +\ Isinl(lub(range(%i.(Iwhen (LAM x.x) (LAM y.UU))(Y i))))" (fn prems => [ (cut_facts_tac prems 1), @@ -352,7 +352,7 @@ qed_goal "lub_ssum1b" Ssum2.thy "[|is_chain(Y);(!i.? x.Y(i)=Isinr(x))|] ==>\ \ range(Y) <<|\ -\ Isinr(lub(range(%i.(Iwhen (LAM y.UU) (LAM x.x))(Y(i)))))" +\ Isinr(lub(range(%i.(Iwhen (LAM y.UU) (LAM x.x))(Y i))))" (fn prems => [ (cut_facts_tac prems 1), @@ -391,12 +391,18 @@ val thelub_ssum1a = lub_ssum1a RS thelubI; -(* [| is_chain(?Y1); ! i. ? x. ?Y1(i) = Isinl(x) |] ==> *) -(* lub(range(?Y1)) = Isinl(lub(range(%i. Iwhen(LAM x. x,LAM y. UU,?Y1(i)))))*) +(* +[| is_chain ?Y1; ! i. ? x. ?Y1 i = Isinl x |] ==> + lub (range ?Y1) = Isinl + (lub (range (%i. Iwhen (LAM x. x) (LAM y. UU) (?Y1 i)))) +*) val thelub_ssum1b = lub_ssum1b RS thelubI; -(* [| is_chain(?Y1); ! i. ? x. ?Y1(i) = Isinr(x) |] ==> *) -(* lub(range(?Y1)) = Isinr(lub(range(%i. Iwhen(LAM y. UU,LAM x. x,?Y1(i)))))*) +(* +[| is_chain ?Y1; ! i. ? x. ?Y1 i = Isinr x |] ==> + lub (range ?Y1) = Isinr + (lub (range (%i. Iwhen (LAM y. UU) (LAM x. x) (?Y1 i)))) +*) qed_goal "cpo_ssum" Ssum2.thy "is_chain(Y::nat=>'a ++'b) ==> ? x.range(Y) <<|x" @@ -412,3 +418,4 @@ (etac lub_ssum1b 1), (atac 1) ]); + diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Ssum3.ML --- a/src/HOLCF/Ssum3.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Ssum3.ML Thu Jun 29 16:28:40 1995 +0200 @@ -76,18 +76,18 @@ (asm_simp_tac Ssum_ss 1) ]); -qed_goal "contX_Isinl" Ssum3.thy "contX(Isinl)" +qed_goal "cont_Isinl" Ssum3.thy "cont(Isinl)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Isinl 1), (rtac contlub_Isinl 1) ]); -qed_goal "contX_Isinr" Ssum3.thy "contX(Isinr)" +qed_goal "cont_Isinr" Ssum3.thy "cont(Isinr)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Isinr 1), (rtac contlub_Isinr 1) ]); @@ -188,7 +188,7 @@ 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))))" +\ Iwhen f g (lub(range Y)) = lub(range(%i. Iwhen f g (Y i)))" (fn prems => [ (cut_facts_tac prems 1), @@ -204,8 +204,8 @@ ]); 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))))" +"[| 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), @@ -263,8 +263,8 @@ 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))))" +"[| 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), @@ -342,26 +342,26 @@ (atac 1) ]); -qed_goal "contX_Iwhen1" Ssum3.thy "contX(Iwhen)" +qed_goal "cont_Iwhen1" Ssum3.thy "cont(Iwhen)" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Iwhen1 1), (rtac contlub_Iwhen1 1) ]); -qed_goal "contX_Iwhen2" Ssum3.thy "contX(Iwhen(f))" +qed_goal "cont_Iwhen2" Ssum3.thy "cont(Iwhen(f))" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Iwhen2 1), (rtac contlub_Iwhen2 1) ]); -qed_goal "contX_Iwhen3" Ssum3.thy "contX(Iwhen(f)(g))" +qed_goal "cont_Iwhen3" Ssum3.thy "cont(Iwhen(f)(g))" (fn prems => [ - (rtac monocontlub2contX 1), + (rtac monocontlub2cont 1), (rtac monofun_Iwhen3 1), (rtac contlub_Iwhen3 1) ]); @@ -370,56 +370,56 @@ (* continuous versions of lemmas for 'a ++ 'b *) (* ------------------------------------------------------------------------ *) -qed_goalw "strict_sinl" Ssum3.thy [sinl_def] "sinl[UU]=UU" +qed_goalw "strict_sinl" Ssum3.thy [sinl_def] "sinl`UU =UU" (fn prems => [ - (simp_tac (Ssum_ss addsimps [contX_Isinl]) 1), + (simp_tac (Ssum_ss addsimps [cont_Isinl]) 1), (rtac (inst_ssum_pcpo RS sym) 1) ]); -qed_goalw "strict_sinr" Ssum3.thy [sinr_def] "sinr[UU]=UU" +qed_goalw "strict_sinr" Ssum3.thy [sinr_def] "sinr`UU=UU" (fn prems => [ - (simp_tac (Ssum_ss addsimps [contX_Isinr]) 1), + (simp_tac (Ssum_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 (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1), - (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1) + (asm_simp_tac (Ssum_ss addsimps [cont_Isinr,cont_Isinl]) 1), + (asm_simp_tac (Ssum_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 (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1), - (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1) + (asm_simp_tac (Ssum_ss addsimps [cont_Isinr,cont_Isinl]) 1), + (asm_simp_tac (Ssum_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 (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1), - (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1) + (asm_simp_tac (Ssum_ss addsimps [cont_Isinr,cont_Isinl]) 1), + (asm_simp_tac (Ssum_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), @@ -430,7 +430,7 @@ ]); qed_goal "defined_sinr" Ssum3.thy - "~x=UU ==> ~sinr[x]=UU" + "x~=UU ==> sinr`x ~= UU" (fn prems => [ (cut_facts_tac prems 1), @@ -441,10 +441,10 @@ ]); 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 (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1), + (asm_simp_tac (Ssum_ss addsimps [cont_Isinr,cont_Isinl]) 1), (rtac (inst_ssum_pcpo RS ssubst) 1), (rtac Exh_Ssum 1) ]); @@ -452,8 +452,8 @@ 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" +\ !!x.[|p=sinl`x; x~=UU |] ==> Q;\ +\ !!y.[|p=sinr`y; y~=UU |] ==> Q|] ==> Q" (fn prems => [ (rtac IssumE 1), @@ -462,165 +462,165 @@ (atac 1), (resolve_tac prems 1), (atac 2), - (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1), + (asm_simp_tac (Ssum_ss addsimps [cont_Isinr,cont_Isinl]) 1), (resolve_tac prems 1), (atac 2), - (asm_simp_tac (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1) + (asm_simp_tac (Ssum_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" + "[|!!x.[|p=sinl`x|] ==> Q;\ +\ !!y.[|p=sinr`y|] ==> Q|] ==> Q" (fn prems => [ (rtac IssumE2 1), (resolve_tac prems 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Isinl 1), + (rtac cont_Isinl 1), (atac 1), (resolve_tac prems 1), (rtac (beta_cfun RS ssubst) 1), - (rtac contX_Isinr 1), + (rtac cont_Isinr 1), (atac 1) ]); -qed_goalw "when1" Ssum3.thy [when_def,sinl_def,sinr_def] - "when[f][g][UU] = UU" +qed_goalw "sswhen1" Ssum3.thy [sswhen_def,sinl_def,sinr_def] + "sswhen`f`g`UU = UU" (fn prems => [ (rtac (inst_ssum_pcpo RS ssubst) 1), (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX2contX_CF1L]) 1)), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont2cont_CF1L]) 1)), (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX2contX_CF1L]) 1)), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont2cont_CF1L]) 1)), (rtac (beta_cfun RS ssubst) 1), - (REPEAT (resolve_tac (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX2contX_CF1L]) 1)), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont2cont_CF1L]) 1)), (simp_tac Ssum_ss 1) ]); -qed_goalw "when2" Ssum3.thy [when_def,sinl_def,sinr_def] - "~x=UU==>when[f][g][sinl[x]] = f[x]" +qed_goalw "sswhen2" Ssum3.thy [sswhen_def,sinl_def,sinr_def] + "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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), (asm_simp_tac Ssum_ss 1) ]); -qed_goalw "when3" Ssum3.thy [when_def,sinl_def,sinr_def] - "~x=UU==>when[f][g][sinr[x]] = g[x]" +qed_goalw "sswhen3" Ssum3.thy [sswhen_def,sinl_def,sinr_def] + "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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 1)), + (REPEAT (resolve_tac (cont_lemmas @ [cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3,cont_Isinl,cont_Isinr,cont2cont_CF1L]) 1)), (asm_simp_tac Ssum_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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (Ssum_ss addsimps [contX_Isinr,contX_Isinl]) 1), + (asm_simp_tac (Ssum_ss addsimps [cont_Isinr,cont_Isinl]) 1), (etac ssum_lemma4 1) ]); -qed_goalw "thelub_ssum2a" Ssum3.thy [sinl_def,sinr_def,when_def] -"[| is_chain(Y); !i.? x. Y(i) = sinl[x] |] ==>\ -\ lub(range(Y)) = sinl[lub(range(%i. when[LAM x. x][LAM y. UU][Y(i)]))]" +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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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), @@ -629,27 +629,27 @@ (rtac exI 1), (etac box_equals 1), (rtac refl 1), - (asm_simp_tac (Ssum_ss addsimps [contX_Isinl]) 1) + (asm_simp_tac (Ssum_ss addsimps [cont_Isinl]) 1) ]); -qed_goalw "thelub_ssum2b" Ssum3.thy [sinl_def,sinr_def,when_def] -"[| is_chain(Y); !i.? x. Y(i) = sinr[x] |] ==>\ -\ lub(range(Y)) = sinr[lub(range(%i. when[LAM y. UU][LAM x. x][Y(i)]))]" +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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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 (contX_lemmas @ [contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3,contX_Isinl,contX_Isinr,contX2contX_CF1L]) 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), @@ -659,42 +659,42 @@ (etac box_equals 1), (rtac refl 1), (asm_simp_tac (Ssum_ss addsimps - [contX_Isinr,contX_Isinl,contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3]) 1) + [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 (Ssum_ss addsimps - [contX_Isinr,contX_Isinl,contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3]) 1), + [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3]) 1), (etac ssum_lemma9 1), (asm_simp_tac (Ssum_ss addsimps - [contX_Isinr,contX_Isinl,contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3]) 1) + [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 (Ssum_ss addsimps - [contX_Isinr,contX_Isinl,contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3]) 1), + [cont_Isinr,cont_Isinl,cont_Iwhen1,cont_Iwhen2, + cont_Iwhen3]) 1), (etac ssum_lemma10 1), (asm_simp_tac (Ssum_ss addsimps - [contX_Isinr,contX_Isinl,contX_Iwhen1,contX_Iwhen2, - contX_Iwhen3]) 1) + [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. when[LAM x. x][LAM y. UU][Y(i)]))]\ -\ | lub(range(Y)) = sinr[lub(range(%i. when[LAM y. UU][LAM x. x][Y(i)]))]" +\ 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), @@ -709,14 +709,14 @@ ]); -qed_goal "when4" Ssum3.thy - "when[sinl][sinr][z]=z" +qed_goal "sswhen4" Ssum3.thy + "sswhen`sinl`sinr`z=z" (fn prems => [ (res_inst_tac [("p","z")] ssumE 1), - (asm_simp_tac (Cfun_ss addsimps [when1,when2,when3]) 1), - (asm_simp_tac (Cfun_ss addsimps [when1,when2,when3]) 1), - (asm_simp_tac (Cfun_ss addsimps [when1,when2,when3]) 1) + (asm_simp_tac (Cfun_ss addsimps [sswhen1,sswhen2,sswhen3]) 1), + (asm_simp_tac (Cfun_ss addsimps [sswhen1,sswhen2,sswhen3]) 1), + (asm_simp_tac (Cfun_ss addsimps [sswhen1,sswhen2,sswhen3]) 1) ]); @@ -724,5 +724,5 @@ (* install simplifier for Ssum *) (* ------------------------------------------------------------------------ *) -val Ssum_rews = [strict_sinl,strict_sinr,when1,when2,when3]; +val Ssum_rews = [strict_sinl,strict_sinr,sswhen1,sswhen2,sswhen3]; val Ssum_ss = Cfun_ss addsimps Ssum_rews; diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Ssum3.thy --- a/src/HOLCF/Ssum3.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Ssum3.thy Thu Jun 29 16:28:40 1995 +0200 @@ -11,19 +11,19 @@ arities "++" :: (pcpo,pcpo)pcpo (* Witness ssum2.ML *) consts - sinl :: "'a -> ('a++'b)" - sinr :: "'b -> ('a++'b)" - when :: "('a->'c)->('b->'c)->('a ++ 'b)-> 'c" + sinl :: "'a -> ('a++'b)" + sinr :: "'b -> ('a++'b)" + sswhen :: "('a->'c)->('b->'c)->('a ++ 'b)-> 'c" rules inst_ssum_pcpo "(UU::'a++'b) = Isinl(UU)" + +defs + sinl_def "sinl == (LAM x.Isinl(x))" sinr_def "sinr == (LAM x.Isinr(x))" -when_def "when == (LAM f g s.Iwhen(f)(g)(s))" +sswhen_def "sswhen == (LAM f g s.Iwhen(f)(g)(s))" end - - - diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Stream.ML --- a/src/HOLCF/Stream.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Stream.ML Thu Jun 29 16:28:40 1995 +0200 @@ -32,9 +32,9 @@ val stream_copy = [ - prover [stream_copy_def] "stream_copy[f][UU]=UU", + 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]]" + "x~=UU ==> stream_copy`f`(scons`x`xs)= scons`x`(f`xs)" ]; val stream_rews = stream_copy @ stream_rews; @@ -44,12 +44,12 @@ (* ------------------------------------------------------------------------*) 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 HOLCF_ss 1), (rtac (stream_rep_iso RS subst) 1), - (res_inst_tac [("p","stream_rep[s]")] sprodE 1), + (res_inst_tac [("p","stream_rep`s")] sprodE 1), (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1), (asm_simp_tac HOLCF_ss 1), (res_inst_tac [("p","y")] liftE1 1), @@ -62,7 +62,7 @@ ]); 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), @@ -88,9 +88,9 @@ val stream_when = [ - prover [stream_when_def] "stream_when[f][UU]=UU", + 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]" + "x~=UU ==> stream_when`f`(scons`x`xs)= f`x`xs" ]; val stream_rews = stream_when @ stream_rews; @@ -106,9 +106,9 @@ ]); 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 @@ -119,9 +119,9 @@ ]); 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" +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; val stream_rews = stream_discsel @ stream_rews; @@ -141,7 +141,7 @@ ]); 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 @@ -151,7 +151,7 @@ ]); val stream_constrdef = [ - prover [scons_def] "scons[UU][xs]=UU" + prover [scons_def] "scons`UU`xs=UU" ] @ stream_constrdef; val stream_rews = stream_constrdef @ stream_rews; @@ -169,16 +169,16 @@ val stream_invert = [ prove_goal Stream.thy "[|x1~=UU; y1~=UU;\ -\ scons[x1][x2] << scons[y1][y2]|] ==> x1<< y1 & x2 << y2" +\ 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), + (dres_inst_tac [("fo5","stream_when`(LAM x l.x)")] monofun_cfun_arg 1), (etac box_less 1), (asm_simp_tac (HOLCF_ss addsimps stream_when) 1), (asm_simp_tac (HOLCF_ss addsimps stream_when) 1), - (dres_inst_tac [("fo5","stream_when[LAM x l.l]")] monofun_cfun_arg 1), + (dres_inst_tac [("fo5","stream_when`(LAM x l.l)")] monofun_cfun_arg 1), (etac box_less 1), (asm_simp_tac (HOLCF_ss addsimps stream_when) 1), (asm_simp_tac (HOLCF_ss addsimps stream_when) 1) @@ -192,16 +192,16 @@ val stream_inject = [ prove_goal Stream.thy "[|x1~=UU; y1~=UU;\ -\ scons[x1][x2] = scons[y1][y2]|] ==> x1= y1 & x2 = y2" +\ 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), + (dres_inst_tac [("f","stream_when`(LAM x l.x)")] cfun_arg_cong 1), (etac box_equals 1), (asm_simp_tac (HOLCF_ss addsimps stream_when) 1), (asm_simp_tac (HOLCF_ss addsimps stream_when) 1), - (dres_inst_tac [("f","stream_when[LAM x l.l]")] cfun_arg_cong 1), + (dres_inst_tac [("f","stream_when`(LAM x l.l)")] cfun_arg_cong 1), (etac box_equals 1), (asm_simp_tac (HOLCF_ss addsimps stream_when) 1), (asm_simp_tac (HOLCF_ss addsimps stream_when) 1) @@ -223,8 +223,8 @@ 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; @@ -236,7 +236,7 @@ val stream_take = [ -prove_goalw Stream.thy [stream_take_def] "stream_take(n)[UU]=UU" +prove_goalw Stream.thy [stream_take_def] "stream_take n`UU = UU" (fn prems => [ (res_inst_tac [("n","n")] natE 1), @@ -244,7 +244,7 @@ (asm_simp_tac iterate_ss 1), (simp_tac (HOLCF_ss addsimps stream_rews) 1) ]), -prove_goalw Stream.thy [stream_take_def] "stream_take(0)[xs]=UU" +prove_goalw Stream.thy [stream_take_def] "stream_take 0`xs=UU" (fn prems => [ (asm_simp_tac iterate_ss 1) @@ -260,7 +260,7 @@ val stream_take = [ prover - "x~=UU ==> stream_take(Suc(n))[scons[x][xs]]=scons[x][stream_take(n)[xs]]" + "x~=UU ==> stream_take (Suc n)`(scons`x`xs) = scons`x`(stream_take n`xs)" ] @ stream_take; val stream_rews = stream_take @ stream_rews; @@ -270,7 +270,7 @@ (* ------------------------------------------------------------------------*) qed_goal "stream_copy2" Stream.thy - "stream_copy[f][scons[x][xs]]= scons[x][f[xs]]" + "stream_copy`f`(scons`x`xs) = scons`x`(f`xs)" (fn prems => [ (res_inst_tac [("Q","x=UU")] classical2 1), @@ -278,7 +278,7 @@ (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1) ]); -qed_goal "shd2" Stream.thy "shd[scons[x][xs]]=x" +qed_goal "shd2" Stream.thy "shd`(scons`x`xs) = x" (fn prems => [ (res_inst_tac [("Q","x=UU")] classical2 1), @@ -287,7 +287,7 @@ ]); qed_goal "stream_take2" Stream.thy - "stream_take(Suc(n))[scons[x][xs]]=scons[x][stream_take(n)[xs]]" + "stream_take (Suc n)`(scons`x`xs) = scons`x`(stream_take n`xs)" (fn prems => [ (res_inst_tac [("Q","x=UU")] classical2 1), @@ -327,10 +327,10 @@ ]); 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" +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), @@ -346,7 +346,7 @@ (* ------------------------------------------------------------------------*) 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]" +"stream_bisim R ==> ! p q. R p q --> stream_take n`p = stream_take n`q" (fn prems => [ (cut_facts_tac prems 1), @@ -365,7 +365,7 @@ (fast_tac HOL_cs 1) ]); -qed_goal "stream_coind" Stream.thy "[|stream_bisim(R);R(p,q)|] ==> p = q" +qed_goal "stream_coind" Stream.thy "[|stream_bisim R ;R p q|] ==> p = q" (fn prems => [ (rtac stream_take_lemma 1), @@ -380,8 +380,8 @@ qed_goal "stream_finite_ind" Stream.thy "[|P(UU);\ -\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons[x][s1])\ -\ |] ==> !s.P(stream_take(n)[s])" +\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ +\ |] ==> !s.P(stream_take n`s)" (fn prems => [ (nat_ind_tac "n" 1), @@ -398,7 +398,7 @@ ]); qed_goalw "stream_finite_ind2" Stream.thy [stream_finite_def] -"(!!n.P(stream_take(n)[s])) ==> stream_finite(s) -->P(s)" +"(!!n.P(stream_take n`s)) ==> stream_finite(s) -->P(s)" (fn prems => [ (strip_tac 1), @@ -409,7 +409,7 @@ qed_goal "stream_finite_ind3" Stream.thy "[|P(UU);\ -\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons[x][s1])\ +\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ \ |] ==> stream_finite(s) --> P(s)" (fn prems => [ @@ -426,7 +426,7 @@ qed_goal "stream_ind" Stream.thy "[|adm(P);\ \ P(UU);\ -\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons[x][s1])\ +\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ \ |] ==> P(s)" (fn prems => [ @@ -446,7 +446,7 @@ qed_goal "stream_ind" Stream.thy "[|adm(P);\ \ P(UU);\ -\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons[x][s1])\ +\ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ \ |] ==> P(s)" (fn prems => [ @@ -456,7 +456,7 @@ (rtac adm_impl_admw 1), (REPEAT (resolve_tac adm_thms 1)), (rtac adm_subst 1), - (contX_tacR 1), + (cont_tacR 1), (resolve_tac prems 1), (rtac allI 1), (rtac (rewrite_rule [stream_take_def] stream_finite_ind) 1), @@ -469,7 +469,7 @@ (* simplify use of Co-induction *) (* ------------------------------------------------------------------------*) -qed_goal "surjectiv_scons" Stream.thy "scons[shd[s]][stl[s]]=s" +qed_goal "surjectiv_scons" Stream.thy "scons`(shd`s)`(stl`s)=s" (fn prems => [ (res_inst_tac [("s","s")] streamE 1), @@ -479,7 +479,7 @@ 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)" +"!s1 s2. R s1 s2 --> shd`s1 = shd`s2 & R (stl`s1) (stl`s2) ==> stream_bisim R" (fn prems => [ (cut_facts_tac prems 1), @@ -495,21 +495,21 @@ (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), + (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 (HOLCF_ss addsimps stream_rews addsimps [surjectiv_scons]) 1), - (eres_inst_tac [("s","shd[s1]"),("t","shd[s2]")] subst 1), + (eres_inst_tac [("s","shd`s1"),("t","shd`s2")] subst 1), (simp_tac (HOLCF_ss 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), + (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 (HOLCF_ss addsimps stream_rews addsimps [surjectiv_scons]) 1), - (res_inst_tac [("s","shd[s1]"),("t","shd[s2]")] ssubst 1), + (res_inst_tac [("s","shd`s1"),("t","shd`s2")] ssubst 1), (etac sym 1), (simp_tac (HOLCF_ss addsimps stream_rews addsimps [surjectiv_scons]) 1) ]); @@ -545,7 +545,7 @@ (* a lemma about shd *) (* ----------------------------------------------------------------------- *) -qed_goal "stream_shd_lemma1" Stream.thy "shd[s]=UU --> s=UU" +qed_goal "stream_shd_lemma1" Stream.thy "shd`s=UU --> s=UU" (fn prems => [ (res_inst_tac [("s","s")] streamE 1), @@ -561,7 +561,7 @@ qed_goal "stream_take_lemma1" Stream.thy "!x xs.x~=UU --> \ -\ stream_take(Suc(n))[scons[x][xs]] = scons[x][xs] --> stream_take(n)[xs]=xs" +\ stream_take (Suc n)`(scons`x`xs) = scons`x`xs --> stream_take n`xs=xs" (fn prems => [ (rtac allI 1), @@ -577,7 +577,7 @@ qed_goal "stream_take_lemma2" Stream.thy - "! s2. stream_take(n)[s2] = s2 --> stream_take(Suc(n))[s2]=s2" + "! s2. stream_take n`s2 = s2 --> stream_take (Suc n)`s2=s2" (fn prems => [ (nat_ind_tac "n" 1), @@ -590,7 +590,7 @@ (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1), (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1), (strip_tac 1 ), - (subgoal_tac "stream_take(n1)[xs] = xs" 1), + (subgoal_tac "stream_take n1`xs = xs" 1), (rtac ((hd stream_inject) RS conjunct2) 2), (atac 4), (atac 2), @@ -601,13 +601,13 @@ 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" +\ stream_take n`(scons`x`xs) = scons`x`xs --> stream_take n`xs=xs" (fn prems => [ (nat_ind_tac "n" 1), (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1), (strip_tac 1 ), - (res_inst_tac [("P","scons[x][xs]=UU")] notE 1), + (res_inst_tac [("P","scons`x`xs=UU")] notE 1), (eresolve_tac stream_constrdef 1), (etac sym 1), (strip_tac 1 ), @@ -620,7 +620,7 @@ qed_goal "stream_take_lemma4" Stream.thy "!x xs.\ -\stream_take(n)[xs]=xs --> stream_take(Suc(n))[scons[x][xs]] = scons[x][xs]" +\stream_take n`xs=xs --> stream_take (Suc n)`(scons`x`xs) = scons`x`xs" (fn prems => [ (nat_ind_tac "n" 1), @@ -631,7 +631,7 @@ (* ---- *) qed_goal "stream_take_lemma5" Stream.thy -"!s. stream_take(n)[s]=s --> iterate(n,stl,s)=UU" +"!s. stream_take n`s=s --> iterate n stl s=UU" (fn prems => [ (nat_ind_tac "n" 1), @@ -652,7 +652,7 @@ ]); qed_goal "stream_take_lemma6" Stream.thy -"!s.iterate(n,stl,s)=UU --> stream_take(n)[s]=s" +"!s.iterate n stl s =UU --> stream_take n`s=s" (fn prems => [ (nat_ind_tac "n" 1), @@ -669,7 +669,7 @@ ]); qed_goal "stream_take_lemma7" Stream.thy -"(iterate(n,stl,s)=UU) = (stream_take(n)[s]=s)" +"(iterate n stl s=UU) = (stream_take n`s=s)" (fn prems => [ (rtac iffI 1), @@ -679,7 +679,7 @@ qed_goal "stream_take_lemma8" Stream.thy -"[|adm(P); !n. ? m. n < m & P(stream_take(m)[s])|] ==> P(s)" +"[|adm(P); !n. ? m. n < m & P (stream_take m`s)|] ==> P(s)" (fn prems => [ (cut_facts_tac prems 1), @@ -697,7 +697,7 @@ (* ----------------------------------------------------------------------- *) qed_goalw "stream_finite_lemma1" Stream.thy [stream_finite_def] - "stream_finite(xs) ==> stream_finite(scons[x][xs])" + "stream_finite(xs) ==> stream_finite(scons`x`xs)" (fn prems => [ (cut_facts_tac prems 1), @@ -707,7 +707,7 @@ ]); qed_goalw "stream_finite_lemma2" Stream.thy [stream_finite_def] - "[|x~=UU; stream_finite(scons[x][xs])|] ==> stream_finite(xs)" + "[|x~=UU; stream_finite(scons`x`xs)|] ==> stream_finite(xs)" (fn prems => [ (cut_facts_tac prems 1), @@ -718,7 +718,7 @@ ]); qed_goal "stream_finite_lemma3" Stream.thy - "x~=UU ==> stream_finite(scons[x][xs]) = stream_finite(xs)" + "x~=UU ==> stream_finite(scons`x`xs) = stream_finite(xs)" (fn prems => [ (cut_facts_tac prems 1), @@ -730,7 +730,7 @@ qed_goalw "stream_finite_lemma5" Stream.thy [stream_finite_def] - "(!n. s1 << s2 --> stream_take(n)[s2] = s2 --> stream_finite(s1))\ + "(!n. s1 << s2 --> stream_take n`s2 = s2 --> stream_finite(s1))\ \=(s1 << s2 --> stream_finite(s2) --> stream_finite(s1))" (fn prems => [ @@ -740,7 +740,7 @@ ]); qed_goal "stream_finite_lemma6" Stream.thy - "!s1 s2. s1 << s2 --> stream_take(n)[s2] = s2 --> stream_finite(s1)" + "!s1 s2. s1 << s2 --> stream_take n`s2 = s2 --> stream_finite(s1)" (fn prems => [ (nat_ind_tac "n" 1), @@ -767,7 +767,7 @@ (strip_tac 1 ), (rtac stream_finite_lemma1 1), (subgoal_tac "xs << xsa" 1), - (subgoal_tac "stream_take(n1)[xsa] = 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), @@ -791,7 +791,7 @@ ]); qed_goalw "stream_finite_lemma8" Stream.thy [stream_finite_def] -"stream_finite(s) = (? n. iterate(n,stl,s)=UU)" +"stream_finite(s) = (? n. iterate n stl s = UU)" (fn prems => [ (simp_tac (HOL_ss addsimps [stream_take_lemma7]) 1) @@ -819,8 +819,8 @@ (* ----------------------------------------------------------------------- *) (* alternative prove for admissibility of ~stream_finite *) -(* show that stream_finite(s) = (? n. iterate(n, stl, s) = UU) *) -(* and prove adm. of ~(? n. iterate(n, stl, s) = UU) *) +(* show that stream_finite(s) = (? n. iterate n stl s = UU) *) +(* and prove adm. of ~(? n. iterate n stl s = UU) *) (* proof uses theorems stream_take_lemma5-7; stream_finite_lemma8 *) (* ----------------------------------------------------------------------- *) @@ -828,10 +828,10 @@ 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), + (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 contX_iterate2 1), + (rtac cont_iterate2 1), (rtac allI 1), (rtac (stream_finite_lemma8 RS ssubst) 1), (fast_tac HOL_cs 1) diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Stream.thy --- a/src/HOLCF/Stream.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Stream.thy Thu Jun 29 16:28:40 1995 +0200 @@ -62,40 +62,41 @@ (* stream_abs is an isomorphism with inverse stream_rep *) (* identity is the least endomorphism on 'a stream *) -stream_abs_iso "stream_rep[stream_abs[x]] = x" -stream_rep_iso "stream_abs[stream_rep[x]] = x" +stream_abs_iso "stream_rep`(stream_abs`x) = x" +stream_rep_iso "stream_abs`(stream_rep`x) = x" stream_copy_def "stream_copy == (LAM f. stream_abs oo - (ssplit[LAM x y. x ## (lift[up oo f])[y]] oo stream_rep))" -stream_reach "(fix[stream_copy])[x]=x" + (ssplit`(LAM x y. (|x , (lift`(up oo f))`y|) )) oo stream_rep)" +stream_reach "(fix`stream_copy)`x = x" +defs (* ----------------------------------------------------------------------- *) (* properties of additional constants *) (* ----------------------------------------------------------------------- *) (* constructors *) -scons_def "scons == (LAM x l. stream_abs[x##up[l]])" +scons_def "scons == (LAM x l. stream_abs`(| x, up`l |))" (* ----------------------------------------------------------------------- *) (* discriminator functional *) stream_when_def -"stream_when == (LAM f l.ssplit[LAM x l.f[x][lift[ID][l]]][stream_rep[l]])" +"stream_when == (LAM f l.ssplit `(LAM x l.f`x`(lift`ID`l)) `(stream_rep`l))" (* ----------------------------------------------------------------------- *) (* discriminators and selectors *) -is_scons_def "is_scons == stream_when[LAM x l.TT]" -shd_def "shd == stream_when[LAM x l.x]" -stl_def "stl == stream_when[LAM x l.l]" +is_scons_def "is_scons == stream_when`(LAM x l.TT)" +shd_def "shd == stream_when`(LAM x l.x)" +stl_def "stl == stream_when`(LAM x l.l)" (* ----------------------------------------------------------------------- *) (* the taker for streams *) -stream_take_def "stream_take == (%n.iterate(n,stream_copy,UU))" +stream_take_def "stream_take == (%n.iterate n stream_copy UU)" (* ----------------------------------------------------------------------- *) -stream_finite_def "stream_finite == (%s.? n.stream_take(n)[s]=s)" +stream_finite_def "stream_finite == (%s.? n.stream_take n `s=s)" (* ----------------------------------------------------------------------- *) (* definition of bisimulation is determined by domain equation *) @@ -103,9 +104,9 @@ stream_bisim_def "stream_bisim == (%R.!s1 s2. - R(s1,s2) --> + R s1 s2 --> ((s1=UU & s2=UU) | - (? x s11 s21. x~=UU & s1=scons[x][s11] & s2 = scons[x][s21] & R(s11,s21))))" + (? x s11 s21. x~=UU & s1=scons`x`s11 & s2 = scons`x`s21 & R s11 s21)))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Stream2.ML --- a/src/HOLCF/Stream2.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Stream2.ML Thu Jun 29 16:28:40 1995 +0200 @@ -12,8 +12,8 @@ (* expand fixed point properties *) (* ------------------------------------------------------------------------- *) -val smap_def2 = fix_prover Stream2.thy smap_def - "smap = (LAM f s. stream_when[LAM x l.scons[f[x]][smap[f][l]]][s])"; +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)"; (* ------------------------------------------------------------------------- *) @@ -21,7 +21,7 @@ (* ------------------------------------------------------------------------- *) -qed_goal "smap1" Stream2.thy "smap[f][UU] = UU" +qed_goal "smap1" Stream2.thy "smap`f`UU = UU" (fn prems => [ (rtac (smap_def2 RS ssubst) 1), @@ -29,7 +29,7 @@ ]); 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), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Stream2.thy --- a/src/HOLCF/Stream2.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Stream2.thy Thu Jun 29 16:28:40 1995 +0200 @@ -12,18 +12,18 @@ smap :: "('a -> 'b) -> 'a stream -> 'b stream" -rules +defs smap_def - "smap = fix[LAM h f s. stream_when[LAM x l.scons[f[x]][h[f][l]]][s]]" + "smap == fix`(LAM h f s. stream_when`(LAM x l.scons `(f`x) `(h`f`l)) `s)" end (* - smap[f][UU] = UU - x~=UU --> smap[f][scons[x][xs]] = scons[f[x]][smap[f][xs]] + smap`f`UU = UU + x~=UU --> smap`f`(scons`x`xs) = scons `(f`x) `(smap`f`xs) *) diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Tr1.ML --- a/src/HOLCF/Tr1.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Tr1.ML Thu Jun 29 16:28:40 1995 +0200 @@ -70,7 +70,7 @@ (resolve_tac dist_less_tr 1) ]); -val dist_eq_tr = map prover ["~TT=UU","~FF=UU","~TT=FF"]; +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); (* ------------------------------------------------------------------------ *) @@ -80,7 +80,7 @@ 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), + (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), @@ -155,12 +155,8 @@ ]); 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 cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Tr1.thy --- a/src/HOLCF/Tr1.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Tr1.thy Thu Jun 29 16:28:40 1995 +0200 @@ -31,22 +31,14 @@ rules - abs_tr_iso "abs_tr[rep_tr[u]] = u" - rep_tr_iso "rep_tr[abs_tr[x]] = x" + abs_tr_iso "abs_tr`(rep_tr`u) = u" + rep_tr_iso "rep_tr`(abs_tr`x) = x" - TT_def "TT == abs_tr[sinl[one]]" - FF_def "FF == abs_tr[sinr[one]]" - - tr_when_def "tr_when == (LAM e1 e2 t.when[LAM x.e1][LAM y.e2][rep_tr[t]])" - -end +defs - - - - + TT_def "TT == abs_tr`(sinl`one)" + FF_def "FF == abs_tr`(sinr`one)" - - - - + tr_when_def "tr_when == + (LAM e1 e2 t. sswhen`(LAM x.e1)`(LAM y.e2)`(rep_tr`t))" +end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Tr2.ML --- a/src/HOLCF/Tr2.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Tr2.ML Thu Jun 29 16:28:40 1995 +0200 @@ -72,9 +72,9 @@ ]); val neg_thms = map prover [ - "neg[TT] = FF", - "neg[FF] = TT", - "neg[UU] = UU" + "neg`TT = FF", + "neg`FF = TT", + "neg`UU = UU" ]; (* ------------------------------------------------------------------------ *) diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Tr2.thy --- a/src/HOLCF/Tr2.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Tr2.thy Thu Jun 29 16:28:40 1995 +0200 @@ -19,20 +19,15 @@ "@andalso" :: "tr => tr => tr" ("_ andalso _" [36,35] 35) "@orelse" :: "tr => tr => tr" ("_ orelse _" [31,30] 30) -translations "x andalso y" == "trand[x][y]" - "x orelse y" == "tror[x][y]" - "If b then e1 else e2 fi" == "Icifte[b][e1][e2]" +translations "x andalso y" == "trand`x`y" + "x orelse y" == "tror`x`y" + "If b then e1 else e2 fi" == "Icifte`b`e1`e2" -rules +defs - ifte_def "Icifte == (LAM t e1 e2.tr_when[e1][e2][t])" - andalso_def "trand == (LAM t1 t2.tr_when[t2][FF][t1])" - orelse_def "tror == (LAM t1 t2.tr_when[TT][t2][t1])" - neg_def "neg == (LAM t. tr_when[FF][TT][t])" + ifte_def "Icifte == (LAM t e1 e2.tr_when`e1`e2`t)" + andalso_def "trand == (LAM t1 t2.tr_when`t2`FF`t1)" + orelse_def "tror == (LAM t1 t2.tr_when`TT`t2`t1)" + neg_def "neg == (LAM t. tr_when`FF`TT`t)" end - - - - - diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Void.ML --- a/src/HOLCF/Void.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Void.ML Thu Jun 29 16:28:40 1995 +0200 @@ -27,14 +27,14 @@ (* less_void is a partial ordering on void *) (* ------------------------------------------------------------------------ *) -qed_goalw "refl_less_void" Void.thy [ less_void_def ] "less_void(x,x)" +qed_goalw "refl_less_void" Void.thy [ less_void_def ] "less_void x x" (fn prems => [ (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), @@ -44,7 +44,7 @@ ]); 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), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/Void.thy --- a/src/HOLCF/Void.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/Void.thy Thu Jun 29 16:28:40 1995 +0200 @@ -23,13 +23,20 @@ UU_void :: "void" less_void :: "[void,void] => bool" -rules +defs (* The unique element in Void is False:bool *) UU_void_Rep_def "UU_void_Rep == False" Void_def "Void == {x. x = UU_void_Rep}" + (*defining the abstract constants*) + + UU_void_def "UU_void == Abs_Void(UU_void_Rep)" + less_void_def "less_void x y == (Rep_Void x = Rep_Void y)" + +rules + (*faking a type definition... *) (* void is isomorphic to Void *) @@ -37,10 +44,6 @@ Rep_Void_inverse "Abs_Void(Rep_Void(x)) = x" Abs_Void_inverse "y:Void ==> Rep_Void(Abs_Void(y)) = y" - (*defining the abstract constants*) - - UU_void_def "UU_void == Abs_Void(UU_void_Rep)" - less_void_def "less_void(x,y) == (Rep_Void(x) = Rep_Void(y))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ccc1.ML --- a/src/HOLCF/ccc1.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ccc1.ML Thu Jun 29 16:28:40 1995 +0200 @@ -14,36 +14,36 @@ (* ------------------------------------------------------------------------ *) -qed_goalw "ID1" ccc1.thy [ID_def] "ID[x]=x" +qed_goalw "ID1" ccc1.thy [ID_def] "ID`x=x" (fn prems => [ (rtac (beta_cfun RS ssubst) 1), - (rtac contX_id 1), + (rtac cont_id 1), (rtac refl 1) ]); -qed_goalw "cfcomp1" ccc1.thy [oo_def] "(f oo g)=(LAM x.f[g[x]])" +qed_goalw "cfcomp1" ccc1.thy [oo_def] "(f oo g)=(LAM x.f`(g`x))" (fn prems => [ (rtac (beta_cfun RS ssubst) 1), - (contX_tacR 1), + (cont_tacR 1), (rtac (beta_cfun RS ssubst) 1), - (contX_tacR 1), + (cont_tacR 1), (rtac refl 1) ]); -qed_goal "cfcomp2" ccc1.thy "(f oo g)[x]=f[g[x]]" +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), - (contX_tacR 1), + (cont_tacR 1), (rtac refl 1) ]); (* ------------------------------------------------------------------------ *) -(* Show that interpretation of (pcpo,_->_) is a ategory *) +(* Show that interpretation of (pcpo,_->_) is a category *) (* The class of objects is interpretation of syntactical class pcpo *) (* The class of arrows between objects 'a and 'b is interpret. of 'a -> 'b *) (* The identity arrow is interpretation of ID *) @@ -74,7 +74,7 @@ (fn prems => [ (rtac ext_cfun 1), - (res_inst_tac [("s","f[g[h[x]]]")] trans 1), + (res_inst_tac [("s","f`(g`(h`x))")] trans 1), (rtac (cfcomp2 RS ssubst) 1), (rtac (cfcomp2 RS ssubst) 1), (rtac refl 1), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ccc1.thy --- a/src/HOLCF/ccc1.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ccc1.thy Thu Jun 29 16:28:40 1995 +0200 @@ -15,12 +15,12 @@ syntax "@oo" :: "('b->'c)=>('a->'b)=>'a->'c" ("_ oo _" [101,100] 100) -translations "f1 oo f2" == "cfcomp[f1][f2]" +translations "f1 oo f2" == "cfcomp`f1`f2" -rules +defs ID_def "ID ==(LAM x.x)" - oo_def "cfcomp == (LAM f g x.f[g[x]])" + oo_def "cfcomp == (LAM f g x.f`(g`x))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ex/Coind.ML --- a/src/HOLCF/ex/Coind.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ex/Coind.ML Thu Jun 29 16:28:40 1995 +0200 @@ -11,11 +11,11 @@ (* ------------------------------------------------------------------------- *) -val nats_def2 = fix_prover Coind.thy nats_def - "nats = scons[dzero][smap[dsucc][nats]]"; +val nats_def2 = fix_prover2 Coind.thy nats_def + "nats = scons`dzero`(smap`dsucc`nats)"; -val from_def2 = fix_prover Coind.thy from_def - "from = (LAM n.scons[n][from[dsucc[n]]])"; +val from_def2 = fix_prover2 Coind.thy from_def + "from = (LAM n.scons`n`(from`(dsucc`n)))"; @@ -24,7 +24,7 @@ (* ------------------------------------------------------------------------- *) -val from = prove_goal Coind.thy "from[n] = scons[n][from[dsucc[n]]]" +val from = prove_goal Coind.thy "from`n = scons`n`(from`(dsucc`n))" (fn prems => [ (rtac trans 1), @@ -34,7 +34,7 @@ ]); -val from1 = prove_goal Coind.thy "from[UU] = UU" +val from1 = prove_goal Coind.thy "from`UU = UU" (fn prems => [ (rtac trans 1), @@ -49,12 +49,12 @@ (* ------------------------------------------------------------------------- *) (* the example *) -(* prove: nats = from[dzero] *) +(* prove: nats = from`dzero *) (* ------------------------------------------------------------------------- *) -val coind_lemma1 = prove_goal Coind.thy "iterator[n][smap[dsucc]][nats] =\ -\ scons[n][iterator[dsucc[n]][smap[dsucc]][nats]]" +val coind_lemma1 = prove_goal Coind.thy "iterator`n`(smap`dsucc)`nats =\ +\ scons`n`(iterator`(dsucc`n)`(smap`dsucc)`nats)" (fn prems => [ (res_inst_tac [("s","n")] dnat_ind 1), @@ -74,11 +74,11 @@ ]); -val nats_eq_from = prove_goal Coind.thy "nats = from[dzero]" +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), +"% 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 (HOLCF_ss addsimps coind_rews) 2), (rewrite_goals_tac [stream_bisim_def]), @@ -91,24 +91,24 @@ (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), + (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), + (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]" +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), +\ iterator`n`(smap`dsucc)`nats & q = from`n")] stream_coind 1), (rtac stream_coind_lemma2 1), (strip_tac 1), (etac exE 1), @@ -122,7 +122,7 @@ (rtac (coind_lemma1 RS ssubst) 1), (rtac (from RS ssubst) 1), (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1), - (res_inst_tac [("x","dsucc[n]")] exI 1), + (res_inst_tac [("x","dsucc`n")] exI 1), (rtac conjI 1), (rtac trans 1), (rtac (coind_lemma1 RS ssubst) 1), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ex/Coind.thy --- a/src/HOLCF/ex/Coind.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ex/Coind.thy Thu Jun 29 16:28:40 1995 +0200 @@ -11,27 +11,23 @@ consts -nats :: "dnat stream" -from :: "dnat -> dnat stream" + nats :: "dnat stream" + from :: "dnat -> dnat stream" -rules +defs + nats_def "nats == fix`(LAM h.scons`dzero`(smap`dsucc`h))" -nats_def "nats = fix[LAM h.scons[dzero][smap[dsucc][h]]]" - -from_def "from = fix[LAM h n.scons[n][h[dsucc[n]]]]" + from_def "from == fix`(LAM h n.scons`n`(h`(dsucc`n)))" end (* - - smap[f][UU] = UU - x~=UU --> smap[f][scons[x][xs]] = scons[f[x]][smap[f][xs]] + smap`f`UU = UU + x~=UU --> smap`f`(scons`x`xs) = scons`(f`x)`(smap`f`xs) - nats = scons[dzero][smap[dsucc][nats]] + nats = scons`dzero`(smap`dsucc`nats) - from[n] = scons[n][from[dsucc[n]]] - - + from`n = scons`n`(from`(dsucc`n)) *) diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ex/Dagstuhl.ML --- a/src/HOLCF/ex/Dagstuhl.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ex/Dagstuhl.ML Thu Jun 29 16:28:40 1995 +0200 @@ -1,52 +1,49 @@ -(* (* $Id$ *) -*) - open Dagstuhl; -val YS_def2 = fix_prover Dagstuhl.thy YS_def "YS = scons[y][YS]"; -val YYS_def2 = fix_prover Dagstuhl.thy YYS_def "YYS = scons[y][scons[y][YYS]]"; +val YS_def2 = fix_prover2 Dagstuhl.thy YS_def "YS = scons`y`YS"; +val YYS_def2 = fix_prover2 Dagstuhl.thy YYS_def "YYS = scons`y`(scons`y`YYS)"; -val prems = goal Dagstuhl.thy "YYS << scons[y][YYS]"; -by (rtac (YYS_def RS ssubst) 1); +val prems = goal Dagstuhl.thy "YYS << scons`y`YYS"; +by (rewrite_goals_tac [YYS_def]); by (rtac fix_ind 1); by (resolve_tac adm_thms 1); -by (contX_tacR 1); +by (cont_tacR 1); by (rtac minimal 1); by (rtac (beta_cfun RS ssubst) 1); -by (contX_tacR 1); +by (cont_tacR 1); by (rtac monofun_cfun_arg 1); by (rtac monofun_cfun_arg 1); by (atac 1); -qed "lemma3"; +val lemma3 = result(); -val prems = goal Dagstuhl.thy "scons[y][YYS] << YYS"; +val prems = goal Dagstuhl.thy "scons`y`YYS << YYS"; by (rtac (YYS_def2 RS ssubst) 1); back(); by (rtac monofun_cfun_arg 1); by (rtac lemma3 1); -qed "lemma4"; +val lemma4=result(); (* val lemma5 = lemma3 RS (lemma4 RS antisym_less) *) -val prems = goal Dagstuhl.thy "scons[y][YYS] = YYS"; +val prems = goal Dagstuhl.thy "scons`y`YYS = YYS"; by (rtac antisym_less 1); by (rtac lemma4 1); by (rtac lemma3 1); -qed "lemma5"; +val lemma5=result(); val prems = goal Dagstuhl.thy "YS = YYS"; by (rtac stream_take_lemma 1); by (nat_ind_tac "n" 1); by (simp_tac (HOLCF_ss addsimps stream_rews) 1); -by (res_inst_tac [("y1","y")] (YS_def2 RS ssubst) 1); -by (res_inst_tac [("y1","y")] (YYS_def2 RS ssubst) 1); +by (rtac (YS_def2 RS ssubst) 1); +by (rtac (YYS_def2 RS ssubst) 1); by (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1); by (rtac (lemma5 RS sym RS subst) 1); by (rtac refl 1); -qed "wir_moel"; +val wir_moel=result(); (* ------------------------------------------------------------------------ *) (* Zweite L"osung: Bernhard M"oller *) @@ -55,95 +52,24 @@ (* ------------------------------------------------------------------------ *) val prems = goal Dagstuhl.thy "YYS << YS"; -by (rtac (YYS_def RS ssubst) 1); +by (rewrite_goals_tac [YYS_def]); by (rtac fix_least 1); by (rtac (beta_cfun RS ssubst) 1); -by (contX_tacR 1); +by (cont_tacR 1); by (simp_tac (HOLCF_ss addsimps [YS_def2 RS sym]) 1); -qed "lemma6"; +val lemma6=result(); val prems = goal Dagstuhl.thy "YS << YYS"; -by (rtac (YS_def RS ssubst) 1); +by (rewrite_goals_tac [YS_def]); by (rtac fix_ind 1); by (resolve_tac adm_thms 1); -by (contX_tacR 1); +by (cont_tacR 1); by (rtac minimal 1); by (rtac (beta_cfun RS ssubst) 1); -by (contX_tacR 1); -by (res_inst_tac [("y2","y10")] (lemma5 RS sym RS ssubst) 1); +by (cont_tacR 1); +by (rtac (lemma5 RS sym RS ssubst) 1); by (etac monofun_cfun_arg 1); -qed "lemma7"; +val lemma7 = result(); val wir_moel = lemma6 RS (lemma7 RS antisym_less); - -(* ------------------------------------------------------------------------ *) -(* L"osung aus Dagstuhl (F.R.) *) -(* Wie oben, jedoch ohne Konstantendefinition f"ur YS, YYS *) -(* ------------------------------------------------------------------------ *) - -val prems = goal Stream2.thy - "(fix[LAM x. scons[y][x]]) = scons[y][fix[LAM x. scons[y][x]]]"; -by (rtac (fix_eq RS ssubst) 1); -back(); -back(); -by (rtac (beta_cfun RS ssubst) 1); -by (contX_tacR 1); -by (rtac refl 1); -qed "lemma1"; - -val prems = goal Stream2.thy - "(fix[ LAM z. scons[y][scons[y][z]]]) = \ -\ scons[y][scons[y][(fix[ LAM z. scons[y][scons[y][z]]])]]"; -by (rtac (fix_eq RS ssubst) 1); -back(); -back(); -by (rtac (beta_cfun RS ssubst) 1); -by (contX_tacR 1); -by (rtac refl 1); -qed "lemma2"; - -val prems = goal Stream2.thy -"fix[LAM z. scons[y][scons[y][z]]] << \ -\ scons[y][fix[LAM z. scons[y][scons[y][z]]]]"; -by (rtac fix_ind 1); -by (resolve_tac adm_thms 1); -by (contX_tacR 1); -by (rtac minimal 1); -by (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1); -by (rtac monofun_cfun_arg 1); -by (rtac monofun_cfun_arg 1); -by (atac 1); -qed "lemma3"; - -val prems = goal Stream2.thy -" scons[y][fix[LAM z. scons[y][scons[y][z]]]] <<\ -\ fix[LAM z. scons[y][scons[y][z]]]"; -by (rtac (lemma2 RS ssubst) 1); -back(); -by (rtac monofun_cfun_arg 1); -by (rtac lemma3 1); -qed "lemma4"; - -val prems = goal Stream2.thy -" scons[y][fix[LAM z. scons[y][scons[y][z]]]] =\ -\ fix[LAM z. scons[y][scons[y][z]]]"; -by (rtac antisym_less 1); -by (rtac lemma4 1); -by (rtac lemma3 1); -qed "lemma5"; - -val prems = goal Stream2.thy - "(fix[LAM x. scons[y][x]]) = (fix[ LAM z. scons[y][scons[y][z]]])"; -by (rtac stream_take_lemma 1); -by (nat_ind_tac "n" 1); -by (simp_tac (HOLCF_ss addsimps stream_rews) 1); -by (rtac (lemma1 RS ssubst) 1); -by (rtac (lemma2 RS ssubst) 1); -by (asm_simp_tac (HOLCF_ss addsimps stream_rews) 1); -by (rtac (lemma5 RS sym RS subst) 1); -by (rtac refl 1); -qed "wir_moel"; - - - diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ex/Dagstuhl.thy --- a/src/HOLCF/ex/Dagstuhl.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ex/Dagstuhl.thy Thu Jun 29 16:28:40 1995 +0200 @@ -4,13 +4,14 @@ Dagstuhl = Stream2 + consts + y :: "'a" YS :: "'a stream" YYS :: "'a stream" -rules +defs -YS_def "YS = fix[LAM x. scons[y][x]]" -YYS_def "YYS = fix[LAM z. scons[y][scons[y][z]]]" +YS_def "YS == fix`(LAM x. scons`y`x)" +YYS_def "YYS == fix`(LAM z. scons`y`(scons`y`z))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ex/Hoare.ML --- a/src/HOLCF/ex/Hoare.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ex/Hoare.ML Thu Jun 29 16:28:40 1995 +0200 @@ -8,7 +8,7 @@ (* --------- pure HOLCF logic, some little lemmas ------ *) -val hoare_lemma2 = prove_goal HOLCF.thy "~b=TT ==> b=FF | b=UU" +val hoare_lemma2 = prove_goal HOLCF.thy "b~=TT ==> b=FF | b=UU" (fn prems => [ (cut_facts_tac prems 1), @@ -20,15 +20,15 @@ ]); val hoare_lemma3 = prove_goal HOLCF.thy -" (!k.b1[iterate(k,g,x)]=TT) | (? k.~ b1[iterate(k,g,x)]=TT)" +" (!k.b1`(iterate k g x) = TT) | (? k. b1`(iterate k g x)~=TT)" (fn prems => [ (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" +"(? 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), @@ -39,9 +39,9 @@ ]); 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" +"[|(? 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), @@ -51,7 +51,7 @@ (etac theleast1 1) ]); -val hoare_lemma6 = prove_goal HOLCF.thy "b=UU ==> ~b=TT" +val hoare_lemma6 = prove_goal HOLCF.thy "b=UU ==> b~=TT" (fn prems => [ (cut_facts_tac prems 1), @@ -59,7 +59,7 @@ (resolve_tac dist_eq_tr 1) ]); -val hoare_lemma7 = prove_goal HOLCF.thy "b=FF ==> ~b=TT" +val hoare_lemma7 = prove_goal HOLCF.thy "b=FF ==> b~=TT" (fn prems => [ (cut_facts_tac prems 1), @@ -68,16 +68,16 @@ ]); 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 b1[iterate(m,g,x)]=TT" +"[|(? 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), + (res_inst_tac [("p","b1`(iterate m g x)")] trE 1), (atac 2), (rtac (le_less_trans RS less_anti_refl) 1), (atac 2), @@ -89,8 +89,9 @@ (etac hoare_lemma7 1) ]); + val hoare_lemma28 = prove_goal HOLCF.thy -"b1[y::'a]=(UU::tr) ==> b1[UU] = UU" +"b1`(y::'a)=(UU::tr) ==> b1`UU = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -103,15 +104,15 @@ (* ----- access to definitions ----- *) val p_def2 = prove_goalw Hoare.thy [p_def] -"p = fix[LAM f x. If b1[x] then f[g[x]] else x fi]" +"p = fix`(LAM f x. If b1`x then f`(g`x) else x fi)" (fn prems => [ (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]" +"q = fix`(LAM f x. If b1`x orelse b2`x then \ +\ f`(g`x) else x fi)" (fn prems => [ (rtac refl 1) @@ -119,7 +120,7 @@ val p_def3 = prove_goal Hoare.thy -"p[x] = If b1[x] then p[g[x]] else x fi" +"p`x = If b1`x then p`(g`x) else x fi" (fn prems => [ (fix_tac3 p_def 1), @@ -127,7 +128,7 @@ ]); val q_def3 = prove_goal Hoare.thy -"q[x] = If b1[x] orelse b2[x] then q[g[x]] else x fi" +"q`x = If b1`x orelse b2`x then q`(g`x) else x fi" (fn prems => [ (fix_tac3 q_def 1), @@ -137,18 +138,18 @@ (** --------- proves about iterations of p and q ---------- **) val hoare_lemma9 = prove_goal Hoare.thy -"(! m. m b1[iterate(m,g,x)]=TT) -->\ -\ p[iterate(k,g,x)]=p[x]" +"(! 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 iterate_ss 1), (simp_tac iterate_ss 1), (strip_tac 1), - (res_inst_tac [("s","p[iterate(k1,g,x)]")] trans 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), + (res_inst_tac [("s","TT"),("t","b1`(iterate k1 g x)")] ssubst 1), (rtac mp 1), (etac spec 1), (simp_tac nat_ss 1), @@ -162,18 +163,18 @@ ]); val hoare_lemma24 = prove_goal Hoare.thy -"(! m. m b1[iterate(m,g,x)]=TT) --> \ -\ q[iterate(k,g,x)]=q[x]" +"(! 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 iterate_ss 1), (simp_tac iterate_ss 1), (strip_tac 1), - (res_inst_tac [("s","q[iterate(k1,g,x)]")] trans 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), + (res_inst_tac [("s","TT"),("t","b1`(iterate k1 g x)")] ssubst 1), (rtac mp 1), (etac spec 1), (simp_tac nat_ss 1), @@ -186,20 +187,21 @@ (simp_tac nat_ss 1) ]); -(* -------- results about p for case (? k.~ b1[iterate(k,g,x)]=TT) ------- *) +(* -------- results about p for case (? k. b1`(iterate k g x)~=TT) ------- *) val hoare_lemma10 = (hoare_lemma8 RS (hoare_lemma9 RS mp)); (* -[| ? k. ~ b1[iterate(k,g,?x1)] = TT; - Suc(?k3) = theleast(%n. ~ b1[iterate(n,g,?x1)] = TT) |] ==> -p[iterate(?k3,g,?x1)] = p[?x1] +val hoare_lemma10 = "[| ? k. b1`(iterate k g ?x1) ~= TT; + Suc ?k3 = theleast (%n. b1`(iterate n g ?x1) ~= TT) |] ==> + p`(iterate ?k3 g ?x1) = p`?x1" : thm + *) val hoare_lemma11 = 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)]=FF \ -\ --> p[x] = iterate(k,g,x)" +"(? n.b1`(iterate n g x) ~= TT) ==>\ +\ 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), @@ -211,7 +213,7 @@ (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)")] + (eres_inst_tac [("s","0"),("t","theleast(%n. b1`(iterate n g x) ~= TT)")] subst 1), (simp_tac iterate_ss 1), (hyp_subst_tac 1), @@ -222,7 +224,7 @@ (atac 1), (rtac trans 1), (rtac p_def3 1), - (res_inst_tac [("s","TT"),("t","b1[iterate(xa,g,x)]")] ssubst 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), @@ -236,9 +238,9 @@ ]); 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" +"(? 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), @@ -260,7 +262,7 @@ (atac 1), (rtac trans 1), (rtac p_def3 1), - (res_inst_tac [("s","TT"),("t","b1[iterate(xa,g,x)]")] ssubst 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), @@ -271,10 +273,10 @@ (asm_simp_tac HOLCF_ss 1) ]); -(* -------- results about p for case (! k. b1[iterate(k,g,x)]=TT) ------- *) +(* -------- 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" +"(! k. b1`(iterate k g x)=TT) ==> !k.p`(iterate k g x) = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -283,13 +285,13 @@ (rtac adm_all 1), (rtac allI 1), (rtac adm_eq 1), - (contX_tacR 1), + (cont_tacR 1), (rtac allI 1), (rtac (strict_fapp1 RS ssubst) 1), (rtac refl 1), (simp_tac iterate_ss 1), (rtac allI 1), - (res_inst_tac [("s","TT"),("t","b1[iterate(k,g,x)]")] ssubst 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), @@ -297,7 +299,7 @@ ]); val hoare_lemma16 = prove_goal Hoare.thy -"(! k. b1[iterate(k,g,x)]=TT) ==> p[x] = UU" +"(! k. b1`(iterate k g x)=TT) ==> p`x = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -305,10 +307,10 @@ (etac (fernpass_lemma RS spec) 1) ]); -(* -------- results about q for case (! k. b1[iterate(k,g,x)]=TT) ------- *) +(* -------- 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" +"(! k. b1`(iterate k g x)=TT) ==> !k.q`(iterate k g x) = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -317,13 +319,13 @@ (rtac adm_all 1), (rtac allI 1), (rtac adm_eq 1), - (contX_tacR 1), + (cont_tacR 1), (rtac allI 1), (rtac (strict_fapp1 RS ssubst) 1), (rtac refl 1), (rtac allI 1), (simp_tac iterate_ss 1), - (res_inst_tac [("s","TT"),("t","b1[iterate(k,g,x)]")] ssubst 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), @@ -331,7 +333,7 @@ ]); val hoare_lemma18 = prove_goal Hoare.thy -"(! k. b1[iterate(k,g,x)]=TT) ==> q[x] = UU" +"(! k. b1`(iterate k g x)=TT) ==> q`x = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -340,7 +342,7 @@ ]); 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)" +"(!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), @@ -350,7 +352,7 @@ ]); val hoare_lemma20 = prove_goal Hoare.thy -"(! y. b1[y::'a]=TT) ==> !k.q[iterate(k,g,x::'a)] = UU" +"(! y. b1`(y::'a)=TT) ==> !k.q`(iterate k g (x::'a)) = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -359,13 +361,13 @@ (rtac adm_all 1), (rtac allI 1), (rtac adm_eq 1), - (contX_tacR 1), + (cont_tacR 1), (rtac allI 1), (rtac (strict_fapp1 RS ssubst) 1), (rtac refl 1), (rtac allI 1), (simp_tac iterate_ss 1), - (res_inst_tac [("s","TT"),("t","b1[iterate(k,g,x::'a)]")] ssubst 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), @@ -373,7 +375,7 @@ ]); val hoare_lemma21 = prove_goal Hoare.thy -"(! y. b1[y::'a]=TT) ==> q[x::'a] = UU" +"(! y. b1`(y::'a)=TT) ==> q`(x::'a) = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -382,7 +384,7 @@ ]); val hoare_lemma22 = prove_goal Hoare.thy -"b1[UU::'a]=UU ==> q[UU::'a] = UU" +"b1`(UU::'a)=UU ==> q`(UU::'a) = UU" (fn prems => [ (cut_facts_tac prems 1), @@ -390,19 +392,19 @@ (asm_simp_tac HOLCF_ss 1) ]); -(* -------- results about q for case (? k.~ b1[iterate(k,g,x)]=TT) ------- *) +(* -------- results about q for case (? k. b1`(iterate k g x) ~= TT) ------- *) val hoare_lemma25 = (hoare_lemma8 RS (hoare_lemma24 RS mp) ); (* -[| ? k. ~ ?b1.1[iterate(k,?g1,?x1)] = TT; - Suc(?k3) = theleast(%n. ~ ?b1.1[iterate(n,?g1,?x1)] = TT) |] ==> -q[iterate(?k3,?g1,?x1)] = q[?x1] +val hoare_lemma25 = "[| ? k. b1`(iterate k g ?x1) ~= TT; + Suc ?k3 = theleast (%n. b1`(iterate n g ?x1) ~= TT) |] ==> + q`(iterate ?k3 g ?x1) = q`?x1" : thm *) val hoare_lemma26 = 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)]=FF \ -\ --> q[x] = q[iterate(k,g,x)]" +"(? n. b1`(iterate n g x)~=TT) ==>\ +\ 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), @@ -419,7 +421,7 @@ (atac 1), (rtac trans 1), (rtac q_def3 1), - (res_inst_tac [("s","TT"),("t","b1[iterate(xa,g,x)]")] ssubst 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), @@ -429,9 +431,9 @@ val hoare_lemma27 = 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 \ -\ --> q[x] = UU" +"(? n. b1`(iterate n g x) ~= TT) ==>\ +\ 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), @@ -452,7 +454,7 @@ (atac 1), (rtac trans 1), (rtac q_def3 1), - (res_inst_tac [("s","TT"),("t","b1[iterate(xa,g,x)]")] ssubst 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), @@ -463,10 +465,10 @@ (asm_simp_tac HOLCF_ss 1) ]); -(* ------- (! k. b1[iterate(k,g,x)]=TT) ==> q o p = q ----- *) +(* ------- (! 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]" +"(! k. b1`(iterate k g x)=TT) ==> q`(p`x) = q`x" (fn prems => [ (cut_facts_tac prems 1), @@ -486,10 +488,10 @@ (rtac refl 1) ]); -(* ------------ ? k. ~ b1[iterate(k,g,x)] = TT ==> q o p = q ----- *) +(* ------------ ? 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]" +"? k. b1`(iterate k g x) ~= TT ==> q`(p`x) = q`x" (fn prems => [ (cut_facts_tac prems 1), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ex/Hoare.thy --- a/src/HOLCF/ex/Hoare.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ex/Hoare.thy Thu Jun 29 16:28:40 1995 +0200 @@ -5,12 +5,12 @@ Theory for an example by C.A.R. Hoare -p x = if b1(x) - then p(g(x)) +p x = if b1 x + then p (g x) else x fi -q x = if b1(x) orelse b2(x) - then q (g(x)) +q x = if b1 x orelse b2 x + then q (g x) else x fi Prove: for all b1 b2 g . @@ -30,14 +30,13 @@ p :: "'a -> 'a" q :: "'a -> 'a" -rules +defs - p_def "p == fix[LAM f. LAM x. - If b1[x] then f[g[x]] else x fi]" + p_def "p == fix`(LAM f. LAM x. + If b1`x then f`(g`x) else x fi)" - q_def "q == fix[LAM f. LAM x. - If b1[x] orelse b2[x] then f[g[x]] else x fi]" - + q_def "q == fix`(LAM f. LAM x. + If b1`x orelse b2`x then f`(g`x) else x fi)" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ex/Loop.ML --- a/src/HOLCF/ex/Loop.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ex/Loop.ML Thu Jun 29 16:28:40 1995 +0200 @@ -13,14 +13,14 @@ (* --------------------------------------------------------------------------- *) val step_def2 = prove_goalw Loop.thy [step_def] -"step[b][g][x] = If b[x] then g[x] else x fi" +"step`b`g`x = If b`x then g`x else x fi" (fn prems => [ (simp_tac Cfun_ss 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]" +"while`b`g = fix`(LAM f x. If b`x then f`(g`x) else x fi)" (fn prems => [ (simp_tac Cfun_ss 1) @@ -32,7 +32,7 @@ (* ------------------------------------------------------------------------- *) val while_unfold = prove_goal Loop.thy -"while[b][g][x] = If b[x] then while[b][g][g[x]] else x fi" +"while`b`g`x = If b`x then while`b`g`(g`x) else x fi" (fn prems => [ (fix_tac5 while_def2 1), @@ -40,7 +40,7 @@ ]); 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), @@ -53,10 +53,10 @@ (rtac trans 1), (etac spec 2), (rtac (step_def2 RS ssubst) 1), - (res_inst_tac [("p","b[x]")] trE 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), + (res_inst_tac [("s","UU"),("t","b`UU")]ssubst 1), (etac (flat_tr RS flat_codom RS disjE) 1), (atac 1), (etac spec 1), @@ -68,10 +68,11 @@ ]); 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), + (res_inst_tac [("s", + "while`b`g`(iterate (Suc 0) (step`b`g) x)")] trans 1), (rtac (while_unfold2 RS spec) 1), (simp_tac iterate_ss 1) ]); @@ -82,7 +83,7 @@ (* --------------------------------------------------------------------------- *) 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" +"[|? 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), @@ -97,8 +98,8 @@ ]); 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" +"[|? 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), @@ -109,9 +110,9 @@ ]); 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))" +"[|!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), @@ -119,15 +120,15 @@ (asm_simp_tac iterate_ss 1), (strip_tac 1), (simp_tac (iterate_ss addsimps [step_def2]) 1), - (res_inst_tac [("p","b[iterate(k1, step[b][g], x)]")] trE 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,iterate_Suc] ) 1), (asm_simp_tac HOLCF_ss 1), (rtac mp 1), (etac spec 1), (asm_simp_tac (HOLCF_ss 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), + (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 [iterate_Suc,step_def2] ) 1), (asm_simp_tac (HOLCF_ss addsimps [loop_lemma2] ) 1) @@ -135,7 +136,7 @@ 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)" +"!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), @@ -152,8 +153,8 @@ ]); 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" +"!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), @@ -161,14 +162,14 @@ (rtac fix_ind 1), (rtac (allI RS adm_all) 1), (rtac adm_eq 1), - (contX_tacR 1), + (cont_tacR 1), (simp_tac HOLCF_ss 1), (rtac allI 1), (simp_tac HOLCF_ss 1), - (res_inst_tac [("p","b[iterate(m, step[b][g],x)]")] trE 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), + (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), @@ -180,7 +181,7 @@ val loop_lemma6 = prove_goal Loop.thy -"!k. ~b[iterate(k,step[b][g],x)]=FF ==> while[b][g][x]=UU" +"!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), @@ -189,7 +190,7 @@ ]); val loop_lemma7 = prove_goal Loop.thy -"~while[b][g][x]=UU ==> ? k. b[iterate(k,step[b][g],x)]=FF" +"while`b`g`x ~= UU ==> ? k. b`(iterate k (step`b`g) x) = FF" (fn prems => [ (cut_facts_tac prems 1), @@ -199,7 +200,7 @@ ]); val loop_lemma8 = prove_goal Loop.thy -"~while[b][g][x]=UU ==> ? y. b[y]=FF" +"while`b`g`x ~= UU ==> ? y. b`y=FF" (fn prems => [ (cut_facts_tac prems 1), @@ -208,17 +209,17 @@ ]); -(* --------------------------------------------------------------------------- *) -(* an invariant rule for loops *) -(* --------------------------------------------------------------------------- *) +(* ------------------------------------------------------------------------- *) +(* an invariant rule for loops *) +(* ------------------------------------------------------------------------- *) val loop_inv2 = 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])" +"[| (!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 => [ - (res_inst_tac [("P","%k.b[iterate(k,step[b][g],x)]=FF")] exE 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), @@ -240,9 +241,9 @@ ]); 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])" +"[| !!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), @@ -263,10 +264,10 @@ val loop_inv = prove_goal Loop.thy "[| P(x);\ -\ !!y.P(y) ==> INV(y);\ -\ !!y.[| INV(y); b[y]=TT; ~g[y]=UU|] ==> INV(g[y]);\ -\ !!y.[| INV(y); b[y]=FF|]==> Q(y);\ -\ ~while[b][g][x]=UU |] ==> Q(while[b][g][x])" +\ !!y.P y ==> INV y;\ +\ !!y.[| INV y; b`y=TT; g`y~=UU|] ==> INV (g`y);\ +\ !!y.[| INV y; b`y=FF|]==> Q y;\ +\ while`b`g`x ~= UU |] ==> Q (while`b`g`x)" (fn prems => [ (rtac loop_inv3 1), diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ex/Loop.thy --- a/src/HOLCF/ex/Loop.thy Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ex/Loop.thy Thu Jun 29 16:28:40 1995 +0200 @@ -13,12 +13,11 @@ step :: "('a -> tr)->('a -> 'a)->'a->'a" while :: "('a -> tr)->('a -> 'a)->'a->'a" -rules +defs - step_def "step == (LAM b g x. If b[x] then g[x] else x fi)" - while_def "while == (LAM b g. fix[LAM f x. - If b[x] then f[g[x]] else x fi])" - + step_def "step == (LAM b g x. If b`x then g`x else x fi)" + while_def "while == (LAM b g. fix`(LAM f x. + If b`x then f`(g`x) else x fi))" end diff -r cbd32a0f2f41 -r 74be52691d62 src/HOLCF/ex/loeckx.ML --- a/src/HOLCF/ex/loeckx.ML Thu Jun 29 16:16:24 1995 +0200 +++ b/src/HOLCF/ex/loeckx.ML Thu Jun 29 16:28:40 1995 +0200 @@ -3,7 +3,7 @@ (* Loeckx & Sieber S.88 *) val prems = goalw Fix.thy [Ifix_def] -"Ifix(F)=lub(range(%i.%G.iterate(i,G,UU)))(F)"; +"Ifix F= lub (range (%i.%G.iterate i G UU)) F"; by (rtac (thelub_fun RS ssubst) 1); by (rtac ch2ch_fun 1); back(); @@ -19,72 +19,76 @@ (* -Idea: %i.%G.iterate(i,G,UU)) is a chain of continuous functions and +Idea: (%i.%G.iterate i G UU) is a chain of continuous functions and Ifix is the lub of this chain. Hence Ifix is continuous. ----- The proof in HOLCF ----------------------- Since we already proved the theorem -val contX_lubcfun = prove_goal Cfun2.thy - "is_chain(F) ==> contX(% x.lub(range(% j.F(j)[x])))" +val cont_lubcfun = prove_goal Cfun2.thy + "is_chain ?F ==> cont (%x. lub (range (%j. ?F j`x)))" -we suffices to prove: +it suffices to prove: -Ifix = (%f.lub(range(%j.(LAM f. iterate(j, f, UU))[f]))) +Ifix = (%f.lub (range (%j. (LAM f. iterate j f UU)`f))) and -contX(%f.lub(range(%j.(LAM f. iterate(j, f, UU))[f]))) +cont (%f.lub (range (%j. (LAM f. iterate j f UU)`f))) Note that if we use the term -%i.%G.iterate(i,G,UU)) instead of (%j.(LAM f. iterate(j, f, UU))) +%i.%G.iterate i G UU instead of (%j.(LAM f. iterate j f UU)) -we cannot use the theorem contX_lubcfun +we cannot use the theorem cont_lubcfun *) -val prems = goal Fix.thy "contX(Ifix)"; -by (res_inst_tac [("t","Ifix"),("s","(%f.lub(range(%j.(LAM f. iterate(j, f, UU))[f])))")] ssubst 1); +val prems = goal Fix.thy "cont(Ifix)"; +by (res_inst_tac + [("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 (subgoal_tac " range(% i.iterate(i, f, UU)) = range(%j.(LAM f. iterate(j, f, UU))[f])" 1); +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); +by (subgoal_tac + "(% 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); -by (rtac contX2contX_CF1L 1); -by (rtac contX_iterate 1); +by (rtac cont2cont_CF1L 1); +by (rtac cont_iterate 1); by (rtac refl 1); -by (rtac contX_lubcfun 1); +by (rtac cont_lubcfun 1); by (rtac is_chainI 1); by (strip_tac 1); by (rtac less_cfun2 1); by (rtac (beta_cfun RS ssubst) 1); -by (rtac contX2contX_CF1L 1); -by (rtac contX_iterate 1); +by (rtac cont2cont_CF1L 1); +by (rtac cont_iterate 1); by (rtac (beta_cfun RS ssubst) 1); -by (rtac contX2contX_CF1L 1); -by (rtac contX_iterate 1); +by (rtac cont2cont_CF1L 1); +by (rtac cont_iterate 1); by (rtac (is_chain_iterate RS is_chainE RS spec) 1); -val contX_Ifix2 = result(); +val cont_Ifix2 = result(); (* the proof in little steps *) val prems = goal Fix.thy -"(% i.iterate(i, f, UU)) = (%j.(LAM f. iterate(j, f, UU))[f])"; +"(% i.iterate i f UU) = (%j.(LAM f. iterate j f UU)`f)"; by (rtac ext 1); by (rtac (beta_cfun RS ssubst) 1); -by (rtac contX2contX_CF1L 1); -by (rtac contX_iterate 1); +by (rtac cont2cont_CF1L 1); +by (rtac cont_iterate 1); by (rtac refl 1); val fix_lemma1 = result(); val prems = goal Fix.thy -" Ifix = (%f.lub(range(%j.(LAM f.iterate(j,f,UU))[f])))"; +" Ifix = (%f.lub(range(%j.(LAM f.iterate j f UU)`f)))"; by (rtac ext 1); by (rewrite_goals_tac [Ifix_def] ); by (rtac (fix_lemma1 RS ssubst) 1); @@ -92,30 +96,23 @@ val fix_lemma2 = result(); (* -- contX_lubcfun; -val it = "is_chain(?F) ==> contX(%x. lub(range(%j. ?F(j)[x])))" : thm +- cont_lubcfun; +val it = "is_chain ?F ==> cont (%x. lub (range (%j. ?F j`x)))" : thm *) -val prems = goal Fix.thy "contX(Ifix)"; +val prems = goal Fix.thy "cont Ifix"; by (rtac ( fix_lemma2 RS ssubst) 1); -by (rtac contX_lubcfun 1); +by (rtac cont_lubcfun 1); by (rtac is_chainI 1); by (strip_tac 1); by (rtac less_cfun2 1); by (rtac (beta_cfun RS ssubst) 1); -by (rtac contX2contX_CF1L 1); -by (rtac contX_iterate 1); +by (rtac cont2cont_CF1L 1); +by (rtac cont_iterate 1); by (rtac (beta_cfun RS ssubst) 1); -by (rtac contX2contX_CF1L 1); -by (rtac contX_iterate 1); +by (rtac cont2cont_CF1L 1); +by (rtac cont_iterate 1); by (rtac (is_chain_iterate RS is_chainE RS spec) 1); -val contX_Ifix2 = result(); +val cont_Ifix2 = result(); - - - - - - -