src/HOLCF/ex/Stream.thy
changeset 17291 94f6113fe9ed
parent 16745 5608017ee28b
child 18075 43000d7a017c
--- a/src/HOLCF/ex/Stream.thy	Tue Sep 06 19:22:31 2005 +0200
+++ b/src/HOLCF/ex/Stream.thy	Tue Sep 06 19:28:58 2005 +0200
@@ -1,53 +1,55 @@
-(*  Title: 	HOLCF/ex/Stream.thy
+(*  Title:      HOLCF/ex/Stream.thy
     ID:         $Id$
-    Author: 	Franz Regensburger, David von Oheimb, Borislav Gajanovic
-
-General Stream domain.
+    Author:     Franz Regensburger, David von Oheimb, Borislav Gajanovic
 *)
 
-theory Stream imports HOLCF Nat_Infinity begin
+header {* General Stream domain *}
+
+theory Stream
+imports HOLCF Nat_Infinity
+begin
 
 domain 'a stream = "&&" (ft::'a) (lazy rt::"'a stream") (infixr 65)
 
 consts
 
-  smap		:: "('a \<rightarrow> 'b) \<rightarrow> 'a stream \<rightarrow> 'b stream"
-  sfilter	:: "('a \<rightarrow> tr) \<rightarrow> 'a stream \<rightarrow> 'a stream"
-  slen		:: "'a stream \<Rightarrow> inat"			("#_" [1000] 1000)
+  smap          :: "('a \<rightarrow> 'b) \<rightarrow> 'a stream \<rightarrow> 'b stream"
+  sfilter       :: "('a \<rightarrow> tr) \<rightarrow> 'a stream \<rightarrow> 'a stream"
+  slen          :: "'a stream \<Rightarrow> inat"                   ("#_" [1000] 1000)
 
 defs
 
-  smap_def:	"smap	 \<equiv> fix\<cdot>(\<Lambda> h f s. case s of x && xs \<Rightarrow> f\<cdot>x && h\<cdot>f\<cdot>xs)"
-  sfilter_def:	"sfilter \<equiv> fix\<cdot>(\<Lambda> h p s. case s of x && xs \<Rightarrow> 
-				     If p\<cdot>x then x && h\<cdot>p\<cdot>xs else h\<cdot>p\<cdot>xs fi)"
-  slen_def:	"#s \<equiv> if stream_finite s 
-		      then Fin (LEAST n. stream_take n\<cdot>s = s) else \<infinity>"
+  smap_def:     "smap    \<equiv> fix\<cdot>(\<Lambda> h f s. case s of x && xs \<Rightarrow> f\<cdot>x && h\<cdot>f\<cdot>xs)"
+  sfilter_def:  "sfilter \<equiv> fix\<cdot>(\<Lambda> h p s. case s of x && xs \<Rightarrow>
+                                     If p\<cdot>x then x && h\<cdot>p\<cdot>xs else h\<cdot>p\<cdot>xs fi)"
+  slen_def:     "#s \<equiv> if stream_finite s
+                      then Fin (LEAST n. stream_take n\<cdot>s = s) else \<infinity>"
 
 (* 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) 
+  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" 
+  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)" 
+  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 
+  sconc_def: "s1 ooo s2 == case #s1 of
                        Fin n \<Rightarrow> (SOME s. (stream_take n$s=s1) & (i_rt n s = s2))
-                     | \<infinity>     \<Rightarrow> s1" 
+                     | \<infinity>     \<Rightarrow> s1"
 
-  constr_sconc_def: "constr_sconc s1 s2 == case #s1 of 
-                                             Fin n \<Rightarrow> constr_sconc' n s1 s2 
+  constr_sconc_def: "constr_sconc s1 s2 == case #s1 of
+                                             Fin n \<Rightarrow> constr_sconc' n s1 s2
                                            | \<infinity>    \<Rightarrow> s1"
-primrec 
+primrec
   constr_sconc'_0:   "constr_sconc' 0 s1 s2 = s2"
