# HG changeset patch # User oheimb # Date 1094565762 -7200 # Node ID 9d57263faf9edb2867fbc41fb1cd7dc73492d7d8 # Parent 8b74a39dba89e79ba3901d51157c879248e82c46 integrated Streams with ex/Stream.*; added FOCUS/Fstreams.thy diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/FOCUS/Buffer.ML --- a/src/HOLCF/FOCUS/Buffer.ML Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/FOCUS/Buffer.ML Tue Sep 07 16:02:42 2004 +0200 @@ -175,7 +175,7 @@ by (Fast_tac 1); qed_spec_mp "BufAC_Asm_cong_lemma"; Goal "\\f \\ BufEq; ff \\ BufEq; s \\ BufAC_Asm\\ \\ f\\s = ff\\s"; -by (resolve_tac stream.take_lemmas 1); +by (resolve_tac (thms "stream.take_lemmas") 1); by (eatac BufAC_Asm_cong_lemma 2 1); qed "BufAC_Asm_cong"; diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/FOCUS/Buffer_adm.ML --- a/src/HOLCF/FOCUS/Buffer_adm.ML Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/FOCUS/Buffer_adm.ML Tue Sep 07 16:02:42 2004 +0200 @@ -30,7 +30,7 @@ subgoal_tac "!d x. (s = Md d\\\\\\x --> (? y. t = d\\y & (x,y):C)) = \ \ (s = Md d\\\\\\x --> ft\\t = Def d & (x,rt\\t):C)" 1, Asm_simp_tac 1, - auto_tac (claset() addIs [surjectiv_scons RS sym], simpset())]); + auto_tac (claset() addIs [thm "surjectiv_scons" RS sym], simpset())]); val cont_BufAC_Cmt_F = prove_goal thy "down_cont BufAC_Cmt_F" (K [ auto_tac (claset(),simpset() addsimps [down_cont_def,BufAC_Cmt_F_def3])]); @@ -56,7 +56,7 @@ b y res_inst_tac [("x","Suc (Suc 0)")] exI 1; b y rtac conjI 1; b y strip_tac 2; -b y dtac slen_mono 2; +b y dtac (thm "slen_mono") 2; b y datac (thm "ile_trans") 1 2; b y ALLGOALS Force_tac; qed "BufAC_Asm_F_stream_antiP"; diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/FOCUS/FOCUS.ML --- a/src/HOLCF/FOCUS/FOCUS.ML Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/FOCUS/FOCUS.ML Tue Sep 07 16:02:42 2004 +0200 @@ -10,10 +10,10 @@ AddSIs [ex_eqI, ex2_eqI]; Addsimps[eq_UU_symf]; Goal "(#x ~= 0) = (? a y. x = a~> y)"; -by (simp_tac (simpset() addsimps [slen_empty_eq, fstream_exhaust_eq]) 1); +by (simp_tac (simpset() addsimps [thm "slen_empty_eq", fstream_exhaust_eq]) 1); qed "fstream_exhaust_slen_eq"; -Addsimps[slen_less_1_eq, fstream_exhaust_slen_eq, - slen_fscons_eq, slen_fscons_less_eq]; +Addsimps[thm "slen_less_1_eq", fstream_exhaust_slen_eq, + thm "slen_fscons_eq", thm "slen_fscons_less_eq"]; Addsimps[thm "Suc_ile_eq"]; AddEs [strictI]; diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/FOCUS/Fstream.ML --- a/src/HOLCF/FOCUS/Fstream.ML Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/FOCUS/Fstream.ML Tue Sep 07 16:02:42 2004 +0200 @@ -3,6 +3,8 @@ Author: David von Oheimb, TU Muenchen *) +Addsimps[eq_UU_iff RS sym]; + Goal "a = Def d \\ a\\b \\ b = Def d"; by (rtac (flat_eq RS iffD1 RS sym) 1); by (rtac Def_not_UU 1); @@ -19,7 +21,7 @@ qed_goal "fstream_exhaust" thy "x = UU | (? a y. x = a~> y)" (fn _ => [ simp_tac (simpset() addsimps [fscons_def2]) 1, - cut_facts_tac [stream.exhaust] 1, + cut_facts_tac [thm "stream.exhaust"] 1, fast_tac (claset() addDs [not_Undef_is_Def RS iffD1]) 1]); qed_goal "fstream_cases" thy "[| x = UU ==> P; !!a y. x = a~> y ==> P |] ==> P" @@ -34,7 +36,7 @@ THEN hyp_subst_tac i THEN hyp_subst_tac (i+1); qed_goal "fstream_exhaust_eq" thy "(x ~= UU) = (? a y. x = a~> y)" (fn _ => [ - simp_tac(simpset() addsimps [fscons_def2,stream_exhaust_eq]) 1, + simp_tac(simpset() addsimps [fscons_def2,thm "stream_exhaust_eq"]) 1, fast_tac (claset() addDs [not_Undef_is_Def RS iffD1] addEs [DefE]) 1]); @@ -52,17 +54,17 @@ qed_goal "fstream_prefix" thy "a~> s << t ==> ? tt. t = a~> tt & s << tt" (fn prems =>[ cut_facts_tac prems 1, - stream_case_tac "t" 1, + res_inst_tac [("x","t")] (thm "stream.casedist") 1, cut_facts_tac [fscons_not_empty] 1, fast_tac (HOL_cs addDs [eq_UU_iff RS iffD2]) 1, asm_full_simp_tac (HOL_ss addsimps [fscons_def2]) 1, - dtac stream_flat_prefix 1, + dtac (thm "stream_flat_prefix") 1, rtac Def_not_UU 1, fast_tac HOL_cs 1]); qed_goal "fstream_prefix'" thy "x << a~> z = (x = <> | (? y. x = a~> y & y << z))" (fn _ => [ - simp_tac(HOL_ss addsimps [fscons_def2, Def_not_UU RS stream_prefix'])1, + simp_tac(HOL_ss addsimps [fscons_def2, Def_not_UU RS thm "stream_prefix'"])1, Step_tac 1, ALLGOALS(etac swap), fast_tac (HOL_cs addIs [refl_less] addEs [DefE]) 2, @@ -77,12 +79,12 @@ section "ft & rt"; -bind_thm("ft_empty",hd stream.sel_rews); +bind_thm("ft_empty",hd (thms "stream.sel_rews")); qed_goalw "ft_fscons" thy [fscons_def] "ft\\(m~> s) = Def m" (fn _ => [ (Simp_tac 1)]); Addsimps[ft_fscons]; -bind_thm("rt_empty",hd (tl stream.sel_rews)); +bind_thm("rt_empty",hd (tl (thms "stream.sel_rews"))); qed_goalw "rt_fscons" thy [fscons_def] "rt\\(m~> s) = s" (fn _ => [ (Simp_tac 1)]); Addsimps[rt_fscons]; @@ -92,7 +94,7 @@ Safe_tac, etac subst 1, rtac exI 1, - rtac (surjectiv_scons RS sym) 1, + rtac (thm "surjectiv_scons" RS sym) 1, Simp_tac 1]); Addsimps[ft_eq]; @@ -125,12 +127,12 @@ qed_goal "slen_fscons_eq" thy "(Fin (Suc n) < #x) = (? a y. x = a~> y & Fin n < #y)" (fn _ => [ - simp_tac (HOL_ss addsimps [fscons_def2, slen_scons_eq]) 1, + simp_tac (HOL_ss addsimps [fscons_def2, thm "slen_scons_eq"]) 1, fast_tac (HOL_cs addDs [not_Undef_is_Def RS iffD1] addEs [DefE]) 1]); qed_goal "slen_fscons_eq_rev" thy "(#x < Fin (Suc (Suc n))) = (!a y. x ~= a~> y | #y < Fin (Suc n))" (K [ - simp_tac (HOL_ss addsimps [fscons_def2, slen_scons_eq_rev]) 1, + simp_tac (HOL_ss addsimps [fscons_def2, thm "slen_scons_eq_rev"]) 1, step_tac (HOL_cs addSEs [DefE]) 1, step_tac (HOL_cs addSEs [DefE]) 1, step_tac (HOL_cs addSEs [DefE]) 1, @@ -153,7 +155,7 @@ qed_goal "fstream_ind" thy "[| adm P; P <>; !!a s. P s ==> P (a~> s) |] ==> P x" (fn prems => [ cut_facts_tac prems 1, - etac stream.ind 1, + etac (thm "stream.ind") 1, atac 1, fast_tac (HOL_cs addDs [not_Undef_is_Def RS iffD1] addEs (tl (tl prems) RL [fscons_def2 RS subst])) 1]); @@ -161,7 +163,7 @@ qed_goal "fstream_ind2" thy "[| adm P; P UU; !!a. P (a~> UU); !!a b s. P s ==> P (a~> b~> s) |] ==> P x" (fn prems => [ cut_facts_tac prems 1, - etac stream_ind2 1, + etac (thm "stream_ind2") 1, atac 1, fast_tac (HOL_cs addDs [not_Undef_is_Def RS iffD1] addIs ([hd(tl(tl prems))]RL[fscons_def2 RS subst]))1, @@ -175,11 +177,11 @@ section "fsfilter"; qed_goalw "fsfilter_empty" thy [fsfilter_def] "A(C)UU = UU" (fn _ => [ - rtac sfilter_empty 1]); + rtac (thm "sfilter_empty") 1]); qed_goalw "fsfilter_fscons" thy [fsfilter_def] "A(C)x~> xs = (if x:A then x~> (A(C)xs) else A(C)xs)" (fn _ => [ - simp_tac (simpset() addsimps [fscons_def2,sfilter_scons,If_and_if]) 1]); + simp_tac (simpset() addsimps [fscons_def2,thm "sfilter_scons",If_and_if]) 1]); qed_goal "fsfilter_emptys" thy "{}(C)x = UU" (fn _ => [ res_inst_tac [("x","x")] fstream_ind 1, @@ -189,7 +191,7 @@ asm_simp_tac (simpset()addsimps [fsfilter_fscons]) 1]); qed_goal "fsfilter_insert" thy "(insert a A)(C)a~> x = a~> ((insert a A)(C)x)" (fn _=> [ - simp_tac (simpset() addsimps [fsfilter_fscons]) 1]); + simp_tac (simpset() addsimps [thm "fsfilter_fscons"]) 1]); qed_goal "fsfilter_single_in" thy "{a}(C)a~> x = a~> ({a}(C)x)" (fn _=> [ rtac fsfilter_insert 1]); diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/FOCUS/Fstream.thy --- a/src/HOLCF/FOCUS/Fstream.thy Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/FOCUS/Fstream.thy Tue Sep 07 16:02:42 2004 +0200 @@ -3,11 +3,12 @@ Author: David von Oheimb, TU Muenchen FOCUS streams (with lifted elements) +###TODO: integrate Fstreams.thy *) (* FOCUS flat streams *) -Fstream = Streams + +Fstream = Stream + default type diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/FOCUS/Fstreams.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/FOCUS/Fstreams.thy Tue Sep 07 16:02:42 2004 +0200 @@ -0,0 +1,357 @@ +(* Title: HOLCF/FOCUS/Fstreams.thy + ID: $Id$ + Author: Borislav Gajanovic + +FOCUS flat streams (with lifted elements) +###TODO: integrate this with Fstream.* +*) + +theory Fstreams = Stream: + +defaultsort type + +types 'a fstream = "('a lift) stream" + +consts + + fsingleton :: "'a => 'a fstream" ("<_>" [1000] 999) + fsfilter :: "'a set \ 'a fstream \ 'a fstream" + fsmap :: "('a => 'b) => 'a fstream -> 'b fstream" + + jth :: "nat => 'a fstream => 'a" + + first :: "'a fstream => 'a" + last :: "'a fstream => 'a" + + +syntax + + "@emptystream":: "'a fstream" ("<>") + "@fsfilter" :: "'a set \ 'a fstream \ 'a fstream" ("(_'(C')_)" [64,63] 63) + +syntax (xsymbols) + + "@fsfilter" :: "'a set \ 'a fstream \ 'a fstream" ("(_\_)" + [64,63] 63) +translations + + "<>" => "\" + "A(C)s" == "fsfilter A\s" + +defs + + fsingleton_def2: "fsingleton == %a. Def a && UU" + + jth_def: "jth == %n s. if Fin n < #s then THE a. i_th n s = Def a else arbitrary" + + first_def: "first == %s. jth 0 s" + last_def: "last == %s. case #s of Fin n => (if n~=0 then jth (THE k. Suc k = n) s else arbitrary) + | Infty => arbitrary" + + fsfilter_def: "fsfilter A \ sfilter\(flift2 (\x. x\A))" + fsmap_def: "fsmap f == smap$(flift2 f)" + + +lemma ft_fsingleton[simp]: "ft$() = Def a" +by (simp add: fsingleton_def2) + +lemma slen_fsingleton[simp]: "#() = Fin 1" +by (simp add: fsingleton_def2 inat_defs) + +lemma slen_fstreams[simp]: "#( ooo s) = iSuc (#s)" +by (simp add: fsingleton_def2) + +lemma slen_fstreams2[simp]: "#(s ooo ) = iSuc (#s)" +apply (case_tac "#s", auto) +apply (insert slen_sconc [of _ s "Suc 0" ""], auto) +by (simp add: sconc_def) + +lemma j_th_0_fsingleton[simp]:"jth 0 () = a" +apply (simp add: fsingleton_def2 jth_def) +by (simp add: i_th_def Fin_0) + +lemma jth_0[simp]: "jth 0 ( ooo s) = a" +apply (simp add: fsingleton_def2 jth_def) +by (simp add: i_th_def Fin_0) + +lemma first_sconc[simp]: "first ( ooo s) = a" +by (simp add: first_def) + +lemma first_fsingleton[simp]: "first () = a" +by (simp add: first_def) + +lemma jth_n[simp]: "Fin n = #s ==> jth n (s ooo ) = a" +apply (simp add: jth_def, auto) +apply (simp add: i_th_def rt_sconc1) +by (simp add: inat_defs split: inat_splits) + +lemma last_sconc[simp]: "Fin n = #s ==> last (s ooo ) = a" +apply (simp add: last_def) +apply (simp add: inat_defs split:inat_splits) +by (drule sym, auto) + +lemma last_fsingleton[simp]: "last () = a" +by (simp add: last_def) + +lemma first_UU[simp]: "first UU = arbitrary" +by (simp add: first_def jth_def) + +lemma last_UU[simp]:"last UU = arbitrary" +by (simp add: last_def jth_def inat_defs) + +lemma last_infinite[simp]:"#s = Infty ==> last s = arbitrary" +by (simp add: last_def) + +lemma jth_slen_lemma1:"n <= k & Fin n = #s ==> jth k s = arbitrary" +by (simp add: jth_def inat_defs split:inat_splits, auto) + +lemma jth_UU[simp]:"jth n UU = arbitrary" +by (simp add: jth_def) + +lemma ext_last:"[|s ~= UU; Fin (Suc n) = #s|] ==> (stream_take n$s) ooo <(last s)> = s" +apply (simp add: last_def) +apply (case_tac "#s", auto) +apply (simp add: fsingleton_def2) +apply (subgoal_tac "Def (jth n s) = i_th n s") +apply (auto simp add: i_th_last) +apply (drule slen_take_lemma1, auto) +apply (simp add: jth_def) +apply (case_tac "i_th n s = UU") +apply auto +apply (simp add: i_th_def) +apply (case_tac "i_rt n s = UU", auto) +apply (drule i_rt_slen [THEN iffD1]) +apply (drule slen_take_eq_rev [rule_format, THEN iffD2],auto) +by (drule not_Undef_is_Def [THEN iffD1], auto) + + +lemma fsingleton_lemma1[simp]: "( = ) = (a=b)" +by (simp add: fsingleton_def2) + +lemma fsingleton_lemma2[simp]: " ~= <>" +by (simp add: fsingleton_def2) + +lemma fsingleton_sconc:" ooo s = Def a && s" +by (simp add: fsingleton_def2) + +lemma fstreams_ind: + "[| adm P; P <>; !!a s. P s ==> P ( ooo s) |] ==> P x" +apply (simp add: fsingleton_def2) +apply (rule stream.ind, auto) +by (drule not_Undef_is_Def [THEN iffD1], auto) + +lemma fstreams_ind2: + "[| adm P; P <>; !!a. P (); !!a b s. P s ==> P ( ooo ooo s) |] ==> P x" +apply (simp add: fsingleton_def2) +apply (rule stream_ind2, auto) +by (drule not_Undef_is_Def [THEN iffD1], auto)+ + +lemma fstreams_take_Suc[simp]: "stream_take (Suc n)$( ooo s) = ooo stream_take n$s" +by (simp add: fsingleton_def2) + +lemma fstreams_not_empty[simp]: " ooo s ~= <>" +by (simp add: fsingleton_def2) + +lemma fstreams_not_empty2[simp]: "s ooo ~= <>" +by (case_tac "s=UU", auto) + +lemma fstreams_exhaust: "x = UU | (EX a s. x = ooo s)" +apply (simp add: fsingleton_def2, auto) +apply (erule contrapos_pp, auto) +apply (drule stream_exhaust_eq [THEN iffD1], auto) +by (drule not_Undef_is_Def [THEN iffD1], auto) + +lemma fstreams_cases: "[| x = UU ==> P; !!a y. x = ooo y ==> P |] ==> P" +by (insert fstreams_exhaust [of x], auto) + +lemma fstreams_exhaust_eq: "(x ~= UU) = (? a y. x = ooo y)" +apply (simp add: fsingleton_def2, auto) +apply (drule stream_exhaust_eq [THEN iffD1], auto) +by (drule not_Undef_is_Def [THEN iffD1], auto) + +lemma fstreams_inject: "( ooo s = ooo t) = (a=b & s=t)" +by (simp add: fsingleton_def2) + +lemma fstreams_prefix: " ooo s << t ==> EX tt. t = ooo tt & s << tt" +apply (simp add: fsingleton_def2) +apply (insert stream_prefix [of "Def a" s t], auto) +by (drule stream.inverts, auto) + +lemma fstreams_prefix': "x << ooo z = (x = <> | (EX y. x = ooo y & y << z))" +apply (auto, case_tac "x=UU", auto) +apply (drule stream_exhaust_eq [THEN iffD1], auto) +apply (simp add: fsingleton_def2, auto) +apply (drule stream.inverts, auto) +apply (drule ax_flat [rule_format], simp) +apply (drule stream.inverts, auto) +by (erule sconc_mono) + +lemma ft_fstreams[simp]: "ft$( ooo s) = Def a" +by (simp add: fsingleton_def2) + +lemma rt_fstreams[simp]: "rt$( ooo s) = s" +by (simp add: fsingleton_def2) + +lemma ft_eq[simp]: "(ft$s = Def a) = (EX t. s = ooo t)" +apply (rule stream.casedist [of s],auto) +by (drule sym, auto simp add: fsingleton_def2) + +lemma surjective_fstreams: "( ooo y = x) = (ft$x = Def d & rt$x = y)" +by auto + +lemma fstreams_mono: " ooo b << ooo c ==> b << c" +apply (simp add: fsingleton_def2) +by (drule stream.inverts, auto) + +lemma fsmap_UU[simp]: "fsmap f$UU = UU" +by (simp add: fsmap_def) + +lemma fsmap_fsingleton_sconc: "fsmap f$( ooo xs) = <(f x)> ooo (fsmap f$xs)" +by (simp add: fsmap_def fsingleton_def2 flift2_def) + +lemma fsmap_fsingleton[simp]: "fsmap f$() = <(f x)>" +by (simp add: fsmap_def fsingleton_def2 flift2_def) + + +declare range_composition[simp del] + + +lemma fstreams_chain_lemma[rule_format]: + "ALL s x y. stream_take n$(s::'a fstream) << x & x << y & y << s & x ~= y --> stream_take (Suc n)$s << y" +apply (induct_tac n, auto) +apply (case_tac "s=UU", auto) +apply (drule stream_exhaust_eq [THEN iffD1], auto) +apply (case_tac "y=UU", auto) +apply (drule stream_exhaust_eq [THEN iffD1], auto) +apply (drule stream.inverts, auto) +apply (simp add: less_lift_def, auto) +apply (rule monofun_cfun, auto) +apply (case_tac "s=UU", auto) +apply (drule stream_exhaust_eq [THEN iffD1], auto) +apply (erule_tac x="ya" in allE) +apply (drule stream_prefix, auto) +apply (case_tac "y=UU",auto) +apply (drule stream_exhaust_eq [THEN iffD1], clarsimp) +apply (drule stream.inverts, auto) +apply (simp add: less_lift_def) +apply (drule stream.inverts, auto)+ +apply (erule_tac x="tt" in allE) +apply (erule_tac x="yb" in allE, auto) +apply (simp add: less_lift_def) +by (rule monofun_cfun, auto) + +lemma fstreams_lub_lemma1: "[| chain Y; lub (range Y) = ooo s |] ==> EX j t. Y j = ooo t" +apply (subgoal_tac "lub(range Y) ~= UU") +apply (drule chain_UU_I_inverse2, auto) +apply (drule_tac x="i" in is_ub_thelub, auto) +by (drule fstreams_prefix' [THEN iffD1], auto) + +lemma fstreams_lub1: + "[| chain Y; lub (range Y) = ooo s |] + ==> (EX j t. Y j = ooo t) & (EX X. chain X & (ALL i. EX j. ooo X i << Y j) & lub (range X) = s)" +apply (auto simp add: fstreams_lub_lemma1) +apply (rule_tac x="%n. stream_take n$s" in exI, auto) +apply (simp add: chain_stream_take) +apply (induct_tac i, auto) +apply (drule fstreams_lub_lemma1, auto) +apply (rule_tac x="j" in exI, auto) +apply (case_tac "max_in_chain j Y") +apply (frule lub_finch1 [THEN thelubI], auto) +apply (rule_tac x="j" in exI) +apply (erule subst) back back +apply (simp add: less_cprod_def sconc_mono) +apply (simp add: max_in_chain_def, auto) +apply (rule_tac x="ja" in exI) +apply (subgoal_tac "Y j << Y ja") +apply (drule fstreams_prefix, auto)+ +apply (rule sconc_mono) +apply (rule fstreams_chain_lemma, auto) +apply (subgoal_tac "Y ja << (LUB i. (Y i))", clarsimp) +apply (drule fstreams_mono, simp) +apply (rule is_ub_thelub, simp) +apply (blast intro: chain_mono3) +by (rule stream_reach2) + + +lemma lub_Pair_not_UU_lemma: + "[| chain Y; lub (range Y) = ((a::'a::flat), b); a ~= UU; b ~= UU |] + ==> EX j c d. Y j = (c, d) & c ~= UU & d ~= UU" +apply (frule thelub_cprod, clarsimp) +apply (drule chain_UU_I_inverse2, clarsimp) +apply (case_tac "Y i", clarsimp) +apply (case_tac "max_in_chain i Y") +apply (drule maxinch_is_thelub, auto) +apply (rule_tac x="i" in exI, auto) +apply (simp add: max_in_chain_def, auto) +apply (subgoal_tac "Y i << Y j",auto) +apply (simp add: less_cprod_def, clarsimp) +apply (drule ax_flat [rule_format], auto) +apply (case_tac "snd (Y j) = UU",auto) +apply (case_tac "Y j", auto) +apply (rule_tac x="j" in exI) +apply (case_tac "Y j",auto) +by (drule chain_mono3, auto) + +lemma fstreams_lub_lemma2: + "[| chain Y; lub (range Y) = (a, ooo ms); (a::'a::flat) ~= UU |] ==> EX j t. Y j = (a, ooo t)" +apply (frule lub_Pair_not_UU_lemma, auto) +apply (drule_tac x="j" in is_ub_thelub, auto) +apply (simp add: less_cprod_def, clarsimp) +apply (drule ax_flat [rule_format], clarsimp) +by (drule fstreams_prefix' [THEN iffD1], auto) + +lemma fstreams_lub2: + "[| chain Y; lub (range Y) = (a, ooo ms); (a::'a::flat) ~= UU |] + ==> (EX j t. Y j = (a, ooo t)) & (EX X. chain X & (ALL i. EX j. (a, ooo X i) << Y j) & lub (range X) = ms)" +apply (auto simp add: fstreams_lub_lemma2) +apply (rule_tac x="%n. stream_take n$ms" in exI, auto) +apply (simp add: chain_stream_take) +apply (induct_tac i, auto) +apply (drule fstreams_lub_lemma2, auto) +apply (rule_tac x="j" in exI, auto) +apply (simp add: less_cprod_def) +apply (case_tac "max_in_chain j Y") +apply (frule lub_finch1 [THEN thelubI], auto) +apply (rule_tac x="j" in exI) +apply (erule subst) back back +apply (simp add: less_cprod_def sconc_mono) +apply (simp add: max_in_chain_def, auto) +apply (rule_tac x="ja" in exI) +apply (subgoal_tac "Y j << Y ja") +apply (simp add: less_cprod_def, auto) +apply (drule trans_less) +apply (simp add: ax_flat, auto) +apply (drule fstreams_prefix, auto)+ +apply (rule sconc_mono) +apply (subgoal_tac "tt ~= tta" "tta << ms") +apply (blast intro: fstreams_chain_lemma) +apply (frule lub_cprod [THEN thelubI], auto) +apply (subgoal_tac "snd (Y ja) << (LUB i. snd (Y i))", clarsimp) +apply (drule fstreams_mono, simp) +apply (rule is_ub_thelub chainI) +apply (simp add: chain_def less_cprod_def) +apply (subgoal_tac "fst (Y j) ~= fst (Y ja) | snd (Y j) ~= snd (Y ja)", simp) +apply (drule ax_flat[rule_format], simp)+ +apply (drule prod_eqI, auto) +apply (simp add: chain_mono3) +by (rule stream_reach2) + + +lemma cpo_cont_lemma: + "[| monofun (f::'a::cpo => 'b::cpo); (!Y. chain Y --> f (lub(range Y)) << (LUB i. f (Y i))) |] ==> cont f" +apply (rule monocontlub2cont, auto) +apply (simp add: contlub, auto) +apply (erule_tac x="Y" in allE, auto) +apply (simp add: po_eq_conv) +apply (frule cpo,auto) +apply (frule is_lubD1) +apply (frule ub2ub_monofun, auto) +apply (drule thelubI, auto) +apply (rule is_lub_thelub, auto) +by (erule ch2ch_monofun, simp) + +end + + + + diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/FOCUS/ROOT.ML --- a/src/HOLCF/FOCUS/ROOT.ML Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/FOCUS/ROOT.ML Tue Sep 07 16:02:42 2004 +0200 @@ -9,5 +9,6 @@ val banner = "HOLCF/FOCUS"; writeln banner; +use_thy "Fstreams"; use_thy "FOCUS"; use_thy "Buffer_adm"; diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/FOCUS/Stream_adm.ML --- a/src/HOLCF/FOCUS/Stream_adm.ML Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/FOCUS/Stream_adm.ML Tue Sep 07 16:02:42 2004 +0200 @@ -24,12 +24,12 @@ by ( dtac spec 1); by ( Safe_tac); by ( rtac exI 1); -by ( rtac slen_strict_mono 1); +by ( rtac (thm "slen_strict_mono") 1); by ( etac spec 1); by ( atac 1); by ( atac 1); by (dtac (not_ex RS iffD1) 1); -by (stac slen_infinite 1); +by (stac (thm "slen_infinite") 1); by (etac thin_rl 1); by (dtac spec 1); by (fold_goals_tac [thm "ile_def"]); @@ -67,16 +67,16 @@ by (dres_inst_tac [("P","%x. x")] subst 1 THEN atac 1); by (etac allE 1 THEN dtac mp 1 THEN rtac ile_lemma 1); by ( etac (thm "ile_trans") 1); -by ( etac slen_mono 1); +by ( etac (thm "slen_mono") 1); by (etac ssubst 1); by (safe_tac HOL_cs); -by ( eatac (ile_lemma RS slen_take_lemma3 RS subst) 2 1); +by ( eatac (ile_lemma RS thm "slen_take_lemma3" RS subst) 2 1); by (etac allE 1); by (etac allE 1); by (dtac mp 1); -by ( etac slen_rt_mult 1); +by ( etac (thm "slen_rt_mult") 1); by (dtac mp 1); -by (etac monofun_rt_mult 1); +by (etac (thm "monofun_rt_mult") 1); by (mp_tac 1); by (atac 1); qed "stream_monoP2I"; @@ -123,9 +123,9 @@ etac allE 1 THEN mp_tac 1, dres_inst_tac [("P","%x. x")] subst 1 THEN atac 1, etac conjE 1 THEN rtac conjI 1, - etac (slen_take_lemma3 RS ssubst) 1 THEN atac 1, + etac (thm "slen_take_lemma3" RS ssubst) 1 THEN atac 1, atac 1, - etac allE 1 THEN etac allE 1 THEN dtac mp 1 THEN etac monofun_rt_mult 1, + etac allE 1 THEN etac allE 1 THEN dtac mp 1 THEN etac (thm "monofun_rt_mult") 1, mp_tac 1, atac 1]); diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/IsaMakefile --- a/src/HOLCF/IsaMakefile Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/IsaMakefile Tue Sep 07 16:02:42 2004 +0200 @@ -39,7 +39,7 @@ Up1.thy Up2.ML Up2.thy Up3.ML Up3.thy adm.ML cont_consts.ML \ domain/axioms.ML domain/extender.ML domain/interface.ML \ domain/library.ML domain/syntax.ML domain/theorems.ML holcf_logic.ML \ - ex/Stream.thy ex/Stream.ML Streams.thy + ex/Stream.thy @$(ISATOOL) usedir -b -r $(OUT)/HOL HOLCF @@ -68,7 +68,8 @@ HOLCF-FOCUS: HOLCF $(LOG)/HOLCF-FOCUS.gz -$(LOG)/HOLCF-FOCUS.gz: $(OUT)/HOLCF FOCUS/Fstream.thy FOCUS/Fstream.ML \ +$(LOG)/HOLCF-FOCUS.gz: $(OUT)/HOLCF FOCUS/Fstreams.thy \ + FOCUS/Fstream.thy FOCUS/Fstream.ML \ FOCUS/FOCUS.thy FOCUS/FOCUS.ML \ FOCUS/Stream_adm.thy FOCUS/Stream_adm.ML ../HOL/Library/Continuity.thy \ FOCUS/Buffer.thy FOCUS/Buffer.ML FOCUS/Buffer_adm.thy FOCUS/Buffer_adm.ML diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/ROOT.ML --- a/src/HOLCF/ROOT.ML Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/ROOT.ML Tue Sep 07 16:02:42 2004 +0200 @@ -24,6 +24,5 @@ use "domain/interface.ML"; path_add "~~/src/HOLCF/ex"; -use_thy "Streams"; print_depth 10; diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/Streams.thy --- a/src/HOLCF/Streams.thy Tue Sep 07 15:59:16 2004 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,564 +0,0 @@ -(* Title: HOLCF/Streams.thy - ID: $Id$ - Author: Borislav Gajanovic and David von Oheimb - -Stream domains with concatenation. -TODO: HOLCF/ex/Stream.* should be integrated into this file. -*) - -theory Streams = Stream : - -(* ----------------------------------------------------------------------- *) - -lemma stream_neq_UU: "x~=UU ==> EX a as. x=a&&as & a~=UU" -by (simp add: stream_exhaust_eq,auto) - -lemma stream_prefix1: "[| x< x&&xs << y&&ys" -by (insert stream_prefix' [of y "x&&xs" ys],force) - -lemma stream_take_le_mono : "k<=n ==> stream_take k$s1 << stream_take n$s1" -apply (insert chain_stream_take [of s1]) -by (drule chain_mono3,auto) - -lemma mono_stream_take: "s1 << s2 ==> stream_take n$s1 << stream_take n$s2" -by (simp add: monofun_cfun_arg) - -lemma stream_take_prefix [simp]: "stream_take n$s << s" -apply (subgoal_tac "s=(LUB n. stream_take n$s)") - apply (erule ssubst, rule is_ub_thelub) - apply (simp only: chain_stream_take) -by (simp only: stream_reach2) - -lemma stream_take_take_less:"stream_take k$(stream_take n$s) << stream_take k$s" -by (rule monofun_cfun_arg,auto) - -(* ----------------------------------------------------------------------- *) - -lemma slen_rt_mono: "#s2 <= #s1 ==> #(rt$s2) <= #(rt$s1)" -apply (rule stream.casedist [of s1]) - apply (rule stream.casedist [of s2],simp+) -by (rule stream.casedist [of s2],auto) - -lemma slen_take_lemma4 [rule_format]: - "!s. stream_take n$s ~= s --> #(stream_take n$s) = Fin n" -apply (induct_tac n,auto simp add: Fin_0) -apply (case_tac "s=UU",simp) -by (drule stream_neq_UU,auto) - -lemma slen_take_lemma5: "#(stream_take n$s) <= Fin n"; -apply (case_tac "stream_take n$s = s") - apply (simp add: slen_take_eq_rev) -by (simp add: slen_take_lemma4) - -lemma stream_take_idempotent [simp]: - "stream_take n$(stream_take n$s) = stream_take n$s" -apply (case_tac "stream_take n$s = s") -apply (auto,insert slen_take_lemma4 [of n s]); -by (auto,insert slen_take_lemma1 [of "stream_take n$s" n],simp) - -lemma stream_take_take_Suc [simp]: "stream_take n$(stream_take (Suc n)$s) = - stream_take n$s" -apply (simp add: po_eq_conv,auto) - apply (simp add: stream_take_take_less) -apply (subgoal_tac "stream_take n$s = stream_take n$(stream_take n$s)") - apply (erule ssubst) - apply (rule_tac monofun_cfun_arg) - apply (insert chain_stream_take [of s]) -by (simp add: chain_def,simp) - -lemma mono_stream_take_pred: - "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==> - stream_take n$s1 << stream_take n$s2" -by (drule mono_stream_take [of _ _ n],simp) - -lemma stream_take_lemma10 [rule_format]: - "ALL k<=n. stream_take n$s1 << stream_take n$s2 - --> stream_take k$s1 << stream_take k$s2" -apply (induct_tac n,simp,clarsimp) -apply (case_tac "k=Suc n",blast) -apply (erule_tac x="k" in allE) -by (drule mono_stream_take_pred,simp) - -lemma stream_take_finite [simp]: "stream_finite (stream_take n$s)" -apply (simp add: stream.finite_def) -by (rule_tac x="n" in exI,simp) - -lemma slen_stream_take_finite [simp]: "#(stream_take n$s) ~= \" -by (simp add: slen_def) - -lemma stream_take_Suc_neq: "stream_take (Suc n)$s ~=s ==> - stream_take n$s ~= stream_take (Suc n)$s" -apply auto -apply (subgoal_tac "stream_take n$s ~=s") - apply (insert slen_take_lemma4 [of n s],auto) -apply (rule stream.casedist [of s],simp) -apply (simp add: inat_defs split:inat_splits) -by (simp add: slen_take_lemma4) - - -(* ----------------------------------------------------------------------- *) - -consts - - i_rt :: "nat => 'a stream => 'a stream" (* chops the first i elements *) - i_th :: "nat => 'a stream => 'a" (* the i-th element *) - - sconc :: "'a stream => 'a stream => 'a stream" (infixr "ooo" 65) - constr_sconc :: "'a stream => 'a stream => 'a stream" (* constructive *) - constr_sconc' :: "nat => 'a stream => 'a stream => 'a stream" - -defs - i_rt_def: "i_rt == %i s. iterate i rt s" - i_th_def: "i_th == %i s. ft$(i_rt i s)" - - sconc_def: "s1 ooo s2 == case #s1 of - Fin n => (SOME s. (stream_take n$s=s1) & (i_rt n s = s2)) - | \ => s1" - - constr_sconc_def: "constr_sconc s1 s2 == case #s1 of - Fin n => constr_sconc' n s1 s2 - | \ => s1" -primrec - constr_sconc'_0: "constr_sconc' 0 s1 s2 = s2" - constr_sconc'_Suc: "constr_sconc' (Suc n) s1 s2 = ft$s1 && - constr_sconc' n (rt$s1) s2" - - -(* ----------------------------------------------------------------------- *) - section "i_rt" -(* ----------------------------------------------------------------------- *) - -lemma i_rt_UU [simp]: "i_rt n UU = UU" -apply (simp add: i_rt_def) -by (rule iterate.induct,auto) - -lemma i_rt_0 [simp]: "i_rt 0 s = s" -by (simp add: i_rt_def) - -lemma i_rt_Suc [simp]: "a ~= UU ==> i_rt (Suc n) (a&&s) = i_rt n s" -by (simp add: i_rt_def iterate_Suc2 del: iterate_Suc) - -lemma i_rt_Suc_forw: "i_rt (Suc n) s = i_rt n (rt$s)" -by (simp only: i_rt_def iterate_Suc2) - -lemma i_rt_Suc_back:"i_rt (Suc n) s = rt$(i_rt n s)" -by (simp only: i_rt_def,auto) - -lemma i_rt_mono: "x << s ==> i_rt n x << i_rt n s" -by (simp add: i_rt_def monofun_rt_mult) - -lemma i_rt_ij_lemma: "Fin (i + j) <= #x ==> Fin j <= #(i_rt i x)" -by (simp add: i_rt_def slen_rt_mult) - -lemma slen_i_rt_mono: "#s2 <= #s1 ==> #(i_rt n s2) <= #(i_rt n s1)" -apply (induct_tac n,auto) -apply (simp add: i_rt_Suc_back) -by (drule slen_rt_mono,simp) - -lemma i_rt_take_lemma1 [rule_format]: "ALL s. i_rt n (stream_take n$s) = UU" -apply (induct_tac n); - apply (simp add: i_rt_Suc_back,auto) -apply (case_tac "s=UU",auto) -by (drule stream_neq_UU,simp add: i_rt_Suc_forw,auto) - -lemma i_rt_slen: "(i_rt n s = UU) = (stream_take n$s = s)" -apply auto - apply (insert i_rt_ij_lemma [of n "Suc 0" s]); - apply (subgoal_tac "#(i_rt n s)=0") - apply (case_tac "stream_take n$s = s",simp+) - apply (insert slen_take_eq [of n s],simp) - apply (simp add: inat_defs split:inat_splits) - apply (simp add: slen_take_eq ) -by (simp, insert i_rt_take_lemma1 [of n s],simp) - -lemma i_rt_lemma_slen: "#s=Fin n ==> i_rt n s = UU" -by (simp add: i_rt_slen slen_take_lemma1) - -lemma stream_finite_i_rt [simp]: "stream_finite (i_rt n s) = stream_finite s" -apply (induct_tac n, auto) - apply (rule stream.casedist [of "s"], auto simp del: i_rt_Suc) -by (simp add: i_rt_Suc_back stream_finite_rt_eq)+ - -lemma take_i_rt_len_lemma: "ALL sl x j t. Fin sl = #x & n <= sl & - #(stream_take n$x) = Fin t & #(i_rt n x)= Fin j - --> Fin (j + t) = #x" -apply (induct_tac n,auto) - apply (simp add: inat_defs) -apply (case_tac "x=UU",auto) - apply (simp add: inat_defs) -apply (drule stream_neq_UU,auto) -apply (subgoal_tac "EX k. Fin k = #as",clarify) - apply (erule_tac x="k" in allE) - apply (erule_tac x="as" in allE,auto) - apply (erule_tac x="THE p. Suc p = t" in allE,auto) - apply (simp add: inat_defs split:inat_splits) - apply (simp add: inat_defs split:inat_splits) - apply (simp only: the_equality) - apply (simp add: inat_defs split:inat_splits) - apply force -by (simp add: inat_defs split:inat_splits) - -lemma take_i_rt_len: -"[| Fin sl = #x; n <= sl; #(stream_take n$x) = Fin t; #(i_rt n x) = Fin j |] ==> - Fin (j + t) = #x" -by (blast intro: take_i_rt_len_lemma [rule_format]) - - -(* ----------------------------------------------------------------------- *) - section "i_th" -(* ----------------------------------------------------------------------- *) - -lemma i_th_i_rt_step: -"[| i_th n s1 << i_th n s2; i_rt (Suc n) s1 << i_rt (Suc n) s2 |] ==> - i_rt n s1 << i_rt n s2" -apply (simp add: i_th_def i_rt_Suc_back) -apply (rule stream.casedist [of "i_rt n s1"],simp) -apply (rule stream.casedist [of "i_rt n s2"],auto) -by (drule stream_prefix1,auto) - -lemma i_th_stream_take_Suc [rule_format]: - "ALL s. i_th n (stream_take (Suc n)$s) = i_th n s" -apply (induct_tac n,auto) - apply (simp add: i_th_def) - apply (case_tac "s=UU",auto) - apply (drule stream_neq_UU,auto) -apply (case_tac "s=UU",simp add: i_th_def) -apply (drule stream_neq_UU,auto) -by (simp add: i_th_def i_rt_Suc_forw) - -lemma last_lemma10: "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==> - i_th n s1 << i_th n s2" -apply (rule i_th_stream_take_Suc [THEN subst]) -apply (rule i_th_stream_take_Suc [THEN subst]) back -apply (simp add: i_th_def) -apply (rule monofun_cfun_arg) -by (erule i_rt_mono) - -lemma i_th_last: "i_th n s && UU = i_rt n (stream_take (Suc n)$s)" -apply (insert surjectiv_scons [of "i_rt n (stream_take (Suc n)$s)"]) -apply (rule i_th_stream_take_Suc [THEN subst]) -apply (simp add: i_th_def i_rt_Suc_back [symmetric]) -by (simp add: i_rt_take_lemma1) - -lemma i_th_last_eq: -"i_th n s1 = i_th n s2 ==> i_rt n (stream_take (Suc n)$s1) = i_rt n (stream_take (Suc n)$s2)" -apply (insert i_th_last [of n s1]) -apply (insert i_th_last [of n s2]) -by auto - -lemma i_th_prefix_lemma: -"[| k <= n; stream_take (Suc n)$s1 << stream_take (Suc n)$s2 |] ==> - i_th k s1 << i_th k s2" -apply (subgoal_tac "stream_take (Suc k)$s1 << stream_take (Suc k)$s2") - apply (simp add: last_lemma10) -by (blast intro: stream_take_lemma10) - -lemma take_i_rt_prefix_lemma1: - "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==> - i_rt (Suc n) s1 << i_rt (Suc n) s2 ==> - i_rt n s1 << i_rt n s2 & stream_take n$s1 << stream_take n$s2" -apply auto - apply (insert i_th_prefix_lemma [of n n s1 s2]) - apply (rule i_th_i_rt_step,auto) -by (drule mono_stream_take_pred,simp) - -lemma take_i_rt_prefix_lemma: -"[| stream_take n$s1 << stream_take n$s2; i_rt n s1 << i_rt n s2 |] ==> s1 << s2" -apply (case_tac "n=0",simp) -apply (insert neq0_conv [of n]) -apply (insert not0_implies_Suc [of n],auto) -apply (subgoal_tac "stream_take 0$s1 << stream_take 0$s2 & - i_rt 0 s1 << i_rt 0 s2") - defer 1 - apply (rule zero_induct,blast) - apply (blast dest: take_i_rt_prefix_lemma1) -by simp - -lemma streams_prefix_lemma: "(s1 << s2) = - (stream_take n$s1 << stream_take n$s2 & i_rt n s1 << i_rt n s2)"; -apply auto - apply (simp add: monofun_cfun_arg) - apply (simp add: i_rt_mono) -by (erule take_i_rt_prefix_lemma,simp) - -lemma streams_prefix_lemma1: - "[| stream_take n$s1 = stream_take n$s2; i_rt n s1 = i_rt n s2 |] ==> s1 = s2" -apply (simp add: po_eq_conv,auto) - apply (insert streams_prefix_lemma) - by blast+ - - -(* ----------------------------------------------------------------------- *) - section "sconc" -(* ----------------------------------------------------------------------- *) - -lemma UU_sconc [simp]: " UU ooo s = s " -by (simp add: sconc_def inat_defs) - -lemma scons_neq_UU: "a~=UU ==> a && s ~=UU" -by auto - -lemma singleton_sconc [rule_format, simp]: "x~=UU --> (x && UU) ooo y = x && y" -apply (simp add: sconc_def inat_defs split:inat_splits,auto) -apply (rule someI2_ex,auto) - apply (rule_tac x="x && y" in exI,auto) -apply (simp add: i_rt_Suc_forw) -apply (case_tac "xa=UU",simp) -by (drule stream_neq_UU,auto) - -lemma ex_sconc [rule_format]: - "ALL k y. #x = Fin k --> (EX w. stream_take k$w = x & i_rt k w = y)" -apply (case_tac "#x") - apply (rule stream_finite_ind [of x],auto) - apply (simp add: stream.finite_def) - apply (drule slen_take_lemma1,blast) - apply (simp add: inat_defs split:inat_splits)+ -apply (erule_tac x="y" in allE,auto) -by (rule_tac x="a && w" in exI,auto) - -lemma rt_sconc1: "Fin n = #x ==> i_rt n (x ooo y) = y"; -apply (simp add: sconc_def inat_defs split:inat_splits , arith?,auto) -apply (rule someI2_ex,auto) -by (drule ex_sconc,simp) - -lemma sconc_inj2: "\Fin n = #x; x ooo y = x ooo z\ \ y = z" -apply (frule_tac y=y in rt_sconc1) -by (auto elim: rt_sconc1) - -lemma sconc_UU [simp]:"s ooo UU = s" -apply (case_tac "#s") - apply (simp add: sconc_def inat_defs) - apply (rule someI2_ex) - apply (rule_tac x="s" in exI) - apply auto - apply (drule slen_take_lemma1,auto) - apply (simp add: i_rt_lemma_slen) - apply (drule slen_take_lemma1,auto) - apply (simp add: i_rt_slen) -by (simp add: sconc_def inat_defs) - -lemma stream_take_sconc [simp]: "Fin n = #x ==> stream_take n$(x ooo y) = x" -apply (simp add: sconc_def) -apply (simp add: inat_defs split:inat_splits,auto) -apply (rule someI2_ex,auto) -by (drule ex_sconc,simp) - -lemma scons_sconc [rule_format,simp]: "a~=UU --> (a && x) ooo y = a && x ooo y" -apply (case_tac "#x",auto) - apply (simp add: sconc_def) - apply (rule someI2_ex) - apply (drule ex_sconc,simp) - apply (rule someI2_ex,auto) - apply (simp add: i_rt_Suc_forw) - apply (rule_tac x="a && x" in exI,auto) - apply (case_tac "xa=UU",auto) - apply (drule_tac s="stream_take nat$x" in scons_neq_UU) - apply (simp add: i_rt_Suc_forw) - apply (drule stream_neq_UU,clarsimp) - apply (drule streams_prefix_lemma1,simp+) -by (simp add: sconc_def) - -lemma ft_sconc: "x ~= UU ==> ft$(x ooo y) = ft$x" -by (rule stream.casedist [of x],auto) - -lemma sconc_assoc: "(x ooo y) ooo z = x ooo y ooo z" -apply (case_tac "#x") - apply (rule stream_finite_ind [of x],auto simp del: scons_sconc) - apply (simp add: stream.finite_def del: scons_sconc) - apply (drule slen_take_lemma1,auto simp del: scons_sconc) - apply (case_tac "a = UU", auto) -by (simp add: sconc_def) - - -(* ----------------------------------------------------------------------- *) - -lemma sconc_mono: "y << y' ==> x ooo y << x ooo y'" -apply (case_tac "#x") - apply (rule stream_finite_ind [of "x"]) - apply (auto simp add: stream.finite_def) - apply (drule slen_take_lemma1,blast) - by (simp add: stream_prefix',auto simp add: sconc_def) - -lemma sconc_mono1 [simp]: "x << x ooo y" -by (rule sconc_mono [of UU, simplified]) - -(* ----------------------------------------------------------------------- *) - -lemma empty_sconc [simp]: "(x ooo y = UU) = (x = UU & y = UU)" -apply (case_tac "#x",auto) - by (insert sconc_mono1 [of x y],auto); - -(* ----------------------------------------------------------------------- *) - -lemma rt_sconc [rule_format, simp]: "s~=UU --> rt$(s ooo x) = rt$s ooo x" -by (rule stream.casedist,auto) - -(* ----------------------------------------------------------------------- *) - -lemma sconc_lemma [rule_format, simp]: "ALL s. stream_take n$s ooo i_rt n s = s" -apply (induct_tac n,auto) -apply (case_tac "s=UU",auto) -by (drule stream_neq_UU,auto) - -(* ----------------------------------------------------------------------- *) - subsection "pointwise equality" -(* ----------------------------------------------------------------------- *) - -lemma ex_last_stream_take_scons: "stream_take (Suc n)$s = - stream_take n$s ooo i_rt n (stream_take (Suc n)$s)" -by (insert sconc_lemma [of n "stream_take (Suc n)$s"],simp) - -lemma i_th_stream_take_eq: -"!!n. ALL n. i_th n s1 = i_th n s2 ==> stream_take n$s1 = stream_take n$s2" -apply (induct_tac n,auto) -apply (subgoal_tac "stream_take (Suc na)$s1 = - stream_take na$s1 ooo i_rt na (stream_take (Suc na)$s1)") - apply (subgoal_tac "i_rt na (stream_take (Suc na)$s1) = - i_rt na (stream_take (Suc na)$s2)") - apply (subgoal_tac "stream_take (Suc na)$s2 = - stream_take na$s2 ooo i_rt na (stream_take (Suc na)$s2)") - apply (insert ex_last_stream_take_scons,simp) - apply blast - apply (erule_tac x="na" in allE) - apply (insert i_th_last_eq [of _ s1 s2]) -by blast+ - -lemma pointwise_eq_lemma[rule_format]: "ALL n. i_th n s1 = i_th n s2 ==> s1 = s2" -by (insert i_th_stream_take_eq [THEN stream.take_lemmas],blast) - -(* ----------------------------------------------------------------------- *) - subsection "finiteness" -(* ----------------------------------------------------------------------- *) - -lemma slen_sconc_finite1: - "[| #(x ooo y) = Infty; Fin n = #x |] ==> #y = Infty" -apply (case_tac "#y ~= Infty",auto) -apply (simp only: slen_infinite [symmetric]) -apply (drule_tac y=y in rt_sconc1) -apply (insert stream_finite_i_rt [of n "x ooo y"]) -by (simp add: slen_infinite) - -lemma slen_sconc_infinite1: "#x=Infty ==> #(x ooo y) = Infty" -by (simp add: sconc_def) - -lemma slen_sconc_infinite2: "#y=Infty ==> #(x ooo y) = Infty" -apply (case_tac "#x") - apply (simp add: sconc_def) - apply (rule someI2_ex) - apply (drule ex_sconc,auto) - apply (erule contrapos_pp) - apply (insert stream_finite_i_rt) - apply (simp add: slen_infinite ,auto) -by (simp add: sconc_def) - -lemma sconc_finite: "(#x~=Infty & #y~=Infty) = (#(x ooo y)~=Infty)" -apply auto - apply (case_tac "#x",auto) - apply (erule contrapos_pp,simp) - apply (erule slen_sconc_finite1,simp) - apply (drule slen_sconc_infinite1 [of _ y],simp) -by (drule slen_sconc_infinite2 [of _ x],simp) - -(* ----------------------------------------------------------------------- *) - -lemma slen_sconc_mono3: "[| Fin n = #x; Fin k = #(x ooo y) |] ==> n <= k" -apply (insert slen_mono [of "x" "x ooo y"]) -by (simp add: inat_defs split: inat_splits) - -(* ----------------------------------------------------------------------- *) - subsection "finite slen" -(* ----------------------------------------------------------------------- *) - -lemma slen_sconc: "[| Fin n = #x; Fin m = #y |] ==> #(x ooo y) = Fin (n + m)" -apply (case_tac "#(x ooo y)") - apply (frule_tac y=y in rt_sconc1) - apply (insert take_i_rt_len [of "THE j. Fin j = #(x ooo y)" "x ooo y" n n m],simp) - apply (insert slen_sconc_mono3 [of n x _ y],simp) -by (insert sconc_finite [of x y],auto) - -(* ----------------------------------------------------------------------- *) - subsection "flat prefix" -(* ----------------------------------------------------------------------- *) - -lemma sconc_prefix: "(s1::'a::flat stream) << s2 ==> EX t. s1 ooo t = s2" -apply (case_tac "#s1") - apply (subgoal_tac "stream_take nat$s1 = stream_take nat$s2"); - apply (rule_tac x="i_rt nat s2" in exI) - apply (simp add: sconc_def) - apply (rule someI2_ex) - apply (drule ex_sconc) - apply (simp,clarsimp,drule streams_prefix_lemma1) - apply (simp+,rule slen_take_lemma3 [rule_format, of _ s1 s2]); - apply (simp+,rule_tac x="UU" in exI) -apply (insert slen_take_lemma3 [rule_format, of _ s1 s2]); -by (rule stream.take_lemmas,simp) - -(* ----------------------------------------------------------------------- *) - subsection "continuity" -(* ----------------------------------------------------------------------- *) - -lemma chain_sconc: "chain S ==> chain (%i. (x ooo S i))" -by (simp add: chain_def,auto simp add: sconc_mono) - -lemma chain_scons: "chain S ==> chain (%i. a && S i)" -apply (simp add: chain_def,auto) -by (rule monofun_cfun_arg,simp) - -lemma contlub_scons: "contlub (%x. a && x)" -by (simp add: contlub_Rep_CFun2) - -lemma contlub_scons_lemma: "chain S ==> (LUB i. a && S i) = a && (LUB i. S i)" -apply (insert contlub_scons [of a]) -by (simp only: contlub) - -lemma finite_lub_sconc: "chain Y ==> (stream_finite x) ==> - (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))" -apply (rule stream_finite_ind [of x]) - apply (auto) -apply (subgoal_tac "(LUB i. a && (s ooo Y i)) = a && (LUB i. s ooo Y i)") - by (force,blast dest: contlub_scons_lemma chain_sconc) - -lemma contlub_sconc_lemma: - "chain Y ==> (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))" -apply (case_tac "#x=Infty") - apply (simp add: sconc_def) - prefer 2 - apply (drule finite_lub_sconc,auto simp add: slen_infinite) -apply (simp add: slen_def) -apply (insert lub_const [of x] unique_lub [of _ x _]) -by (auto simp add: lub) - -lemma contlub_sconc: "contlub (%y. x ooo y)"; -by (rule contlubI, insert contlub_sconc_lemma [of _ x], simp); - -lemma monofun_sconc: "monofun (%y. x ooo y)" -by (simp add: monofun sconc_mono) - -lemma cont_sconc: "cont (%y. x ooo y)" -apply (rule monocontlub2cont) - apply (rule monofunI, simp add: sconc_mono) -by (rule contlub_sconc); - - -(* ----------------------------------------------------------------------- *) - section "constr_sconc" -(* ----------------------------------------------------------------------- *) - -lemma constr_sconc_UUs [simp]: "constr_sconc UU s = s" -by (simp add: constr_sconc_def inat_defs) - -lemma "x ooo y = constr_sconc x y" -apply (case_tac "#x") - apply (rule stream_finite_ind [of x],auto simp del: scons_sconc) - defer 1 - apply (simp add: constr_sconc_def del: scons_sconc) - apply (case_tac "#s") - apply (simp add: inat_defs) - apply (case_tac "a=UU",auto simp del: scons_sconc) - apply (simp) - apply (simp add: sconc_def) - apply (simp add: constr_sconc_def) -apply (simp add: stream.finite_def) -by (drule slen_take_lemma1,auto) - -end diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/ex/Dagstuhl.ML --- a/src/HOLCF/ex/Dagstuhl.ML Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/ex/Dagstuhl.ML Tue Sep 07 16:02:42 2004 +0200 @@ -30,12 +30,12 @@ val lemma5=result(); val prems = goal Dagstuhl.thy "YS = YYS"; -by (resolve_tac stream.take_lemmas 1); +by (resolve_tac (thms "stream.take_lemmas") 1); by (induct_tac "n" 1); -by (simp_tac (simpset() addsimps stream.rews) 1); +by (simp_tac (simpset() addsimps (thms "stream.rews")) 1); by (stac YS_def2 1); by (stac YYS_def2 1); -by (asm_simp_tac (simpset() addsimps stream.rews) 1); +by (asm_simp_tac (simpset() addsimps (thms "stream.rews")) 1); by (rtac (lemma5 RS sym RS subst) 1); by (rtac refl 1); val wir_moel=result(); diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/ex/Stream.ML --- a/src/HOLCF/ex/Stream.ML Tue Sep 07 15:59:16 2004 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,594 +0,0 @@ -(* Title: HOLCF/ex//Stream.ML - ID: $Id$ - Author: Franz Regensburger and David von Oheimb, TU Muenchen - -General Stream domain. -*) - -fun stream_case_tac s i = res_inst_tac [("x",s)] stream.casedist i - THEN hyp_subst_tac i THEN hyp_subst_tac (i+1); - - -val [stream_con_rew1,stream_con_rew2] = stream.con_rews; - -Addsimps stream.rews; -Addsimps[eq_UU_iff RS sym]; - -(* ----------------------------------------------------------------------- *) -(* theorems about scons *) -(* ----------------------------------------------------------------------- *) - -section "scons"; - -Goal "(a && s = UU) = (a = UU)"; -by Safe_tac; -by (etac contrapos_pp 1); -by (safe_tac (claset() addSIs stream.con_rews)); -qed "scons_eq_UU"; - -Goal "[| a && x = UU; a ~= UU |] ==> R"; -by (dtac stream_con_rew2 1); -by (contr_tac 1); -qed "scons_not_empty"; - -Goal "(x ~= UU) = (EX a y. a ~= UU & x = a && y)"; -by (cut_facts_tac [stream.exhaust] 1); -by (best_tac (claset() addDs [stream_con_rew2]) 1); -qed "stream_exhaust_eq"; - -Goal -"[| a && s << t; a ~= UU |] ==> EX b tt. t = b && tt & b ~= UU & s << tt"; -by (stream_case_tac "t" 1); -by (fast_tac (claset() addDs [eq_UU_iff RS iffD2, stream_con_rew2]) 1); -by (fast_tac (claset() addDs stream.inverts) 1); -qed "stream_prefix"; - -Goal "[| x && xs << y && ys; (x::'a::flat) ~= UU|] ==> x = y & xs << ys"; -by (cut_inst_tac [("x1","x::'a::flat"), ("x","y::'a::flat")] - (ax_flat RS spec RS spec) 1); -by (forward_tac stream.inverts 1); -by Safe_tac; -by (dtac (hd stream.con_rews RS subst) 1); -by (fast_tac (claset() addDs [eq_UU_iff RS iffD2, stream_con_rew2]) 1); -qed "stream_flat_prefix"; - -Goal "b ~= UU ==> x << b && z = \ -\ (x = UU | (EX a y. x = a && y & a ~= UU & a << b & y << z))"; -by Safe_tac; -by (stream_case_tac "x" 1); -by (safe_tac (claset() addSDs stream.inverts - addSIs [monofun_cfun, monofun_cfun_arg])); -by (Fast_tac 1); -qed "stream_prefix'"; - -Goal "[| a ~= UU; b ~= UU |] ==> (a && s = b && t) = (a = b & s = t)"; -by (best_tac (claset() addSEs stream.injects) 1); -qed "stream_inject_eq"; -Addsimps [stream_inject_eq]; - - -(* ----------------------------------------------------------------------- *) -(* theorems about stream_when *) -(* ----------------------------------------------------------------------- *) - -section "stream_when"; - -Goal "stream_when$UU$s=UU"; -by (stream_case_tac "s" 1); -by (ALLGOALS Asm_simp_tac); -qed "stream_when_strictf"; - - -(* ----------------------------------------------------------------------- *) -(* theorems about ft and rt *) -(* ----------------------------------------------------------------------- *) - -section "ft & rt"; - -(* -Goal "ft$s=UU --> s=UU"; -by (stream_case_tac "s" 1); -by (REPEAT (asm_simp_tac (simpset() addsimps stream.rews) 1)); -qed "stream_ft_lemma1"; -*) - -Goal "s~=UU ==> ft$s~=UU"; -by (stream_case_tac "s" 1); -by (Blast_tac 1); -by (Asm_simp_tac 1); -qed "ft_defin"; - -Goal "rt$s~=UU ==> s~=UU"; -by Auto_tac; -qed "rt_strict_rev"; - -Goal "(ft$s)&&(rt$s)=s"; -by (stream_case_tac "s" 1); -by (ALLGOALS Asm_simp_tac); -qed "surjectiv_scons"; - -Goal "ALL x s. x << s --> iterate i rt x << iterate i rt s"; -by (induct_tac "i" 1); -by (Simp_tac 1); -by (strip_tac 1); -by (stream_case_tac "x" 1); -by (rtac (minimal RS (monofun_iterate2 RS monofunE RS spec RS spec RS mp))1); -by (dtac stream_prefix 1); -by Safe_tac; -by (asm_simp_tac (HOL_ss addsimps iterate_Suc2::stream.rews) 1); -qed_spec_mp "monofun_rt_mult"; - - - -(* ----------------------------------------------------------------------- *) -(* theorems about stream_take *) -(* ----------------------------------------------------------------------- *) - -section "stream_take"; - -Goal "(LUB i. stream_take i$s) = s"; -by (subgoal_tac "(LUB i. stream_take i$s) = fix$stream_copy$s" 1); -by (simp_tac (simpset() addsimps [fix_def2, stream.take_def, - contlub_cfun_fun, chain_iterate]) 2); -by (asm_full_simp_tac (simpset() addsimps [stream.reach]) 1); -qed "stream_reach2"; - -Goal "chain (%i. stream_take i$s)"; -by (rtac chainI 1); -by (subgoal_tac "ALL i s::'a stream. stream_take i$s << stream_take (Suc i)$s" 1); -by (Fast_tac 1); -by (rtac allI 1); -by (induct_tac "ia" 1); -by (Simp_tac 1); -by (rtac allI 1); -by (stream_case_tac "s" 1); -by (Simp_tac 1); -by (asm_simp_tac (simpset() addsimps [monofun_cfun_arg]) 1); -qed "chain_stream_take"; - - -Goalw [stream.take_def] - "stream_take n$x = x ==> stream_take (Suc n)$x = x"; -by (rtac antisym_less 1); -by (subgoal_tac "iterate (Suc n) stream_copy UU$x << fix$stream_copy$x" 1); -by (asm_full_simp_tac (simpset() addsimps [stream.reach]) 1); -by (rtac monofun_cfun_fun 1); -by (stac fix_def2 1); -by (rtac is_ub_thelub 1); -by (rtac chain_iterate 1); -by (etac subst 1 THEN rtac monofun_cfun_fun 1); -by (rtac (rewrite_rule [chain_def] chain_iterate RS spec) 1); -qed "stream_take_more"; - -(* -Goal - "ALL s2. stream_take n$s2 = s2 --> stream_take (Suc n)$s2=s2"; -by (induct_tac "n" 1); -by (simp_tac (simpset() addsimps stream_rews) 1); -by (strip_tac 1); -by (hyp_subst_tac 1); -by (simp_tac (simpset() addsimps stream_rews) 1); -by (rtac allI 1); -by (res_inst_tac [("s","s2")] streamE 1); -by (asm_simp_tac (simpset() addsimps stream_rews) 1); -by (asm_simp_tac (simpset() addsimps stream_rews) 1); -by (strip_tac 1 ); -by (subgoal_tac "stream_take n1$xs = xs" 1); -by (rtac ((hd stream_inject) RS conjunct2) 2); -by (atac 4); -by (atac 2); -by (atac 2); -by (rtac cfun_arg_cong 1); -by (Blast_tac 1); -qed "stream_take_lemma2"; -*) - -Goal -"ALL x xs. x~=UU --> stream_take n$(x && xs) = x && xs --> stream_take n$xs=xs"; -by (induct_tac "n" 1); -by (Asm_simp_tac 1); -by (strip_tac 1); -by (res_inst_tac [("P","x && xs=UU")] notE 1); -by (eresolve_tac stream.con_rews 1); -by (etac sym 1); -by (strip_tac 1); -by (rtac stream_take_more 1); -by (res_inst_tac [("a1","x")] ((hd stream.injects) RS conjunct2) 1); -by (assume_tac 3); -by (etac (hd(tl(tl(stream.take_rews))) RS subst) 1); -by (atac 1); -qed_spec_mp "stream_take_lemma3"; - -Goal - "ALL x xs. stream_take n$xs=xs --> stream_take (Suc n)$(x && xs) = x && xs"; -by (induct_tac "n" 1); -by Auto_tac; -qed_spec_mp "stream_take_lemma4"; - - -(* ------------------------------------------------------------------------- *) -(* special induction rules *) -(* ------------------------------------------------------------------------- *) - -section "induction"; - - -val prems = Goalw [stream.finite_def] - "[| stream_finite x; \ -\ P UU; \ -\ !!a s. [| a ~= UU; P s |] \ -\ ==> P (a && s) |] ==> P x"; -by (cut_facts_tac prems 1); -by (etac exE 1); -by (etac subst 1); -by (rtac stream.finite_ind 1); -by (atac 1); -by (eresolve_tac prems 1); -by (atac 1); -qed "stream_finite_ind"; - -val major::prems = Goal -"[| P UU;\ -\ !! x . x ~= UU ==> P (x && UU); \ -\ !! y z s. [| y ~= UU; z ~= UU; P s |] ==> P (y && z && s ) \ -\ |] ==> !s. P (stream_take n$s)"; -by (res_inst_tac [("n","n")] nat_induct2 1); -by (asm_simp_tac (simpset() addsimps [major]) 1); -by (rtac allI 1); -by (stream_case_tac "s" 1); -by (asm_simp_tac (simpset() addsimps [major]) 1); -by (asm_simp_tac (simpset() addsimps prems) 1); -by (rtac allI 1); -by (stream_case_tac "s" 1); -by (asm_simp_tac (simpset() addsimps [major]) 1); -by (res_inst_tac [("x","sa")] stream.casedist 1); -by (asm_simp_tac (simpset() addsimps prems) 1); -by (asm_simp_tac (simpset() addsimps prems) 1); -qed "stream_finite_ind2"; - - -val prems = Goal -"[| adm P; P UU; !!a. a ~= UU ==> P (a && UU); \ -\ !!a b s. [| a ~= UU; b ~= UU; P s |] ==> P (a && b && s)\ -\ |] ==> P x"; -by (rtac (stream.reach RS subst) 1); -by (rtac (adm_impl_admw RS wfix_ind) 1); -by (rtac adm_subst 1); -by (cont_tacR 1); -by (resolve_tac prems 1); -by (rtac allI 1); -by (rtac (rewrite_rule [stream.take_def] stream_finite_ind2 RS spec) 1); -by (resolve_tac prems 1); -by (eresolve_tac prems 1); -by (eresolve_tac prems 1); -by (atac 1); -by (atac 1); -qed "stream_ind2"; - - -(* ----------------------------------------------------------------------- *) -(* simplify use of coinduction *) -(* ----------------------------------------------------------------------- *) - -section "coinduction"; - - -Goalw [stream.bisim_def] -"!s1 s2. R s1 s2 --> ft$s1 = ft$s2 & R (rt$s1) (rt$s2) ==> stream_bisim R"; -by (strip_tac 1); -by (etac allE 1); -by (etac allE 1); -by (dtac mp 1); -by (atac 1); -by (etac conjE 1); -by (case_tac "x = UU & x' = UU" 1); -by (rtac disjI1 1); -by (Blast_tac 1); -by (rtac disjI2 1); -by (rtac disjE 1); -by (etac (de_Morgan_conj RS subst) 1); -by (res_inst_tac [("x","ft$x")] exI 1); -by (res_inst_tac [("x","rt$x")] exI 1); -by (res_inst_tac [("x","rt$x'")] exI 1); -by (rtac conjI 1); -by (etac ft_defin 1); -by (asm_simp_tac (simpset() addsimps [surjectiv_scons]) 1); -by (eres_inst_tac [("s","ft$x"),("t","ft$x'")] subst 1); -by (simp_tac (simpset() addsimps [surjectiv_scons]) 1); -by (res_inst_tac [("x","ft$x'")] exI 1); -by (res_inst_tac [("x","rt$x")] exI 1); -by (res_inst_tac [("x","rt$x'")] exI 1); -by (rtac conjI 1); -by (etac ft_defin 1); -by (asm_simp_tac (simpset() addsimps [surjectiv_scons]) 1); -by (res_inst_tac [("s","ft$x"),("t","ft$x'")] ssubst 1); -by (etac sym 1); -by (simp_tac (simpset() addsimps [surjectiv_scons]) 1); -qed "stream_coind_lemma2"; - -(* ----------------------------------------------------------------------- *) -(* theorems about stream_finite *) -(* ----------------------------------------------------------------------- *) - -section "stream_finite"; - - -Goalw [stream.finite_def] "stream_finite UU"; -by (rtac exI 1); -by (Simp_tac 1); -qed "stream_finite_UU"; - -Goal "~ stream_finite s ==> s ~= UU"; -by (blast_tac (claset() addIs [stream_finite_UU]) 1); -qed "stream_finite_UU_rev"; - -Goalw [stream.finite_def] "stream_finite xs ==> stream_finite (x && xs)"; -by (blast_tac (claset() addIs [stream_take_lemma4]) 1); -qed "stream_finite_lemma1"; - -Goalw [stream.finite_def] - "[| x ~= UU; stream_finite (x && xs) |] ==> stream_finite xs"; -by (blast_tac (claset() addIs [stream_take_lemma3]) 1); -qed "stream_finite_lemma2"; - -Goal "stream_finite (rt$s) = stream_finite s"; -by (stream_case_tac "s" 1); -by (simp_tac (simpset() addsimps [stream_finite_UU]) 1); -by (Asm_simp_tac 1); -by (fast_tac (claset() addIs [stream_finite_lemma1] - addDs [stream_finite_lemma2]) 1); -qed "stream_finite_rt_eq"; - -Goal "stream_finite s ==> !t. t< stream_finite t"; -by (eres_inst_tac [("x","s")] stream_finite_ind 1); -by (strip_tac 1); -by (dtac UU_I 1); -by (asm_simp_tac (simpset() addsimps [stream_finite_UU]) 1); -by (strip_tac 1); -by (asm_full_simp_tac (simpset() addsimps [stream_prefix']) 1); -by (fast_tac (claset() addSIs [stream_finite_UU, stream_finite_lemma1]) 1); -qed_spec_mp "stream_finite_less"; - -Goal "adm (%x. ~ stream_finite x)"; -by (rtac admI2 1); -by (best_tac (claset() addIs [stream_finite_less, is_ub_thelub]) 1); -qed "adm_not_stream_finite"; - -section "smap"; - -bind_thm ("smap_unfold", fix_prover2 thy smap_def - "smap = (\\ f s. case s of x && l => f\\x && smap\\f\\l)"); - -Goal "smap\\f\\\\ = \\"; -by (stac smap_unfold 1); -by (Simp_tac 1); -qed "smap_empty"; - -Goal "x~=\\ ==> smap\\f\\(x&&xs) = (f\\x)&&(smap\\f\\xs)"; -by (rtac trans 1); -by (stac smap_unfold 1); -by (Asm_simp_tac 1); -by (rtac refl 1); -qed "smap_scons"; - -Addsimps [smap_empty, smap_scons]; - -(* ------------------------------------------------------------------------- *) - -section "slen"; - -Goalw [slen_def, stream.finite_def] "#\\ = 0"; -by (Simp_tac 1); -by (stac Least_equality 1); -by Auto_tac; -by (simp_tac(simpset() addsimps [thm "Fin_0"]) 1); -qed "slen_empty"; - -Goalw [slen_def] "x ~= \\ ==> #(x&&xs) = iSuc (#xs)"; -by (res_inst_tac [("Q1","stream_finite (x && xs)")] (split_if RS iffD2) 1); -by Safe_tac; -by (rtac (split_if RS iffD2) 2); -by Safe_tac; -by (fast_tac (claset() addIs [stream_finite_lemma1]) 2); -by (rtac (thm "iSuc_Infty" RS sym) 2); -by (rtac (split_if RS iffD2) 1); -by (Simp_tac 1); -by Safe_tac; -by (eatac stream_finite_lemma2 1 2); -by (rewtac stream.finite_def); -by (Clarify_tac 1); -by (eatac Least_Suc2 1 1); -by (rtac not_sym 1); -by Auto_tac; -qed "slen_scons"; - -Addsimps [slen_empty, slen_scons]; - -Goal "(#x < Fin (Suc 0)) = (x = UU)"; -by (stream_case_tac "x" 1); -by (auto_tac (claset(), simpset() delsimps [thm "iSuc_Fin"] addsimps - [thm "Fin_0", thm "iSuc_Fin" RS sym, thm "i0_iless_iSuc", thm "iSuc_mono"])); -qed "slen_less_1_eq"; - -Goal "(#x = 0) = (x = \\)"; -by Auto_tac; -by (stream_case_tac "x" 1); -by (rotate_tac ~1 2); -by Auto_tac; -qed "slen_empty_eq"; - -Goal "(Fin (Suc n) < #x) = (? a y. x = a && y & a ~= \\ & Fin n < #y)"; -by (stream_case_tac "x" 1); -by (auto_tac (claset(), simpset() delsimps [thm "iSuc_Fin"] addsimps - [thm "iSuc_Fin" RS sym, thm "iSuc_mono"])); -by (dtac sym 1); -by Auto_tac; -qed "slen_scons_eq"; - - -Goal "#x = iSuc n --> (? a y. x = a&&y & a ~= \\ & #y = n)"; -by (stream_case_tac "x" 1); -by Auto_tac; -qed_spec_mp "slen_iSuc"; - - -Goal -"(#x < Fin (Suc (Suc n))) = (!a y. x ~= a && y | a = \\ | #y < Fin (Suc n))"; -by (stac (thm "iSuc_Fin" RS sym) 1); -by (stac (thm "iSuc_Fin" RS sym) 1); -by (safe_tac HOL_cs); -by (rotate_tac ~1 1); -by (asm_full_simp_tac (HOL_ss addsimps [slen_scons, thm "iSuc_mono"]) 1); -by (Asm_full_simp_tac 1); -by (stream_case_tac "x" 1); -by ( simp_tac (HOL_ss addsimps [slen_empty, thm "i0_iless_iSuc"]) 1); -by (asm_simp_tac (HOL_ss addsimps [slen_scons, thm "iSuc_mono"]) 1); -by (etac allE 1); -by (etac allE 1); -by (safe_tac HOL_cs); -by ( contr_tac 2); -by ( fast_tac HOL_cs 1); -by (Asm_full_simp_tac 1); -qed "slen_scons_eq_rev"; - -Goal "!x. (Fin n < #x) = (stream_take n\\x ~= x)"; -by (induct_tac "n" 1); -by (simp_tac(simpset() addsimps [slen_empty_eq, thm "Fin_0"]) 1); -by (Fast_tac 1); -by (rtac allI 1); -by (asm_simp_tac (simpset() addsimps [slen_scons_eq]) 1); -by (etac thin_rl 1); -by (safe_tac HOL_cs); -by (Asm_full_simp_tac 1); -by (stream_case_tac "x" 1); -by (rotate_tac ~1 2); -by Auto_tac; -qed_spec_mp "slen_take_eq"; - -Goalw [thm "ile_def"] "(#x <= Fin n) = (stream_take n\\x = x)"; -by (simp_tac (simpset() addsimps [slen_take_eq]) 1); -qed "slen_take_eq_rev"; - -Goal "#x = Fin n ==> stream_take n\\x = x"; -by (asm_simp_tac (simpset() addsimps [slen_take_eq_rev RS sym]) 1); -qed "slen_take_lemma1"; - -Goal "!x. ~stream_finite x --> #(stream_take i\\x) = Fin i"; -by (induct_tac "i" 1); -by (simp_tac(simpset() addsimps [thm "Fin_0"]) 1); -by (rtac allI 1); -by (stream_case_tac "x" 1); -by (auto_tac (claset()addSIs [stream_finite_lemma1], simpset() delsimps [thm "iSuc_Fin"] addsimps [thm "iSuc_Fin" RS sym])); -by (force_tac (claset(), simpset() addsimps [stream.finite_def]) 1); -qed_spec_mp "slen_take_lemma2"; - -Goal "stream_finite x = (#x ~= Infty)"; -by (rtac iffI 1); -by (etac stream_finite_ind 1); -by (Simp_tac 1); -by (etac (slen_scons RS ssubst) 1); -by (stac (thm "iSuc_Infty" RS sym) 1); -by (etac contrapos_nn 1); -by (etac (thm "iSuc_inject" RS iffD1) 1); -by (case_tac "#x" 1); -by (auto_tac (claset()addSDs [slen_take_lemma1], - simpset() addsimps [stream.finite_def])); -qed "slen_infinite"; - -Goal "s << t ==> #s <= #t"; -by (case_tac "stream_finite t" 1); -by (EVERY'[dtac (slen_infinite RS subst), dtac notnotD] 2); -by (Asm_simp_tac 2); -by (etac rev_mp 1); -by (res_inst_tac [("x","s")] allE 1); -by (atac 2); -by (etac (stream_finite_ind) 1); -by (Simp_tac 1); -by (rtac allI 1); -by (stream_case_tac "x" 1); -by (asm_simp_tac (HOL_ss addsimps [slen_empty, thm "i0_lb"]) 1); -by (asm_simp_tac (HOL_ss addsimps [slen_scons]) 1); -by (fast_tac(claset() addSIs [thm "iSuc_ile_mono" RS iffD2] addSDs stream.inverts) 1); -qed "slen_mono"; - -Goal "!(x::'a::flat stream) y. Fin n <= #x --> x << y --> \ -\ stream_take n\\x = stream_take n\\y"; -by (induct_tac "n" 1); -by ( simp_tac(simpset() addsimps [slen_empty_eq]) 1); -by (strip_tac 1); -by (stream_case_tac "x" 1); -by (asm_full_simp_tac (HOL_ss addsimps [slen_empty, - thm "iSuc_Fin" RS sym, thm "not_iSuc_ilei0"]) 1); -by (fatac stream_prefix 1 1); -by (safe_tac (claset() addSDs [stream_flat_prefix])); -by (Simp_tac 1); -by (rtac cfun_arg_cong 1); -by (rotate_tac 3 1); -by (asm_full_simp_tac (simpset() delsimps [thm "iSuc_Fin"] addsimps - [thm "iSuc_Fin" RS sym, thm "iSuc_ile_mono"]) 1); -qed_spec_mp "slen_take_lemma3"; - -Goal "stream_finite t ==> \ -\!s. #(s::'a::flat stream) = #t & s << t --> s = t"; -by (etac stream_finite_ind 1); -by (fast_tac (claset() addDs [eq_UU_iff RS iffD2]) 1); -by Safe_tac; -by (stream_case_tac "sa" 1); -by (dtac sym 1); -by (Asm_full_simp_tac 1); -by (safe_tac (claset() addSDs [stream_flat_prefix])); -by (rtac cfun_arg_cong 1); -by (etac allE 1); -by (etac mp 1); -by (Asm_full_simp_tac 1); -qed "slen_strict_mono_lemma"; - -Goal "[|stream_finite t; s ~= t; s << (t::'a::flat stream) |] ==> #s < #t"; -by (rtac (thm "ilessI1") 1); -by (etac slen_mono 1); -by (dtac slen_strict_mono_lemma 1); -by (Fast_tac 1); -qed "slen_strict_mono"; - -Goal "!x. Fin (i + j) <= #x --> Fin j <= #(iterate i rt x)"; -by (induct_tac "i" 1); -by (Simp_tac 1); -by (strip_tac 1); -by (stream_case_tac "x" 1); -by (asm_full_simp_tac (simpset() delsimps [thm "iSuc_Fin"] - addsimps [thm "iSuc_Fin" RS sym]) 1); -by (stac iterate_Suc2 1); -by (rotate_tac 2 1); -by (asm_full_simp_tac (simpset() delsimps [thm "iSuc_Fin"] - addsimps [thm "iSuc_Fin" RS sym, thm "iSuc_ile_mono"]) 1); -qed_spec_mp "slen_rt_mult"; - - -(* ------------------------------------------------------------------------- *) - -section "sfilter"; - -bind_thm ("sfilter_unfold", fix_prover2 thy sfilter_def - "sfilter = (\\ p s. case s of x && xs => \ -\ If p\\x then x && sfilter\\p\\xs else sfilter\\p\\xs fi)"); - -Goal "sfilter\\\\ = \\"; -by (rtac ext_cfun 1); -by (stac sfilter_unfold 1); -by (stream_case_tac "x" 1); -by Auto_tac; -qed "strict_sfilter"; - -Goal "sfilter\\f\\\\ = \\"; -by (stac sfilter_unfold 1); -by (Simp_tac 1); -qed "sfilter_empty"; - -Goal "x ~= \\ ==> sfilter\\f\\(x && xs) = \ -\ If f\\x then x && sfilter\\f\\xs else sfilter\\f\\xs fi"; -by (rtac trans 1); -by (stac sfilter_unfold 1); -by (Asm_simp_tac 1); -by (rtac refl 1); -qed "sfilter_scons"; - diff -r 8b74a39dba89 -r 9d57263faf9e src/HOLCF/ex/Stream.thy --- a/src/HOLCF/ex/Stream.thy Tue Sep 07 15:59:16 2004 +0200 +++ b/src/HOLCF/ex/Stream.thy Tue Sep 07 16:02:42 2004 +0200 @@ -1,29 +1,1012 @@ (* Title: HOLCF/ex/Stream.thy ID: $Id$ - Author: Franz Regensburger, David von Oheimb + Author: Franz Regensburger, David von Oheimb, Borislav Gajanovic General Stream domain. -TODO: should be integrated with HOLCF/Streams *) -Stream = HOLCF + Nat_Infinity + +theory Stream = HOLCF + Nat_Infinity: -domain 'a stream = "&&" (ft::'a) (lazy rt::'a stream) (infixr 65) +domain 'a stream = "&&" (ft::'a) (lazy rt::"'a stream") (infixr 65) consts - smap :: "('a -> 'b) -> 'a stream -> 'b stream" - sfilter :: "('a -> tr) -> 'a stream -> 'a stream" - slen :: "'a stream => inat" ("#_" [1000] 1000) + smap :: "('a \ 'b) \ 'a stream \ 'b stream" + sfilter :: "('a \ tr) \ 'a stream \ 'a stream" + slen :: "'a stream \ inat" ("#_" [1000] 1000) defs - smap_def "smap == fix\\(\\ h f s. case s of x && xs => f\\x && h\\f\\xs)" - sfilter_def "sfilter == fix\\(\\ h p s. case s of x && xs => - If p\\x then x && h\\p\\xs else h\\p\\xs fi)" - slen_def "#s == if stream_finite s - then Fin (LEAST n. stream_take n\\s = s) else \\" + smap_def: "smap \ fix\(\ h f s. case s of x && xs \ f\x && h\f\xs)" + sfilter_def: "sfilter \ fix\(\ h p s. case s of x && xs \ + If p\x then x && h\p\xs else h\p\xs fi)" + slen_def: "#s \ if stream_finite s + then Fin (LEAST n. stream_take n\s = s) else \" + +(* concatenation *) + +consts + + i_rt :: "nat => 'a stream => 'a stream" (* chops the first i elements *) + i_th :: "nat => 'a stream => 'a" (* the i-th element ä*) + + sconc :: "'a stream => 'a stream => 'a stream" (infixr "ooo" 65) + constr_sconc :: "'a stream => 'a stream => 'a stream" (* constructive *) + constr_sconc' :: "nat => 'a stream => 'a stream => 'a stream" + +defs + i_rt_def: "i_rt == %i s. iterate i rt s" + i_th_def: "i_th == %i s. ft$(i_rt i s)" + + sconc_def: "s1 ooo s2 == case #s1 of + Fin n \ (SOME s. (stream_take n$s=s1) & (i_rt n s = s2)) + | \ \ s1" + + constr_sconc_def: "constr_sconc s1 s2 == case #s1 of + Fin n \ constr_sconc' n s1 s2 + | \ \ s1" +primrec + constr_sconc'_0: "constr_sconc' 0 s1 s2 = s2" + constr_sconc'_Suc: "constr_sconc' (Suc n) s1 s2 = ft$s1 && + constr_sconc' n (rt$s1) s2" + + +declare stream.rews [simp add] + +(* ----------------------------------------------------------------------- *) +(* theorems about scons *) +(* ----------------------------------------------------------------------- *) + + +section "scons" + +lemma scons_eq_UU: "(a && s = UU) = (a = UU)" +by (auto, erule contrapos_pp, simp) + +lemma scons_not_empty: "[| a && x = UU; a ~= UU |] ==> R" +by auto + +lemma stream_exhaust_eq: "(x ~= UU) = (EX a y. a ~= UU & x = a && y)" +by (auto,insert stream.exhaust [of x],auto) + +lemma stream_neq_UU: "x~=UU ==> EX a as. x=a&&as & a~=UU" +by (simp add: stream_exhaust_eq,auto) + +lemma stream_inject_eq [simp]: + "[| a ~= UU; b ~= UU |] ==> (a && s = b && t) = (a = b & s = t)" +by (insert stream.injects [of a s b t], auto) + +lemma stream_prefix: + "[| a && s << t; a ~= UU |] ==> EX b tt. t = b && tt & b ~= UU & s << tt" +apply (insert stream.exhaust [of t], auto) +apply (drule eq_UU_iff [THEN iffD2], simp) +by (drule stream.inverts, auto) + +lemma stream_prefix': + "b ~= UU ==> x << b && z = + (x = UU | (EX a y. x = a && y & a ~= UU & a << b & y << z))" +apply (case_tac "x=UU",auto) +apply (drule stream_exhaust_eq [THEN iffD1],auto) +apply (drule stream.inverts,auto) +by (intro monofun_cfun,auto) + +(* +lemma stream_prefix1: "[| x< x&&xs << y&&ys" +by (insert stream_prefix' [of y "x&&xs" ys],force) +*) + +lemma stream_flat_prefix: + "[| x && xs << y && ys; (x::'a::flat) ~= UU|] ==> x = y & xs << ys" +apply (case_tac "y=UU",auto) +apply (drule eq_UU_iff [THEN iffD2],auto) +apply (drule stream.inverts,auto) +apply (drule ax_flat [rule_format],simp) +by (drule stream.inverts,auto) + + + +(* ----------------------------------------------------------------------- *) +(* theorems about stream_when *) +(* ----------------------------------------------------------------------- *) + +section "stream_when" + + +lemma stream_when_strictf: "stream_when$UU$s=UU" +by (rule stream.casedist [of s], auto) + + + +(* ----------------------------------------------------------------------- *) +(* theorems about ft and rt *) +(* ----------------------------------------------------------------------- *) + + +section "ft & rt" + + +lemma ft_defin: "s~=UU ==> ft$s~=UU" +by (drule stream_exhaust_eq [THEN iffD1],auto) + +lemma rt_strict_rev: "rt$s~=UU ==> s~=UU" +by auto + +lemma surjectiv_scons: "(ft$s)&&(rt$s)=s" +by (rule stream.casedist [of s], auto) + +lemma monofun_rt_mult: "x << s ==> iterate i rt x << iterate i rt s" +by (insert monofun_iterate2 [of i "rt"], simp add: monofun, auto) + + + +(* ----------------------------------------------------------------------- *) +(* theorems about stream_take *) +(* ----------------------------------------------------------------------- *) + + +section "stream_take"; + + +lemma stream_reach2: "(LUB i. stream_take i$s) = s" +apply (insert stream.reach [of s], erule subst) back +apply (simp add: fix_def2 stream.take_def) +apply (insert contlub_cfun_fun [of "%i. iterate i stream_copy UU" s,THEN sym]) +by (simp add: chain_iterate) + +lemma chain_stream_take: "chain (%i. stream_take i$s)" +apply (rule chainI) +apply (rule monofun_cfun_fun) +apply (simp add: stream.take_def del: iterate_Suc) +by (rule chainE, simp add: chain_iterate) + +lemma stream_take_prefix [simp]: "stream_take n$s << s" +apply (insert stream_reach2 [of s]) +apply (erule subst) back +apply (rule is_ub_thelub) +by (simp only: chain_stream_take) + +lemma stream_take_more [rule_format]: + "ALL x. stream_take n$x = x --> stream_take (Suc n)$x = x" +apply (induct_tac n,auto) +apply (case_tac "x=UU",auto) +by (drule stream_exhaust_eq [THEN iffD1],auto) + +lemma stream_take_lemma3 [rule_format]: + "ALL x xs. x~=UU --> stream_take n$(x && xs) = x && xs --> stream_take n$xs=xs" +apply (induct_tac n,clarsimp) +apply (drule sym, erule scons_not_empty, simp) +apply (clarify, rule stream_take_more) +apply (erule_tac x="x" in allE) +by (erule_tac x="xs" in allE,simp) + +lemma stream_take_lemma4: + "ALL x xs. stream_take n$xs=xs --> stream_take (Suc n)$(x && xs) = x && xs" +by auto + +lemma stream_take_idempotent [rule_format, simp]: + "ALL s. stream_take n$(stream_take n$s) = stream_take n$s" +apply (induct_tac n, auto) +apply (case_tac "s=UU", auto) +by (drule stream_exhaust_eq [THEN iffD1], auto) + +lemma stream_take_take_Suc [rule_format, simp]: + "ALL s. stream_take n$(stream_take (Suc n)$s) = + stream_take n$s" +apply (induct_tac n, auto) +apply (case_tac "s=UU", auto) +by (drule stream_exhaust_eq [THEN iffD1], auto) + +lemma mono_stream_take_pred: + "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==> + stream_take n$s1 << stream_take n$s2" +by (insert monofun_cfun_arg [of "stream_take (Suc n)$s1" + "stream_take (Suc n)$s2" "stream_take n"], auto) +(* +lemma mono_stream_take_pred: + "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==> + stream_take n$s1 << stream_take n$s2" +by (drule mono_stream_take [of _ _ n],simp) +*) + +lemma stream_take_lemma10 [rule_format]: + "ALL k<=n. stream_take n$s1 << stream_take n$s2 + --> stream_take k$s1 << stream_take k$s2" +apply (induct_tac n,simp,clarsimp) +apply (case_tac "k=Suc n",blast) +apply (erule_tac x="k" in allE) +by (drule mono_stream_take_pred,simp) + +lemma stream_take_le_mono : "k<=n ==> stream_take k$s1 << stream_take n$s1" +apply (insert chain_stream_take [of s1]) +by (drule chain_mono3,auto) + +lemma mono_stream_take: "s1 << s2 ==> stream_take n$s1 << stream_take n$s2" +by (simp add: monofun_cfun_arg) + +(* +lemma stream_take_prefix [simp]: "stream_take n$s << s" +apply (subgoal_tac "s=(LUB n. stream_take n$s)") + apply (erule ssubst, rule is_ub_thelub) + apply (simp only: chain_stream_take) +by (simp only: stream_reach2) +*) + +lemma stream_take_take_less:"stream_take k$(stream_take n$s) << stream_take k$s" +by (rule monofun_cfun_arg,auto) + + +(* ------------------------------------------------------------------------- *) +(* special induction rules *) +(* ------------------------------------------------------------------------- *) + + +section "induction" + +lemma stream_finite_ind: + "[| stream_finite x; P UU; !!a s. [| a ~= UU; P s |] ==> P (a && s) |] ==> P x" +apply (simp add: stream.finite_def,auto) +apply (erule subst) +by (drule stream.finite_ind [of P _ x], auto) + +lemma stream_finite_ind2: +"[| P UU; !! x. x ~= UU ==> P (x && UU); !! y z s. [| y ~= UU; z ~= UU; P s |] ==> P (y && z && s )|] ==> + !s. P (stream_take n$s)" +apply (rule nat_induct2 [of _ n],auto) +apply (case_tac "s=UU",clarsimp) +apply (drule stream_exhaust_eq [THEN iffD1],clarsimp) +apply (case_tac "s=UU",clarsimp) +apply (drule stream_exhaust_eq [THEN iffD1],clarsimp) +apply (case_tac "y=UU",clarsimp) +by (drule stream_exhaust_eq [THEN iffD1],clarsimp) + +lemma stream_ind2: +"[| adm P; P UU; !!a. a ~= UU ==> P (a && UU); !!a b s. [| a ~= UU; b ~= UU; P s |] ==> P (a && b && s) |] ==> P x" +apply (insert stream.reach [of x],erule subst) +apply (frule adm_impl_admw, rule wfix_ind, auto) +apply (rule adm_subst [THEN adm_impl_admw],auto) +apply (insert stream_finite_ind2 [of P]) +by (simp add: stream.take_def) + + + +(* ----------------------------------------------------------------------- *) +(* simplify use of coinduction *) +(* ----------------------------------------------------------------------- *) + + +section "coinduction" + +lemma stream_coind_lemma2: "!s1 s2. R s1 s2 --> ft$s1 = ft$s2 & R (rt$s1) (rt$s2) ==> stream_bisim R" +apply (simp add: stream.bisim_def,clarsimp) +apply (case_tac "x=UU",clarsimp) +apply (erule_tac x="UU" in allE,simp) +apply (case_tac "x'=UU",simp) +apply (drule stream_exhaust_eq [THEN iffD1],auto)+ +apply (case_tac "x'=UU",auto) +apply (erule_tac x="a && y" in allE) +apply (erule_tac x="UU" in allE)+ +apply (auto,drule stream_exhaust_eq [THEN iffD1],clarsimp) +apply (erule_tac x="a && y" in allE) +apply (erule_tac x="aa && ya" in allE) +by auto + + + +(* ----------------------------------------------------------------------- *) +(* theorems about stream_finite *) +(* ----------------------------------------------------------------------- *) + + +section "stream_finite" + +lemma stream_finite_UU [simp]: "stream_finite UU" +by (simp add: stream.finite_def) + +lemma stream_finite_UU_rev: "~ stream_finite s ==> s ~= UU" +by (auto simp add: stream.finite_def) + +lemma stream_finite_lemma1: "stream_finite xs ==> stream_finite (x && xs)" +apply (simp add: stream.finite_def,auto) +apply (rule_tac x="Suc n" in exI) +by (simp add: stream_take_lemma4) + +lemma stream_finite_lemma2: "[| x ~= UU; stream_finite (x && xs) |] ==> stream_finite xs" +apply (simp add: stream.finite_def, auto) +apply (rule_tac x="n" in exI) +by (erule stream_take_lemma3,simp) + +lemma stream_finite_rt_eq: "stream_finite (rt$s) = stream_finite s" +apply (rule stream.casedist [of s], auto) +apply (rule stream_finite_lemma1, simp) +by (rule stream_finite_lemma2,simp) + +lemma stream_finite_less: "stream_finite s ==> !t. t< stream_finite t" +apply (erule stream_finite_ind [of s]) +apply (clarsimp, drule eq_UU_iff [THEN iffD2], auto) +apply (case_tac "t=UU", auto) +apply (drule stream_exhaust_eq [THEN iffD1],auto) +apply (drule stream.inverts, auto) +apply (erule_tac x="y" in allE, simp) +by (rule stream_finite_lemma1, simp) + +lemma stream_take_finite [simp]: "stream_finite (stream_take n$s)" +apply (simp add: stream.finite_def) +by (rule_tac x="n" in exI,simp) + +lemma adm_not_stream_finite: "adm (%x. ~ stream_finite x)" +apply (rule admI2, auto) +apply (drule stream_finite_less,drule is_ub_thelub) +by auto + + + +(* ----------------------------------------------------------------------- *) +(* theorems about stream length *) +(* ----------------------------------------------------------------------- *) + + +section "slen" + +lemma slen_empty [simp]: "#\ = 0" +apply (simp add: slen_def stream.finite_def) +by (simp add: inat_defs Least_equality) + +lemma slen_scons [simp]: "x ~= \ ==> #(x&&xs) = iSuc (#xs)" +apply (case_tac "stream_finite (x && xs)") +apply (simp add: slen_def, auto) +apply (simp add: stream.finite_def, auto) +apply (rule Least_Suc2,auto) +apply (drule sym) +apply (drule sym scons_eq_UU [THEN iffD1],simp) +apply (erule stream_finite_lemma2, simp) +apply (simp add: slen_def, auto) +by (drule stream_finite_lemma1,auto) + +lemma slen_less_1_eq: "(#x < Fin (Suc 0)) = (x = \)" +by (rule stream.casedist [of x], auto simp del: iSuc_Fin + simp add: Fin_0 iSuc_Fin[THEN sym] i0_iless_iSuc iSuc_mono) + +lemma slen_empty_eq: "(#x = 0) = (x = \)" +by (rule stream.casedist [of x], auto) + +lemma slen_scons_eq: "(Fin (Suc n) < #x) = (? a y. x = a && y & a ~= \ & Fin n < #y)" +apply (auto, case_tac "x=UU",auto) +apply (drule stream_exhaust_eq [THEN iffD1], auto) +apply (rule_tac x="a" in exI) +apply (rule_tac x="y" in exI, simp) +by (simp add: inat_defs split:inat_splits)+ + +lemma slen_iSuc: "#x = iSuc n --> (? a y. x = a&&y & a ~= \ & #y = n)" +by (rule stream.casedist [of x], auto) + +lemma slen_stream_take_finite [simp]: "#(stream_take n$s) ~= \" +by (simp add: slen_def) + +lemma slen_scons_eq_rev: "(#x < Fin (Suc (Suc n))) = (!a y. x ~= a && y | a = \ | #y < Fin (Suc n))" +apply (rule stream.casedist [of x], auto) +apply (drule sym, drule scons_eq_UU [THEN iffD1],auto) +apply (simp add: inat_defs split:inat_splits) +apply (subgoal_tac "s=y & aa=a",simp); +apply (simp add: inat_defs split:inat_splits) +apply (case_tac "aa=UU",auto) +apply (erule_tac x="a" in allE, simp) +by (simp add: inat_defs split:inat_splits) + +lemma slen_take_lemma4 [rule_format]: + "!s. stream_take n$s ~= s --> #(stream_take n$s) = Fin n" +apply (induct_tac n,auto simp add: Fin_0) +apply (case_tac "s=UU",simp) +by (drule stream_exhaust_eq [THEN iffD1], auto) + +(* +lemma stream_take_idempotent [simp]: + "stream_take n$(stream_take n$s) = stream_take n$s" +apply (case_tac "stream_take n$s = s") +apply (auto,insert slen_take_lemma4 [of n s]); +by (auto,insert slen_take_lemma1 [of "stream_take n$s" n],simp) + +lemma stream_take_take_Suc [simp]: "stream_take n$(stream_take (Suc n)$s) = + stream_take n$s" +apply (simp add: po_eq_conv,auto) + apply (simp add: stream_take_take_less) +apply (subgoal_tac "stream_take n$s = stream_take n$(stream_take n$s)") + apply (erule ssubst) + apply (rule_tac monofun_cfun_arg) + apply (insert chain_stream_take [of s]) +by (simp add: chain_def,simp) +*) + +lemma slen_take_eq: "ALL x. (Fin n < #x) = (stream_take n\x ~= x)" +apply (induct_tac n, auto) +apply (simp add: Fin_0, clarsimp) +apply (drule not_sym) +apply (drule slen_empty_eq [THEN iffD1], simp) +apply (case_tac "x=UU", simp) +apply (drule stream_exhaust_eq [THEN iffD1], clarsimp) +apply (erule_tac x="y" in allE, auto) +apply (simp add: inat_defs split:inat_splits) +apply (case_tac "x=UU", simp) +apply (drule stream_exhaust_eq [THEN iffD1], clarsimp) +apply (erule_tac x="y" in allE, simp) +by (simp add: inat_defs split:inat_splits) + +lemma slen_take_eq_rev: "(#x <= Fin n) = (stream_take n\x = x)" +by (simp add: ile_def slen_take_eq) + +lemma slen_take_lemma1: "#x = Fin n ==> stream_take n\x = x" +by (rule slen_take_eq_rev [THEN iffD1], auto) + +lemma slen_rt_mono: "#s2 <= #s1 ==> #(rt$s2) <= #(rt$s1)" +apply (rule stream.casedist [of s1]) + by (rule stream.casedist [of s2],simp+)+ + +lemma slen_take_lemma5: "#(stream_take n$s) <= Fin n"; +apply (case_tac "stream_take n$s = s") + apply (simp add: slen_take_eq_rev) +by (simp add: slen_take_lemma4) + +lemma slen_take_lemma2: "!x. ~stream_finite x --> #(stream_take i\x) = Fin i" +apply (simp add: stream.finite_def, auto) +by (simp add: slen_take_lemma4) + +lemma slen_infinite: "stream_finite x = (#x ~= Infty)" +by (simp add: slen_def) + +lemma slen_mono_lemma: "stream_finite s ==> ALL t. s << t --> #s <= #t" +apply (erule stream_finite_ind [of s], auto) +apply (case_tac "t=UU", auto) +apply (drule eq_UU_iff [THEN iffD2]) +apply (drule scons_eq_UU [THEN iffD2], simp) +apply (drule stream_exhaust_eq [THEN iffD1], auto) +apply (erule_tac x="y" in allE, auto) +by (drule stream.inverts, auto) + +lemma slen_mono: "s << t ==> #s <= #t" +apply (case_tac "stream_finite t") +apply (frule stream_finite_less) +apply (erule_tac x="s" in allE, simp) +apply (drule slen_mono_lemma, auto) +by (simp add: slen_def) + +lemma iterate_lemma: "F$(iterate n F x) = iterate n F (F$x)" +by (insert iterate_Suc2 [of n F x], auto) + +lemma slen_rt_mult [rule_format]: "!x. Fin (i + j) <= #x --> Fin j <= #(iterate i rt x)" +apply (induct_tac i, auto) +apply (case_tac "x=UU", auto) +apply (simp add: inat_defs) +apply (drule stream_exhaust_eq [THEN iffD1], auto) +apply (erule_tac x="y" in allE, auto) +apply (simp add: inat_defs split:inat_splits) +by (simp add: iterate_lemma) + +lemma slen_take_lemma3 [rule_format]: + "!(x::'a::flat stream) y. Fin n <= #x --> x << y --> stream_take n\x = stream_take n\y" +apply (induct_tac n, auto) +apply (case_tac "x=UU", auto) +apply (simp add: inat_defs) +apply (simp add: Suc_ile_eq) +apply (case_tac "y=UU", clarsimp) +apply (drule eq_UU_iff [THEN iffD2],simp) +apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)+ +apply (erule_tac x="ya" in allE, simp) +apply (drule stream.inverts,auto) +by (drule ax_flat [rule_format], simp) + +lemma slen_strict_mono_lemma: + "stream_finite t ==> !s. #(s::'a::flat stream) = #t & s << t --> s = t" +apply (erule stream_finite_ind, auto) +apply (drule eq_UU_iff [THEN iffD2], simp) +apply (case_tac "sa=UU", auto) +apply (drule stream_exhaust_eq [THEN iffD1], clarsimp) +apply (drule stream.inverts, simp, simp, clarsimp) +by (drule ax_flat [rule_format], simp) + +lemma slen_strict_mono: "[|stream_finite t; s ~= t; s << (t::'a::flat stream) |] ==> #s < #t" +apply (intro ilessI1, auto) +apply (simp add: slen_mono) +by (drule slen_strict_mono_lemma, auto) + +lemma stream_take_Suc_neq: "stream_take (Suc n)$s ~=s ==> + stream_take n$s ~= stream_take (Suc n)$s" +apply auto +apply (subgoal_tac "stream_take n$s ~=s") + apply (insert slen_take_lemma4 [of n s],auto) +apply (rule stream.casedist [of s],simp) +apply (simp add: inat_defs split:inat_splits) +by (simp add: slen_take_lemma4) + +(* ----------------------------------------------------------------------- *) +(* theorems about smap *) +(* ----------------------------------------------------------------------- *) + + +section "smap" + +lemma smap_unfold: "smap = (\ f t. case t of x&&xs \ f$x && smap$f$xs)" +by (insert smap_def [THEN fix_eq2], auto) + +lemma smap_empty [simp]: "smap\f\\ = \" +by (subst smap_unfold, simp) + +lemma smap_scons [simp]: "x~=\ ==> smap\f\(x&&xs) = (f\x)&&(smap\f\xs)" +by (subst smap_unfold, force) + + + +(* ----------------------------------------------------------------------- *) +(* theorems about sfilter *) +(* ----------------------------------------------------------------------- *) + +section "sfilter" + +lemma sfilter_unfold: + "sfilter = (\ p s. case s of x && xs \ + If p\x then x && sfilter\p\xs else sfilter\p\xs fi)" +by (insert sfilter_def [THEN fix_eq2], auto) + +lemma strict_sfilter: "sfilter\\ = \" +apply (rule ext_cfun) +apply (subst sfilter_unfold, auto) +apply (case_tac "x=UU", auto) +by (drule stream_exhaust_eq [THEN iffD1], auto) + +lemma sfilter_empty [simp]: "sfilter\f\\ = \" +by (subst sfilter_unfold, force) + +lemma sfilter_scons [simp]: + "x ~= \ ==> sfilter\f\(x && xs) = + If f\x then x && sfilter\f\xs else sfilter\f\xs fi" +by (subst sfilter_unfold, force) + + +(* ----------------------------------------------------------------------- *) + section "i_rt" +(* ----------------------------------------------------------------------- *) + +lemma i_rt_UU [simp]: "i_rt n UU = UU" +apply (simp add: i_rt_def) +by (rule iterate.induct,auto) + +lemma i_rt_0 [simp]: "i_rt 0 s = s" +by (simp add: i_rt_def) + +lemma i_rt_Suc [simp]: "a ~= UU ==> i_rt (Suc n) (a&&s) = i_rt n s" +by (simp add: i_rt_def iterate_Suc2 del: iterate_Suc) + +lemma i_rt_Suc_forw: "i_rt (Suc n) s = i_rt n (rt$s)" +by (simp only: i_rt_def iterate_Suc2) + +lemma i_rt_Suc_back:"i_rt (Suc n) s = rt$(i_rt n s)" +by (simp only: i_rt_def,auto) + +lemma i_rt_mono: "x << s ==> i_rt n x << i_rt n s" +by (simp add: i_rt_def monofun_rt_mult) + +lemma i_rt_ij_lemma: "Fin (i + j) <= #x ==> Fin j <= #(i_rt i x)" +by (simp add: i_rt_def slen_rt_mult) + +lemma slen_i_rt_mono: "#s2 <= #s1 ==> #(i_rt n s2) <= #(i_rt n s1)" +apply (induct_tac n,auto) +apply (simp add: i_rt_Suc_back) +by (drule slen_rt_mono,simp) + +lemma i_rt_take_lemma1 [rule_format]: "ALL s. i_rt n (stream_take n$s) = UU" +apply (induct_tac n); + apply (simp add: i_rt_Suc_back,auto) +apply (case_tac "s=UU",auto) +by (drule stream_exhaust_eq [THEN iffD1],auto) + +lemma i_rt_slen: "(i_rt n s = UU) = (stream_take n$s = s)" +apply auto + apply (insert i_rt_ij_lemma [of n "Suc 0" s]); + apply (subgoal_tac "#(i_rt n s)=0") + apply (case_tac "stream_take n$s = s",simp+) + apply (insert slen_take_eq [rule_format,of n s],simp) + apply (simp add: inat_defs split:inat_splits) + apply (simp add: slen_take_eq ) +by (simp, insert i_rt_take_lemma1 [of n s],simp) + +lemma i_rt_lemma_slen: "#s=Fin n ==> i_rt n s = UU" +by (simp add: i_rt_slen slen_take_lemma1) + +lemma stream_finite_i_rt [simp]: "stream_finite (i_rt n s) = stream_finite s" +apply (induct_tac n, auto) + apply (rule stream.casedist [of "s"], auto simp del: i_rt_Suc) +by (simp add: i_rt_Suc_back stream_finite_rt_eq)+ + +lemma take_i_rt_len_lemma: "ALL sl x j t. Fin sl = #x & n <= sl & + #(stream_take n$x) = Fin t & #(i_rt n x)= Fin j + --> Fin (j + t) = #x" +apply (induct_tac n,auto) + apply (simp add: inat_defs) +apply (case_tac "x=UU",auto) + apply (simp add: inat_defs) +apply (drule stream_exhaust_eq [THEN iffD1],clarsimp) +apply (subgoal_tac "EX k. Fin k = #y",clarify) + apply (erule_tac x="k" in allE) + apply (erule_tac x="y" in allE,auto) + apply (erule_tac x="THE p. Suc p = t" in allE,auto) + apply (simp add: inat_defs split:inat_splits) + apply (simp add: inat_defs split:inat_splits) + apply (simp only: the_equality) + apply (simp add: inat_defs split:inat_splits) + apply force +by (simp add: inat_defs split:inat_splits) + +lemma take_i_rt_len: +"[| Fin sl = #x; n <= sl; #(stream_take n$x) = Fin t; #(i_rt n x) = Fin j |] ==> + Fin (j + t) = #x" +by (blast intro: take_i_rt_len_lemma [rule_format]) + + +(* ----------------------------------------------------------------------- *) + section "i_th" +(* ----------------------------------------------------------------------- *) + +lemma i_th_i_rt_step: +"[| i_th n s1 << i_th n s2; i_rt (Suc n) s1 << i_rt (Suc n) s2 |] ==> + i_rt n s1 << i_rt n s2" +apply (simp add: i_th_def i_rt_Suc_back) +apply (rule stream.casedist [of "i_rt n s1"],simp) +apply (rule stream.casedist [of "i_rt n s2"],auto) +apply (drule eq_UU_iff [THEN iffD2], simp add: scons_eq_UU) +by (intro monofun_cfun, auto) + +lemma i_th_stream_take_Suc [rule_format]: + "ALL s. i_th n (stream_take (Suc n)$s) = i_th n s" +apply (induct_tac n,auto) + apply (simp add: i_th_def) + apply (case_tac "s=UU",auto) + apply (drule stream_exhaust_eq [THEN iffD1],auto) +apply (case_tac "s=UU",simp add: i_th_def) +apply (drule stream_exhaust_eq [THEN iffD1],auto) +by (simp add: i_th_def i_rt_Suc_forw) + +lemma i_th_last: "i_th n s && UU = i_rt n (stream_take (Suc n)$s)" +apply (insert surjectiv_scons [of "i_rt n (stream_take (Suc n)$s)"]) +apply (rule i_th_stream_take_Suc [THEN subst]) +apply (simp add: i_th_def i_rt_Suc_back [symmetric]) +by (simp add: i_rt_take_lemma1) + +lemma i_th_last_eq: +"i_th n s1 = i_th n s2 ==> i_rt n (stream_take (Suc n)$s1) = i_rt n (stream_take (Suc n)$s2)" +apply (insert i_th_last [of n s1]) +apply (insert i_th_last [of n s2]) +by auto + +lemma i_th_prefix_lemma: +"[| k <= n; stream_take (Suc n)$s1 << stream_take (Suc n)$s2 |] ==> + i_th k s1 << i_th k s2" +apply (insert i_th_stream_take_Suc [of k s1, THEN sym]) +apply (insert i_th_stream_take_Suc [of k s2, THEN sym],auto) +apply (simp add: i_th_def) +apply (rule monofun_cfun, auto) +apply (rule i_rt_mono) +by (blast intro: stream_take_lemma10) + +lemma take_i_rt_prefix_lemma1: + "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==> + i_rt (Suc n) s1 << i_rt (Suc n) s2 ==> + i_rt n s1 << i_rt n s2 & stream_take n$s1 << stream_take n$s2" +apply auto + apply (insert i_th_prefix_lemma [of n n s1 s2]) + apply (rule i_th_i_rt_step,auto) +by (drule mono_stream_take_pred,simp) + +lemma take_i_rt_prefix_lemma: +"[| stream_take n$s1 << stream_take n$s2; i_rt n s1 << i_rt n s2 |] ==> s1 << s2" +apply (case_tac "n=0",simp) +apply (insert neq0_conv [of n]) +apply (insert not0_implies_Suc [of n],auto) +apply (subgoal_tac "stream_take 0$s1 << stream_take 0$s2 & + i_rt 0 s1 << i_rt 0 s2") + defer 1 + apply (rule zero_induct,blast) + apply (blast dest: take_i_rt_prefix_lemma1) +by simp + +lemma streams_prefix_lemma: "(s1 << s2) = + (stream_take n$s1 << stream_take n$s2 & i_rt n s1 << i_rt n s2)"; +apply auto + apply (simp add: monofun_cfun_arg) + apply (simp add: i_rt_mono) +by (erule take_i_rt_prefix_lemma,simp) + +lemma streams_prefix_lemma1: + "[| stream_take n$s1 = stream_take n$s2; i_rt n s1 = i_rt n s2 |] ==> s1 = s2" +apply (simp add: po_eq_conv,auto) + apply (insert streams_prefix_lemma) + by blast+ + + +(* ----------------------------------------------------------------------- *) + section "sconc" +(* ----------------------------------------------------------------------- *) + +lemma UU_sconc [simp]: " UU ooo s = s " +by (simp add: sconc_def inat_defs) + +lemma scons_neq_UU: "a~=UU ==> a && s ~=UU" +by auto + +lemma singleton_sconc [rule_format, simp]: "x~=UU --> (x && UU) ooo y = x && y" +apply (simp add: sconc_def inat_defs split:inat_splits,auto) +apply (rule someI2_ex,auto) + apply (rule_tac x="x && y" in exI,auto) +apply (simp add: i_rt_Suc_forw) +apply (case_tac "xa=UU",simp) +by (drule stream_exhaust_eq [THEN iffD1],auto) + +lemma ex_sconc [rule_format]: + "ALL k y. #x = Fin k --> (EX w. stream_take k$w = x & i_rt k w = y)" +apply (case_tac "#x") + apply (rule stream_finite_ind [of x],auto) + apply (simp add: stream.finite_def) + apply (drule slen_take_lemma1,blast) + apply (simp add: inat_defs split:inat_splits)+ +apply (erule_tac x="y" in allE,auto) +by (rule_tac x="a && w" in exI,auto) + +lemma rt_sconc1: "Fin n = #x ==> i_rt n (x ooo y) = y"; +apply (simp add: sconc_def inat_defs split:inat_splits, arith?,auto) +apply (rule someI2_ex,auto) +by (drule ex_sconc,simp) + +lemma sconc_inj2: "\Fin n = #x; x ooo y = x ooo z\ \ y = z" +apply (frule_tac y=y in rt_sconc1) +by (auto elim: rt_sconc1) + +lemma sconc_UU [simp]:"s ooo UU = s" +apply (case_tac "#s") + apply (simp add: sconc_def inat_defs) + apply (rule someI2_ex) + apply (rule_tac x="s" in exI) + apply auto + apply (drule slen_take_lemma1,auto) + apply (simp add: i_rt_lemma_slen) + apply (drule slen_take_lemma1,auto) + apply (simp add: i_rt_slen) +by (simp add: sconc_def inat_defs) + +lemma stream_take_sconc [simp]: "Fin n = #x ==> stream_take n$(x ooo y) = x" +apply (simp add: sconc_def) +apply (simp add: inat_defs split:inat_splits,auto) +apply (rule someI2_ex,auto) +by (drule ex_sconc,simp) + +lemma scons_sconc [rule_format,simp]: "a~=UU --> (a && x) ooo y = a && x ooo y" +apply (case_tac "#x",auto) + apply (simp add: sconc_def) + apply (rule someI2_ex) + apply (drule ex_sconc,simp) + apply (rule someI2_ex,auto) + apply (simp add: i_rt_Suc_forw) + apply (rule_tac x="a && x" in exI,auto) + apply (case_tac "xa=UU",auto) + apply (drule_tac s="stream_take nat$x" in scons_neq_UU) + apply (simp add: i_rt_Suc_forw) + apply (drule stream_exhaust_eq [THEN iffD1],auto) + apply (drule streams_prefix_lemma1,simp+) +by (simp add: sconc_def) + +lemma ft_sconc: "x ~= UU ==> ft$(x ooo y) = ft$x" +by (rule stream.casedist [of x],auto) + +lemma sconc_assoc: "(x ooo y) ooo z = x ooo y ooo z" +apply (case_tac "#x") + apply (rule stream_finite_ind [of x],auto simp del: scons_sconc) + apply (simp add: stream.finite_def del: scons_sconc) + apply (drule slen_take_lemma1,auto simp del: scons_sconc) + apply (case_tac "a = UU", auto) +by (simp add: sconc_def) + + +(* ----------------------------------------------------------------------- *) + +lemma sconc_mono: "y << y' ==> x ooo y << x ooo y'" +apply (case_tac "#x") + apply (rule stream_finite_ind [of "x"]) + apply (auto simp add: stream.finite_def) + apply (drule slen_take_lemma1,blast) + by (simp add: stream_prefix',auto simp add: sconc_def) + +lemma sconc_mono1 [simp]: "x << x ooo y" +by (rule sconc_mono [of UU, simplified]) + +(* ----------------------------------------------------------------------- *) + +lemma empty_sconc [simp]: "(x ooo y = UU) = (x = UU & y = UU)" +apply (case_tac "#x",auto) + apply (insert sconc_mono1 [of x y]); + by (insert eq_UU_iff [THEN iffD2, of x],auto) + +(* ----------------------------------------------------------------------- *) + +lemma rt_sconc [rule_format, simp]: "s~=UU --> rt$(s ooo x) = rt$s ooo x" +by (rule stream.casedist,auto) + +lemma i_th_sconc_lemma [rule_format]: + "ALL x y. Fin n < #x --> i_th n (x ooo y) = i_th n x" +apply (induct_tac n, auto) +apply (simp add: Fin_0 i_th_def) +apply (simp add: slen_empty_eq ft_sconc) +apply (simp add: i_th_def) +apply (case_tac "x=UU",auto) +apply (drule stream_exhaust_eq [THEN iffD1], auto) +apply (erule_tac x="ya" in allE) +by (simp add: inat_defs split:inat_splits) + + + +(* ----------------------------------------------------------------------- *) + +lemma sconc_lemma [rule_format, simp]: "ALL s. stream_take n$s ooo i_rt n s = s" +apply (induct_tac n,auto) +apply (case_tac "s=UU",auto) +by (drule stream_exhaust_eq [THEN iffD1],auto) + +(* ----------------------------------------------------------------------- *) + subsection "pointwise equality" +(* ----------------------------------------------------------------------- *) + +lemma ex_last_stream_take_scons: "stream_take (Suc n)$s = + stream_take n$s ooo i_rt n (stream_take (Suc n)$s)" +by (insert sconc_lemma [of n "stream_take (Suc n)$s"],simp) + +lemma i_th_stream_take_eq: +"!!n. ALL n. i_th n s1 = i_th n s2 ==> stream_take n$s1 = stream_take n$s2" +apply (induct_tac n,auto) +apply (subgoal_tac "stream_take (Suc na)$s1 = + stream_take na$s1 ooo i_rt na (stream_take (Suc na)$s1)") + apply (subgoal_tac "i_rt na (stream_take (Suc na)$s1) = + i_rt na (stream_take (Suc na)$s2)") + apply (subgoal_tac "stream_take (Suc na)$s2 = + stream_take na$s2 ooo i_rt na (stream_take (Suc na)$s2)") + apply (insert ex_last_stream_take_scons,simp) + apply blast + apply (erule_tac x="na" in allE) + apply (insert i_th_last_eq [of _ s1 s2]) +by blast+ + +lemma pointwise_eq_lemma[rule_format]: "ALL n. i_th n s1 = i_th n s2 ==> s1 = s2" +by (insert i_th_stream_take_eq [THEN stream.take_lemmas],blast) + +(* ----------------------------------------------------------------------- *) + subsection "finiteness" +(* ----------------------------------------------------------------------- *) + +lemma slen_sconc_finite1: + "[| #(x ooo y) = Infty; Fin n = #x |] ==> #y = Infty" +apply (case_tac "#y ~= Infty",auto) +apply (simp only: slen_infinite [symmetric]) +apply (drule_tac y=y in rt_sconc1) +apply (insert stream_finite_i_rt [of n "x ooo y"]) +by (simp add: slen_infinite) + +lemma slen_sconc_infinite1: "#x=Infty ==> #(x ooo y) = Infty" +by (simp add: sconc_def) + +lemma slen_sconc_infinite2: "#y=Infty ==> #(x ooo y) = Infty" +apply (case_tac "#x") + apply (simp add: sconc_def) + apply (rule someI2_ex) + apply (drule ex_sconc,auto) + apply (erule contrapos_pp) + apply (insert stream_finite_i_rt) + apply (simp add: slen_infinite,auto) +by (simp add: sconc_def) + +lemma sconc_finite: "(#x~=Infty & #y~=Infty) = (#(x ooo y)~=Infty)" +apply auto + apply (case_tac "#x",auto) + apply (erule contrapos_pp,simp) + apply (erule slen_sconc_finite1,simp) + apply (drule slen_sconc_infinite1 [of _ y],simp) +by (drule slen_sconc_infinite2 [of _ x],simp) + +(* ----------------------------------------------------------------------- *) + +lemma slen_sconc_mono3: "[| Fin n = #x; Fin k = #(x ooo y) |] ==> n <= k" +apply (insert slen_mono [of "x" "x ooo y"]) +by (simp add: inat_defs split: inat_splits) + +(* ----------------------------------------------------------------------- *) + subsection "finite slen" +(* ----------------------------------------------------------------------- *) + +lemma slen_sconc: "[| Fin n = #x; Fin m = #y |] ==> #(x ooo y) = Fin (n + m)" +apply (case_tac "#(x ooo y)") + apply (frule_tac y=y in rt_sconc1) + apply (insert take_i_rt_len [of "THE j. Fin j = #(x ooo y)" "x ooo y" n n m],simp) + apply (insert slen_sconc_mono3 [of n x _ y],simp) +by (insert sconc_finite [of x y],auto) + +(* ----------------------------------------------------------------------- *) + subsection "flat prefix" +(* ----------------------------------------------------------------------- *) + +lemma sconc_prefix: "(s1::'a::flat stream) << s2 ==> EX t. s1 ooo t = s2" +apply (case_tac "#s1") + apply (subgoal_tac "stream_take nat$s1 = stream_take nat$s2"); + apply (rule_tac x="i_rt nat s2" in exI) + apply (simp add: sconc_def) + apply (rule someI2_ex) + apply (drule ex_sconc) + apply (simp,clarsimp,drule streams_prefix_lemma1) + apply (simp+,rule slen_take_lemma3 [of _ s1 s2]); + apply (simp+,rule_tac x="UU" in exI) +apply (insert slen_take_lemma3 [of _ s1 s2]); +by (rule stream.take_lemmas,simp) + +(* ----------------------------------------------------------------------- *) + subsection "continuity" +(* ----------------------------------------------------------------------- *) + +lemma chain_sconc: "chain S ==> chain (%i. (x ooo S i))" +by (simp add: chain_def,auto simp add: sconc_mono) + +lemma chain_scons: "chain S ==> chain (%i. a && S i)" +apply (simp add: chain_def,auto) +by (rule monofun_cfun_arg,simp) + +lemma contlub_scons: "contlub (%x. a && x)" +by (simp add: contlub_Rep_CFun2) + +lemma contlub_scons_lemma: "chain S ==> (LUB i. a && S i) = a && (LUB i. S i)" +apply (insert contlub_scons [of a]) +by (simp only: contlub) + +lemma finite_lub_sconc: "chain Y ==> (stream_finite x) ==> + (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))" +apply (rule stream_finite_ind [of x]) + apply (auto) +apply (subgoal_tac "(LUB i. a && (s ooo Y i)) = a && (LUB i. s ooo Y i)") + by (force,blast dest: contlub_scons_lemma chain_sconc) + +lemma contlub_sconc_lemma: + "chain Y ==> (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))" +apply (case_tac "#x=Infty") + apply (simp add: sconc_def) + prefer 2 + apply (drule finite_lub_sconc,auto simp add: slen_infinite) +apply (simp add: slen_def) +apply (insert lub_const [of x] unique_lub [of _ x _]) +by (auto simp add: lub) + +lemma contlub_sconc: "contlub (%y. x ooo y)"; +by (rule contlubI, insert contlub_sconc_lemma [of _ x], simp); + +lemma monofun_sconc: "monofun (%y. x ooo y)" +by (simp add: monofun sconc_mono) + +lemma cont_sconc: "cont (%y. x ooo y)" +apply (rule monocontlub2cont) + apply (rule monofunI, simp add: sconc_mono) +by (rule contlub_sconc); + + +(* ----------------------------------------------------------------------- *) + section "constr_sconc" +(* ----------------------------------------------------------------------- *) + +lemma constr_sconc_UUs [simp]: "constr_sconc UU s = s" +by (simp add: constr_sconc_def inat_defs) + +lemma "x ooo y = constr_sconc x y" +apply (case_tac "#x") + apply (rule stream_finite_ind [of x],auto simp del: scons_sconc) + defer 1 + apply (simp add: constr_sconc_def del: scons_sconc) + apply (case_tac "#s") + apply (simp add: inat_defs) + apply (case_tac "a=UU",auto simp del: scons_sconc) + apply (simp) + apply (simp add: sconc_def) + apply (simp add: constr_sconc_def) +apply (simp add: stream.finite_def) +by (drule slen_take_lemma1,auto) + +declare eq_UU_iff [THEN sym, simp add] end - -