-  constr_sconc'_Suc: "constr_sconc' (Suc n) s1 s2 = ft$s1 && 
+  constr_sconc'_Suc: "constr_sconc' (Suc n) s1 s2 = ft$s1 &&
                                                     constr_sconc' n (rt$s1) s2"
 
 
@@ -76,14 +78,14 @@
   "[| 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: 
+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 = 
+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)
@@ -95,7 +97,7 @@
 by (insert stream_prefix' [of y "x&&xs" ys],force)
 *)
 
-lemma stream_flat_prefix: 
+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)
@@ -144,7 +146,7 @@
 (* ----------------------------------------------------------------------- *)
 
 
-section "stream_take";
+section "stream_take"
 
 
 lemma stream_reach2: "(LUB i. stream_take i$s) = s"
@@ -154,7 +156,7 @@
 by (simp add: chain_iterate)
 
 lemma chain_stream_take: "chain (%i. stream_take i$s)"
-apply (rule chainI) 
+apply (rule chainI)
 apply (rule monofun_cfun_fun)
 apply (simp add: stream.take_def del: iterate_Suc)
 by (rule chainE, simp add: chain_iterate)
@@ -165,13 +167,13 @@
 apply (rule is_ub_thelub)
 by (simp only: chain_stream_take)
 
-lemma stream_take_more [rule_format]: 
+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]: 
+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)*)
@@ -179,37 +181,37 @@
 apply (erule_tac x="x" in allE)
 by (erule_tac x="xs" in allE,simp)
 
-lemma stream_take_lemma4: 
+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]: 
+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) = 
+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: 
+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" 
+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: 
+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 
+  "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)
@@ -242,14 +244,14 @@
 
 section "induction"
 
-lemma stream_finite_ind: 
+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 )|] ==> 
+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)
@@ -259,7 +261,7 @@
 apply (case_tac "y=UU",clarsimp)
 by (drule stream_exhaust_eq [THEN iffD1],clarsimp)
 
-lemma stream_ind2: 
+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)
@@ -363,7 +365,7 @@
 by (drule stream_finite_lemma1,auto)
 
 lemma slen_less_1_eq: "(#x < Fin (Suc 0)) = (x = \<bottom>)"
-by (rule stream.casedist [of x], auto simp del: iSuc_Fin 
+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 = \<bottom>)"
@@ -386,26 +388,26 @@
 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 (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]: 
+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]: 
+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) = 
+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)
@@ -440,7 +442,7 @@
 apply (rule stream.casedist [of s1])
  by (rule stream.casedist [of s2],simp+)+
 
-lemma slen_take_lemma5: "#(stream_take n$s) <= Fin n"; 
+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)
@@ -463,12 +465,12 @@
 
 lemma slen_mono: "s << t ==> #s <= #t"
 apply (case_tac "stream_finite t")
-apply (frule stream_finite_less) 
+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)" 
+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)"
@@ -480,7 +482,7 @@
 apply (simp add: inat_defs split:inat_splits)
 by (simp add: iterate_lemma)
 
-lemma slen_take_lemma3 [rule_format]: 
+lemma slen_take_lemma3 [rule_format]:
   "!(x::'a::flat stream) y. Fin n <= #x --> x << y --> stream_take n\<cdot>x = stream_take n\<cdot>y"
 apply (induct_tac n, auto)
 apply (case_tac "x=UU", auto)
@@ -493,7 +495,7 @@
 apply (drule stream.inverts,auto)
 by (drule ax_flat [rule_format], simp)
 
-lemma slen_strict_mono_lemma: 
+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)
@@ -507,7 +509,7 @@
 apply (simp add: slen_mono)
 by (drule slen_strict_mono_lemma, auto)
 
-lemma stream_take_Suc_neq: "stream_take (Suc n)$s ~=s ==> 
+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")
@@ -540,7 +542,7 @@
 
 section "sfilter"
 
-lemma sfilter_unfold: 
+lemma sfilter_unfold:
  "sfilter = (\<Lambda> p s. case s of x && xs \<Rightarrow>
   If p\<cdot>x then x && sfilter\<cdot>p\<cdot>xs else sfilter\<cdot>p\<cdot>xs fi)"
 by (insert sfilter_def [THEN fix_eq2], auto)
@@ -554,9 +556,9 @@
 lemma sfilter_empty [simp]: "sfilter\<cdot>f\<cdot>\<bottom> = \<bottom>"
 by (subst sfilter_unfold, force)
 
-lemma sfilter_scons [simp]: 
-  "x ~= \<bottom> ==> sfilter\<cdot>f\<cdot>(x && xs) = 
-                           If f\<cdot>x then x && sfilter\<cdot>f\<cdot>xs else sfilter\<cdot>f\<cdot>xs fi" 
+lemma sfilter_scons [simp]:
+  "x ~= \<bottom> ==> sfilter\<cdot>f\<cdot>(x && xs) =
+                           If f\<cdot>x then x && sfilter\<cdot>f\<cdot>xs else sfilter\<cdot>f\<cdot>xs fi"
 by (subst sfilter_unfold, force)
 
 
@@ -592,14 +594,14 @@
 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 (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 (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)
@@ -616,7 +618,7 @@
 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 
+                            #(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)
@@ -634,7 +636,7 @@
  apply force
 by (simp add: inat_defs split:inat_splits)
 
-lemma take_i_rt_len: 
+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])
@@ -645,7 +647,7 @@
 (* ----------------------------------------------------------------------- *)
 
 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_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)
@@ -653,7 +655,7 @@
 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]: 
+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)
@@ -669,14 +671,14 @@
 apply (simp add: i_th_def  i_rt_Suc_back [symmetric])
 by (simp add: i_rt_take_lemma1)
 
-lemma i_th_last_eq: 
+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 |] ==> 
+"[| 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)
@@ -685,29 +687,29 @@
 apply (rule i_rt_mono)
 by (blast intro: stream_take_lemma10)
 
-lemma take_i_rt_prefix_lemma1: 
+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 (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: 
+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 & 
+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)"; 
+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)
@@ -738,7 +740,7 @@
 apply (case_tac "xa=UU",simp)
 by (drule stream_exhaust_eq [THEN iffD1],auto)
 
-lemma ex_sconc [rule_format]: 
+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)
@@ -748,7 +750,7 @@
 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"; 
+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)
@@ -777,7 +779,7 @@
 
 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 (simp add: sconc_def)
  apply (rule someI2_ex)
   apply (drule ex_sconc,simp)
  apply (rule someI2_ex,auto)
@@ -818,7 +820,7 @@
 
 lemma empty_sconc [simp]: "(x ooo y = UU) = (x = UU & y = UU)"
 apply (case_tac "#x",auto)
-   apply (insert sconc_mono1 [of x y]);
+   apply (insert sconc_mono1 [of x y])
    by (insert eq_UU_iff [THEN iffD2, of x],auto)
 
 (* ----------------------------------------------------------------------- *)
@@ -826,7 +828,7 @@
 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]: 
+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)
@@ -850,18 +852,18 @@
    subsection "pointwise equality"
 (* ----------------------------------------------------------------------- *)
 
-lemma ex_last_stream_take_scons: "stream_take (Suc n)$s = 
+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: 
+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) = 
+ 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 = 
+  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
@@ -928,15 +930,15 @@
 
 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 (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 slen_take_lemma3 [of _ s1 s2])
   apply (simp+,rule_tac x="UU" in exI)
-apply (insert slen_take_lemma3 [of _ s1 s2]);
+apply (insert slen_take_lemma3 [of _ s1 s2])
 by (rule stream.take_lemmas,simp)
 
 (* ----------------------------------------------------------------------- *)
@@ -957,14 +959,14 @@
 apply (insert contlub_scons [of a])
 by (simp only: contlub_def)
 
-lemma finite_lub_sconc: "chain Y ==> (stream_finite x) ==> 
+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: 
+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)
@@ -974,8 +976,8 @@
 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 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_def sconc_mono)
@@ -983,7 +985,7 @@
 lemma cont_sconc: "cont (%y. x ooo y)"
 apply (rule  monocontlub2cont)
  apply (rule monofunI, simp add: sconc_mono)
-by (rule contlub_sconc);
+by (rule contlub_sconc)
 
 
 (* ----------------------------------------------------------------------- *)