src/HOLCF/ex/Stream.thy
author huffman
Wed, 17 Feb 2010 09:08:58 -0800
changeset 35169 31cbcb019003
parent 34928 156925dd67af
child 35174 e15040ae75d7
permissions -rw-r--r--
fix warnings about duplicate simp rules
     1 (*  Title:      HOLCF/ex/Stream.thy
     2     ID:         $Id$
     3     Author:     Franz Regensburger, David von Oheimb, Borislav Gajanovic
     4 *)
     5 
     6 header {* General Stream domain *}
     7 
     8 theory Stream
     9 imports HOLCF Nat_Infinity
    10 begin
    11 
    12 domain 'a stream = scons (ft::'a) (lazy rt::"'a stream") (infixr "&&" 65)
    13 
    14 definition
    15   smap :: "('a \<rightarrow> 'b) \<rightarrow> 'a stream \<rightarrow> 'b stream" where
    16   "smap = fix\<cdot>(\<Lambda> h f s. case s of x && xs \<Rightarrow> f\<cdot>x && h\<cdot>f\<cdot>xs)"
    17 
    18 definition
    19   sfilter :: "('a \<rightarrow> tr) \<rightarrow> 'a stream \<rightarrow> 'a stream" where
    20   "sfilter = fix\<cdot>(\<Lambda> h p s. case s of x && xs \<Rightarrow>
    21                                      If p\<cdot>x then x && h\<cdot>p\<cdot>xs else h\<cdot>p\<cdot>xs fi)"
    22 
    23 definition
    24   slen :: "'a stream \<Rightarrow> inat"  ("#_" [1000] 1000) where
    25   "#s = (if stream_finite s then Fin (LEAST n. stream_take n\<cdot>s = s) else \<infinity>)"
    26 
    27 
    28 (* concatenation *)
    29 
    30 definition
    31   i_rt :: "nat => 'a stream => 'a stream" where (* chops the first i elements *)
    32   "i_rt = (%i s. iterate i$rt$s)"
    33 
    34 definition
    35   i_th :: "nat => 'a stream => 'a" where (* the i-th element *)
    36   "i_th = (%i s. ft$(i_rt i s))"
    37 
    38 definition
    39   sconc :: "'a stream => 'a stream => 'a stream"  (infixr "ooo" 65) where
    40   "s1 ooo s2 = (case #s1 of
    41                   Fin n \<Rightarrow> (SOME s. (stream_take n$s=s1) & (i_rt n s = s2))
    42                | \<infinity>     \<Rightarrow> s1)"
    43 
    44 primrec constr_sconc' :: "nat => 'a stream => 'a stream => 'a stream"
    45 where
    46   constr_sconc'_0:   "constr_sconc' 0 s1 s2 = s2"
    47 | constr_sconc'_Suc: "constr_sconc' (Suc n) s1 s2 = ft$s1 &&
    48                                                     constr_sconc' n (rt$s1) s2"
    49 
    50 definition
    51   constr_sconc  :: "'a stream => 'a stream => 'a stream" where (* constructive *)
    52   "constr_sconc s1 s2 = (case #s1 of
    53                           Fin n \<Rightarrow> constr_sconc' n s1 s2
    54                         | \<infinity>    \<Rightarrow> s1)"
    55 
    56 
    57 (* ----------------------------------------------------------------------- *)
    58 (* theorems about scons                                                    *)
    59 (* ----------------------------------------------------------------------- *)
    60 
    61 
    62 section "scons"
    63 
    64 lemma scons_eq_UU: "(a && s = UU) = (a = UU)"
    65 by simp
    66 
    67 lemma scons_not_empty: "[| a && x = UU; a ~= UU |] ==> R"
    68 by simp
    69 
    70 lemma stream_exhaust_eq: "(x ~= UU) = (EX a y. a ~= UU &  x = a && y)"
    71 by (auto,insert stream.exhaust [of x],auto)
    72 
    73 lemma stream_neq_UU: "x~=UU ==> EX a a_s. x=a&&a_s & a~=UU"
    74 by (simp add: stream_exhaust_eq,auto)
    75 
    76 lemma stream_inject_eq [simp]:
    77   "[| a ~= UU; b ~= UU |] ==> (a && s = b && t) = (a = b &  s = t)"
    78 by (insert stream.injects [of a s b t], auto)
    79 
    80 lemma stream_prefix:
    81   "[| a && s << t; a ~= UU  |] ==> EX b tt. t = b && tt &  b ~= UU &  s << tt"
    82 by (insert stream.exhaust [of t], auto)
    83 
    84 lemma stream_prefix':
    85   "b ~= UU ==> x << b && z =
    86    (x = UU |  (EX a y. x = a && y &  a ~= UU &  a << b &  y << z))"
    87 apply (case_tac "x=UU",auto)
    88 by (drule stream_exhaust_eq [THEN iffD1],auto)
    89 
    90 
    91 (*
    92 lemma stream_prefix1: "[| x<<y; xs<<ys |] ==> x&&xs << y&&ys"
    93 by (insert stream_prefix' [of y "x&&xs" ys],force)
    94 *)
    95 
    96 lemma stream_flat_prefix:
    97   "[| x && xs << y && ys; (x::'a::flat) ~= UU|] ==> x = y & xs << ys"
    98 apply (case_tac "y=UU",auto)
    99 by (drule ax_flat,simp)
   100 
   101 
   102 
   103 
   104 (* ----------------------------------------------------------------------- *)
   105 (* theorems about stream_when                                              *)
   106 (* ----------------------------------------------------------------------- *)
   107 
   108 section "stream_when"
   109 
   110 
   111 lemma stream_when_strictf: "stream_when$UU$s=UU"
   112 by (rule stream.casedist [of s], auto)
   113 
   114 
   115 
   116 (* ----------------------------------------------------------------------- *)
   117 (* theorems about ft and rt                                                *)
   118 (* ----------------------------------------------------------------------- *)
   119 
   120 
   121 section "ft & rt"
   122 
   123 
   124 lemma ft_defin: "s~=UU ==> ft$s~=UU"
   125 by (drule stream_exhaust_eq [THEN iffD1],auto)
   126 
   127 lemma rt_strict_rev: "rt$s~=UU ==> s~=UU"
   128 by auto
   129 
   130 lemma surjectiv_scons: "(ft$s)&&(rt$s)=s"
   131 by (rule stream.casedist [of s], auto)
   132 
   133 lemma monofun_rt_mult: "x << s ==> iterate i$rt$x << iterate i$rt$s"
   134 by (rule monofun_cfun_arg)
   135 
   136 
   137 
   138 (* ----------------------------------------------------------------------- *)
   139 (* theorems about stream_take                                              *)
   140 (* ----------------------------------------------------------------------- *)
   141 
   142 
   143 section "stream_take"
   144 
   145 
   146 lemma stream_reach2: "(LUB i. stream_take i$s) = s"
   147 apply (insert stream.reach [of s], erule subst) back
   148 apply (simp add: fix_def2 stream.take_def)
   149 apply (insert contlub_cfun_fun [of "%i. iterate i$stream_copy$UU" s,THEN sym])
   150 by simp
   151 
   152 lemma chain_stream_take: "chain (%i. stream_take i$s)"
   153 apply (rule chainI)
   154 apply (rule monofun_cfun_fun)
   155 apply (simp add: stream.take_def del: iterate_Suc)
   156 by (rule chainE, simp)
   157 
   158 lemma stream_take_prefix [simp]: "stream_take n$s << s"
   159 apply (insert stream_reach2 [of s])
   160 apply (erule subst) back
   161 apply (rule is_ub_thelub)
   162 by (simp only: chain_stream_take)
   163 
   164 lemma stream_take_more [rule_format]:
   165   "ALL x. stream_take n$x = x --> stream_take (Suc n)$x = x"
   166 apply (induct_tac n,auto)
   167 apply (case_tac "x=UU",auto)
   168 by (drule stream_exhaust_eq [THEN iffD1],auto)
   169 
   170 lemma stream_take_lemma3 [rule_format]:
   171   "ALL x xs. x~=UU --> stream_take n$(x && xs) = x && xs --> stream_take n$xs=xs"
   172 apply (induct_tac n,clarsimp)
   173 (*apply (drule sym, erule scons_not_empty, simp)*)
   174 apply (clarify, rule stream_take_more)
   175 apply (erule_tac x="x" in allE)
   176 by (erule_tac x="xs" in allE,simp)
   177 
   178 lemma stream_take_lemma4:
   179   "ALL x xs. stream_take n$xs=xs --> stream_take (Suc n)$(x && xs) = x && xs"
   180 by auto
   181 
   182 lemma stream_take_idempotent [rule_format, simp]:
   183  "ALL s. stream_take n$(stream_take n$s) = stream_take n$s"
   184 apply (induct_tac n, auto)
   185 apply (case_tac "s=UU", auto)
   186 by (drule stream_exhaust_eq [THEN iffD1], auto)
   187 
   188 lemma stream_take_take_Suc [rule_format, simp]:
   189   "ALL s. stream_take n$(stream_take (Suc n)$s) =
   190                                     stream_take n$s"
   191 apply (induct_tac n, auto)
   192 apply (case_tac "s=UU", auto)
   193 by (drule stream_exhaust_eq [THEN iffD1], auto)
   194 
   195 lemma mono_stream_take_pred:
   196   "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==>
   197                        stream_take n$s1 << stream_take n$s2"
   198 by (insert monofun_cfun_arg [of "stream_take (Suc n)$s1"
   199   "stream_take (Suc n)$s2" "stream_take n"], auto)
   200 (*
   201 lemma mono_stream_take_pred:
   202   "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==>
   203                        stream_take n$s1 << stream_take n$s2"
   204 by (drule mono_stream_take [of _ _ n],simp)
   205 *)
   206 
   207 lemma stream_take_lemma10 [rule_format]:
   208   "ALL k<=n. stream_take n$s1 << stream_take n$s2
   209                              --> stream_take k$s1 << stream_take k$s2"
   210 apply (induct_tac n,simp,clarsimp)
   211 apply (case_tac "k=Suc n",blast)
   212 apply (erule_tac x="k" in allE)
   213 by (drule mono_stream_take_pred,simp)
   214 
   215 lemma stream_take_le_mono : "k<=n ==> stream_take k$s1 << stream_take n$s1"
   216 apply (insert chain_stream_take [of s1])
   217 by (drule chain_mono,auto)
   218 
   219 lemma mono_stream_take: "s1 << s2 ==> stream_take n$s1 << stream_take n$s2"
   220 by (simp add: monofun_cfun_arg)
   221 
   222 (*
   223 lemma stream_take_prefix [simp]: "stream_take n$s << s"
   224 apply (subgoal_tac "s=(LUB n. stream_take n$s)")
   225  apply (erule ssubst, rule is_ub_thelub)
   226  apply (simp only: chain_stream_take)
   227 by (simp only: stream_reach2)
   228 *)
   229 
   230 lemma stream_take_take_less:"stream_take k$(stream_take n$s) << stream_take k$s"
   231 by (rule monofun_cfun_arg,auto)
   232 
   233 
   234 (* ------------------------------------------------------------------------- *)
   235 (* special induction rules                                                   *)
   236 (* ------------------------------------------------------------------------- *)
   237 
   238 
   239 section "induction"
   240 
   241 lemma stream_finite_ind:
   242  "[| stream_finite x; P UU; !!a s. [| a ~= UU; P s |] ==> P (a && s) |] ==> P x"
   243 apply (simp add: stream.finite_def,auto)
   244 apply (erule subst)
   245 by (drule stream.finite_ind [of P _ x], auto)
   246 
   247 lemma stream_finite_ind2:
   248 "[| P UU; !! x. x ~= UU ==> P (x && UU); !! y z s. [| y ~= UU; z ~= UU; P s |] ==> P (y && z && s )|] ==>
   249                                  !s. P (stream_take n$s)"
   250 apply (rule nat_less_induct [of _ n],auto)
   251 apply (case_tac n, auto) 
   252 apply (case_tac nat, auto) 
   253 apply (case_tac "s=UU",clarsimp)
   254 apply (drule stream_exhaust_eq [THEN iffD1],clarsimp)
   255 apply (case_tac "s=UU",clarsimp)
   256 apply (drule stream_exhaust_eq [THEN iffD1],clarsimp)
   257 apply (case_tac "y=UU",clarsimp)
   258 by (drule stream_exhaust_eq [THEN iffD1],clarsimp)
   259 
   260 lemma stream_ind2:
   261 "[| 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"
   262 apply (insert stream.reach [of x],erule subst)
   263 apply (frule adm_impl_admw, rule wfix_ind, auto)
   264 apply (rule adm_subst [THEN adm_impl_admw],auto)
   265 apply (insert stream_finite_ind2 [of P])
   266 by (simp add: stream.take_def)
   267 
   268 
   269 
   270 (* ----------------------------------------------------------------------- *)
   271 (* simplify use of coinduction                                             *)
   272 (* ----------------------------------------------------------------------- *)
   273 
   274 
   275 section "coinduction"
   276 
   277 lemma stream_coind_lemma2: "!s1 s2. R s1 s2 --> ft$s1 = ft$s2 &  R (rt$s1) (rt$s2) ==> stream_bisim R"
   278  apply (simp add: stream.bisim_def,clarsimp)
   279  apply (case_tac "x=UU",clarsimp)
   280   apply (erule_tac x="UU" in allE,simp)
   281   apply (case_tac "x'=UU",simp)
   282   apply (drule stream_exhaust_eq [THEN iffD1],auto)+
   283  apply (case_tac "x'=UU",auto)
   284   apply (erule_tac x="a && y" in allE)
   285   apply (erule_tac x="UU" in allE)+
   286   apply (auto,drule stream_exhaust_eq [THEN iffD1],clarsimp)
   287  apply (erule_tac x="a && y" in allE)
   288  apply (erule_tac x="aa && ya" in allE) back
   289 by auto
   290 
   291 
   292 
   293 (* ----------------------------------------------------------------------- *)
   294 (* theorems about stream_finite                                            *)
   295 (* ----------------------------------------------------------------------- *)
   296 
   297 
   298 section "stream_finite"
   299 
   300 lemma stream_finite_UU [simp]: "stream_finite UU"
   301 by (simp add: stream.finite_def)
   302 
   303 lemma stream_finite_UU_rev: "~  stream_finite s ==> s ~= UU"
   304 by (auto simp add: stream.finite_def)
   305 
   306 lemma stream_finite_lemma1: "stream_finite xs ==> stream_finite (x && xs)"
   307 apply (simp add: stream.finite_def,auto)
   308 apply (rule_tac x="Suc n" in exI)
   309 by (simp add: stream_take_lemma4)
   310 
   311 lemma stream_finite_lemma2: "[| x ~= UU; stream_finite (x && xs) |] ==> stream_finite xs"
   312 apply (simp add: stream.finite_def, auto)
   313 apply (rule_tac x="n" in exI)
   314 by (erule stream_take_lemma3,simp)
   315 
   316 lemma stream_finite_rt_eq: "stream_finite (rt$s) = stream_finite s"
   317 apply (rule stream.casedist [of s], auto)
   318 apply (rule stream_finite_lemma1, simp)
   319 by (rule stream_finite_lemma2,simp)
   320 
   321 lemma stream_finite_less: "stream_finite s ==> !t. t<<s --> stream_finite t"
   322 apply (erule stream_finite_ind [of s], auto)
   323 apply (case_tac "t=UU", auto)
   324 apply (drule stream_exhaust_eq [THEN iffD1],auto)
   325 apply (erule_tac x="y" in allE, simp)
   326 by (rule stream_finite_lemma1, simp)
   327 
   328 lemma stream_take_finite [simp]: "stream_finite (stream_take n$s)"
   329 apply (simp add: stream.finite_def)
   330 by (rule_tac x="n" in exI,simp)
   331 
   332 lemma adm_not_stream_finite: "adm (%x. ~ stream_finite x)"
   333 apply (rule adm_upward)
   334 apply (erule contrapos_nn)
   335 apply (erule (1) stream_finite_less [rule_format])
   336 done
   337 
   338 
   339 
   340 (* ----------------------------------------------------------------------- *)
   341 (* theorems about stream length                                            *)
   342 (* ----------------------------------------------------------------------- *)
   343 
   344 
   345 section "slen"
   346 
   347 lemma slen_empty [simp]: "#\<bottom> = 0"
   348 by (simp add: slen_def stream.finite_def zero_inat_def Least_equality)
   349 
   350 lemma slen_scons [simp]: "x ~= \<bottom> ==> #(x&&xs) = iSuc (#xs)"
   351 apply (case_tac "stream_finite (x && xs)")
   352 apply (simp add: slen_def, auto)
   353 apply (simp add: stream.finite_def, auto simp add: iSuc_Fin)
   354 apply (rule Least_Suc2, auto)
   355 (*apply (drule sym)*)
   356 (*apply (drule sym scons_eq_UU [THEN iffD1],simp)*)
   357 apply (erule stream_finite_lemma2, simp)
   358 apply (simp add: slen_def, auto)
   359 by (drule stream_finite_lemma1,auto)
   360 
   361 lemma slen_less_1_eq: "(#x < Fin (Suc 0)) = (x = \<bottom>)"
   362 by (rule stream.casedist [of x], auto simp del: iSuc_Fin
   363     simp add: Fin_0 iSuc_Fin[THEN sym] i0_iless_iSuc iSuc_mono)
   364 
   365 lemma slen_empty_eq: "(#x = 0) = (x = \<bottom>)"
   366 by (rule stream.casedist [of x], auto)
   367 
   368 lemma slen_scons_eq: "(Fin (Suc n) < #x) = (? a y. x = a && y &  a ~= \<bottom> &  Fin n < #y)"
   369 apply (auto, case_tac "x=UU",auto)
   370 apply (drule stream_exhaust_eq [THEN iffD1], auto)
   371 apply (case_tac "#y") apply simp_all
   372 apply (case_tac "#y") apply simp_all
   373 done
   374 
   375 lemma slen_iSuc: "#x = iSuc n --> (? a y. x = a&&y &  a ~= \<bottom> &  #y = n)"
   376 by (rule stream.casedist [of x], auto)
   377 
   378 lemma slen_stream_take_finite [simp]: "#(stream_take n$s) ~= \<infinity>"
   379 by (simp add: slen_def)
   380 
   381 lemma slen_scons_eq_rev: "(#x < Fin (Suc (Suc n))) = (!a y. x ~= a && y |  a = \<bottom> |  #y < Fin (Suc n))"
   382  apply (rule stream.casedist [of x], auto)
   383    apply (simp add: zero_inat_def)
   384   apply (case_tac "#s") apply (simp_all add: iSuc_Fin)
   385  apply (case_tac "#s") apply (simp_all add: iSuc_Fin)
   386 done
   387 
   388 lemma slen_take_lemma4 [rule_format]:
   389   "!s. stream_take n$s ~= s --> #(stream_take n$s) = Fin n"
   390 apply (induct n, auto simp add: Fin_0)
   391 apply (case_tac "s=UU", simp)
   392 by (drule stream_exhaust_eq [THEN iffD1], auto simp add: iSuc_Fin)
   393 
   394 (*
   395 lemma stream_take_idempotent [simp]:
   396  "stream_take n$(stream_take n$s) = stream_take n$s"
   397 apply (case_tac "stream_take n$s = s")
   398 apply (auto,insert slen_take_lemma4 [of n s]);
   399 by (auto,insert slen_take_lemma1 [of "stream_take n$s" n],simp)
   400 
   401 lemma stream_take_take_Suc [simp]: "stream_take n$(stream_take (Suc n)$s) =
   402                                     stream_take n$s"
   403 apply (simp add: po_eq_conv,auto)
   404  apply (simp add: stream_take_take_less)
   405 apply (subgoal_tac "stream_take n$s = stream_take n$(stream_take n$s)")
   406  apply (erule ssubst)
   407  apply (rule_tac monofun_cfun_arg)
   408  apply (insert chain_stream_take [of s])
   409 by (simp add: chain_def,simp)
   410 *)
   411 
   412 lemma slen_take_eq: "ALL x. (Fin n < #x) = (stream_take n\<cdot>x ~= x)"
   413 apply (induct_tac n, auto)
   414 apply (simp add: Fin_0, clarsimp)
   415 apply (drule not_sym)
   416 apply (drule slen_empty_eq [THEN iffD1], simp)
   417 apply (case_tac "x=UU", simp)
   418 apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)
   419 apply (erule_tac x="y" in allE, auto)
   420 apply (simp_all add: not_less iSuc_Fin)
   421 apply (case_tac "#y") apply simp_all
   422 apply (case_tac "x=UU", simp)
   423 apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)
   424 apply (erule_tac x="y" in allE, simp)
   425 apply (case_tac "#y") by simp_all
   426 
   427 lemma slen_take_eq_rev: "(#x <= Fin n) = (stream_take n\<cdot>x = x)"
   428 by (simp add: linorder_not_less [symmetric] slen_take_eq)
   429 
   430 lemma slen_take_lemma1: "#x = Fin n ==> stream_take n\<cdot>x = x"
   431 by (rule slen_take_eq_rev [THEN iffD1], auto)
   432 
   433 lemma slen_rt_mono: "#s2 <= #s1 ==> #(rt$s2) <= #(rt$s1)"
   434 apply (rule stream.casedist [of s1])
   435  by (rule stream.casedist [of s2],simp+)+
   436 
   437 lemma slen_take_lemma5: "#(stream_take n$s) <= Fin n"
   438 apply (case_tac "stream_take n$s = s")
   439  apply (simp add: slen_take_eq_rev)
   440 by (simp add: slen_take_lemma4)
   441 
   442 lemma slen_take_lemma2: "!x. ~stream_finite x --> #(stream_take i\<cdot>x) = Fin i"
   443 apply (simp add: stream.finite_def, auto)
   444 by (simp add: slen_take_lemma4)
   445 
   446 lemma slen_infinite: "stream_finite x = (#x ~= Infty)"
   447 by (simp add: slen_def)
   448 
   449 lemma slen_mono_lemma: "stream_finite s ==> ALL t. s << t --> #s <= #t"
   450 apply (erule stream_finite_ind [of s], auto)
   451 apply (case_tac "t=UU", auto)
   452 apply (drule stream_exhaust_eq [THEN iffD1], auto)
   453 done
   454 
   455 lemma slen_mono: "s << t ==> #s <= #t"
   456 apply (case_tac "stream_finite t")
   457 apply (frule stream_finite_less)
   458 apply (erule_tac x="s" in allE, simp)
   459 apply (drule slen_mono_lemma, auto)
   460 by (simp add: slen_def)
   461 
   462 lemma iterate_lemma: "F$(iterate n$F$x) = iterate n$F$(F$x)"
   463 by (insert iterate_Suc2 [of n F x], auto)
   464 
   465 lemma slen_rt_mult [rule_format]: "!x. Fin (i + j) <= #x --> Fin j <= #(iterate i$rt$x)"
   466 apply (induct i, auto)
   467 apply (case_tac "x=UU", auto simp add: zero_inat_def)
   468 apply (drule stream_exhaust_eq [THEN iffD1], auto)
   469 apply (erule_tac x="y" in allE, auto)
   470 apply (simp add: not_le) apply (case_tac "#y") apply (simp_all add: iSuc_Fin)
   471 by (simp add: iterate_lemma)
   472 
   473 lemma slen_take_lemma3 [rule_format]:
   474   "!(x::'a::flat stream) y. Fin n <= #x --> x << y --> stream_take n\<cdot>x = stream_take n\<cdot>y"
   475 apply (induct_tac n, auto)
   476 apply (case_tac "x=UU", auto)
   477 apply (simp add: zero_inat_def)
   478 apply (simp add: Suc_ile_eq)
   479 apply (case_tac "y=UU", clarsimp)
   480 apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)+
   481 apply (erule_tac x="ya" in allE, simp)
   482 by (drule ax_flat, simp)
   483 
   484 lemma slen_strict_mono_lemma:
   485   "stream_finite t ==> !s. #(s::'a::flat stream) = #t &  s << t --> s = t"
   486 apply (erule stream_finite_ind, auto)
   487 apply (case_tac "sa=UU", auto)
   488 apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)
   489 by (drule ax_flat, simp)
   490 
   491 lemma slen_strict_mono: "[|stream_finite t; s ~= t; s << (t::'a::flat stream) |] ==> #s < #t"
   492 by (auto simp add: slen_mono less_le dest: slen_strict_mono_lemma)
   493 
   494 lemma stream_take_Suc_neq: "stream_take (Suc n)$s ~=s ==>
   495                      stream_take n$s ~= stream_take (Suc n)$s"
   496 apply auto
   497 apply (subgoal_tac "stream_take n$s ~=s")
   498  apply (insert slen_take_lemma4 [of n s],auto)
   499 apply (rule stream.casedist [of s],simp)
   500 by (simp add: slen_take_lemma4 iSuc_Fin)
   501 
   502 (* ----------------------------------------------------------------------- *)
   503 (* theorems about smap                                                     *)
   504 (* ----------------------------------------------------------------------- *)
   505 
   506 
   507 section "smap"
   508 
   509 lemma smap_unfold: "smap = (\<Lambda> f t. case t of x&&xs \<Rightarrow> f$x && smap$f$xs)"
   510 by (insert smap_def [where 'a='a and 'b='b, THEN eq_reflection, THEN fix_eq2], auto)
   511 
   512 lemma smap_empty [simp]: "smap\<cdot>f\<cdot>\<bottom> = \<bottom>"
   513 by (subst smap_unfold, simp)
   514 
   515 lemma smap_scons [simp]: "x~=\<bottom> ==> smap\<cdot>f\<cdot>(x&&xs) = (f\<cdot>x)&&(smap\<cdot>f\<cdot>xs)"
   516 by (subst smap_unfold, force)
   517 
   518 
   519 
   520 (* ----------------------------------------------------------------------- *)
   521 (* theorems about sfilter                                                  *)
   522 (* ----------------------------------------------------------------------- *)
   523 
   524 section "sfilter"
   525 
   526 lemma sfilter_unfold:
   527  "sfilter = (\<Lambda> p s. case s of x && xs \<Rightarrow>
   528   If p\<cdot>x then x && sfilter\<cdot>p\<cdot>xs else sfilter\<cdot>p\<cdot>xs fi)"
   529 by (insert sfilter_def [where 'a='a, THEN eq_reflection, THEN fix_eq2], auto)
   530 
   531 lemma strict_sfilter: "sfilter\<cdot>\<bottom> = \<bottom>"
   532 apply (rule ext_cfun)
   533 apply (subst sfilter_unfold, auto)
   534 apply (case_tac "x=UU", auto)
   535 by (drule stream_exhaust_eq [THEN iffD1], auto)
   536 
   537 lemma sfilter_empty [simp]: "sfilter\<cdot>f\<cdot>\<bottom> = \<bottom>"
   538 by (subst sfilter_unfold, force)
   539 
   540 lemma sfilter_scons [simp]:
   541   "x ~= \<bottom> ==> sfilter\<cdot>f\<cdot>(x && xs) =
   542                            If f\<cdot>x then x && sfilter\<cdot>f\<cdot>xs else sfilter\<cdot>f\<cdot>xs fi"
   543 by (subst sfilter_unfold, force)
   544 
   545 
   546 (* ----------------------------------------------------------------------- *)
   547    section "i_rt"
   548 (* ----------------------------------------------------------------------- *)
   549 
   550 lemma i_rt_UU [simp]: "i_rt n UU = UU"
   551   by (induct n) (simp_all add: i_rt_def)
   552 
   553 lemma i_rt_0 [simp]: "i_rt 0 s = s"
   554 by (simp add: i_rt_def)
   555 
   556 lemma i_rt_Suc [simp]: "a ~= UU ==> i_rt (Suc n) (a&&s) = i_rt n s"
   557 by (simp add: i_rt_def iterate_Suc2 del: iterate_Suc)
   558 
   559 lemma i_rt_Suc_forw: "i_rt (Suc n) s = i_rt n (rt$s)"
   560 by (simp only: i_rt_def iterate_Suc2)
   561 
   562 lemma i_rt_Suc_back:"i_rt (Suc n) s = rt$(i_rt n s)"
   563 by (simp only: i_rt_def,auto)
   564 
   565 lemma i_rt_mono: "x << s ==> i_rt n x  << i_rt n s"
   566 by (simp add: i_rt_def monofun_rt_mult)
   567 
   568 lemma i_rt_ij_lemma: "Fin (i + j) <= #x ==> Fin j <= #(i_rt i x)"
   569 by (simp add: i_rt_def slen_rt_mult)
   570 
   571 lemma slen_i_rt_mono: "#s2 <= #s1 ==> #(i_rt n s2) <= #(i_rt n s1)"
   572 apply (induct_tac n,auto)
   573 apply (simp add: i_rt_Suc_back)
   574 by (drule slen_rt_mono,simp)
   575 
   576 lemma i_rt_take_lemma1 [rule_format]: "ALL s. i_rt n (stream_take n$s) = UU"
   577 apply (induct_tac n)
   578  apply (simp add: i_rt_Suc_back,auto)
   579 apply (case_tac "s=UU",auto)
   580 by (drule stream_exhaust_eq [THEN iffD1],auto)
   581 
   582 lemma i_rt_slen: "(i_rt n s = UU) = (stream_take n$s = s)"
   583 apply auto
   584  apply (insert i_rt_ij_lemma [of n "Suc 0" s])
   585  apply (subgoal_tac "#(i_rt n s)=0")
   586   apply (case_tac "stream_take n$s = s",simp+)
   587   apply (insert slen_take_eq [rule_format,of n s],simp)
   588   apply (cases "#s") apply (simp_all add: zero_inat_def)
   589   apply (simp add: slen_take_eq)
   590   apply (cases "#s")
   591   using i_rt_take_lemma1 [of n s]
   592   apply (simp_all add: zero_inat_def)
   593   done
   594 
   595 lemma i_rt_lemma_slen: "#s=Fin n ==> i_rt n s = UU"
   596 by (simp add: i_rt_slen slen_take_lemma1)
   597 
   598 lemma stream_finite_i_rt [simp]: "stream_finite (i_rt n s) = stream_finite s"
   599 apply (induct_tac n, auto)
   600  apply (rule stream.casedist [of "s"], auto simp del: i_rt_Suc)
   601 by (simp add: i_rt_Suc_back stream_finite_rt_eq)+
   602 
   603 lemma take_i_rt_len_lemma: "ALL sl x j t. Fin sl = #x & n <= sl &
   604                             #(stream_take n$x) = Fin t & #(i_rt n x)= Fin j
   605                                               --> Fin (j + t) = #x"
   606 apply (induct n, auto)
   607  apply (simp add: zero_inat_def)
   608 apply (case_tac "x=UU",auto)
   609  apply (simp add: zero_inat_def)
   610 apply (drule stream_exhaust_eq [THEN iffD1],clarsimp)
   611 apply (subgoal_tac "EX k. Fin k = #y",clarify)
   612  apply (erule_tac x="k" in allE)
   613  apply (erule_tac x="y" in allE,auto)
   614  apply (erule_tac x="THE p. Suc p = t" in allE,auto)
   615    apply (simp add: iSuc_def split: inat.splits)
   616   apply (simp add: iSuc_def split: inat.splits)
   617   apply (simp only: the_equality)
   618  apply (simp add: iSuc_def split: inat.splits)
   619  apply force
   620 apply (simp add: iSuc_def split: inat.splits)
   621 done
   622 
   623 lemma take_i_rt_len:
   624 "[| Fin sl = #x; n <= sl; #(stream_take n$x) = Fin t; #(i_rt n x) = Fin j |] ==>
   625     Fin (j + t) = #x"
   626 by (blast intro: take_i_rt_len_lemma [rule_format])
   627 
   628 
   629 (* ----------------------------------------------------------------------- *)
   630    section "i_th"
   631 (* ----------------------------------------------------------------------- *)
   632 
   633 lemma i_th_i_rt_step:
   634 "[| i_th n s1 << i_th n s2; i_rt (Suc n) s1 << i_rt (Suc n) s2 |] ==>
   635    i_rt n s1 << i_rt n s2"
   636 apply (simp add: i_th_def i_rt_Suc_back)
   637 apply (rule stream.casedist [of "i_rt n s1"],simp)
   638 apply (rule stream.casedist [of "i_rt n s2"],auto)
   639 done
   640 
   641 lemma i_th_stream_take_Suc [rule_format]:
   642  "ALL s. i_th n (stream_take (Suc n)$s) = i_th n s"
   643 apply (induct_tac n,auto)
   644  apply (simp add: i_th_def)
   645  apply (case_tac "s=UU",auto)
   646  apply (drule stream_exhaust_eq [THEN iffD1],auto)
   647 apply (case_tac "s=UU",simp add: i_th_def)
   648 apply (drule stream_exhaust_eq [THEN iffD1],auto)
   649 by (simp add: i_th_def i_rt_Suc_forw)
   650 
   651 lemma i_th_last: "i_th n s && UU = i_rt n (stream_take (Suc n)$s)"
   652 apply (insert surjectiv_scons [of "i_rt n (stream_take (Suc n)$s)"])
   653 apply (rule i_th_stream_take_Suc [THEN subst])
   654 apply (simp add: i_th_def  i_rt_Suc_back [symmetric])
   655 by (simp add: i_rt_take_lemma1)
   656 
   657 lemma i_th_last_eq:
   658 "i_th n s1 = i_th n s2 ==> i_rt n (stream_take (Suc n)$s1) = i_rt n (stream_take (Suc n)$s2)"
   659 apply (insert i_th_last [of n s1])
   660 apply (insert i_th_last [of n s2])
   661 by auto
   662 
   663 lemma i_th_prefix_lemma:
   664 "[| k <= n; stream_take (Suc n)$s1 << stream_take (Suc n)$s2 |] ==>
   665     i_th k s1 << i_th k s2"
   666 apply (insert i_th_stream_take_Suc [of k s1, THEN sym])
   667 apply (insert i_th_stream_take_Suc [of k s2, THEN sym],auto)
   668 apply (simp add: i_th_def)
   669 apply (rule monofun_cfun, auto)
   670 apply (rule i_rt_mono)
   671 by (blast intro: stream_take_lemma10)
   672 
   673 lemma take_i_rt_prefix_lemma1:
   674   "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==>
   675    i_rt (Suc n) s1 << i_rt (Suc n) s2 ==>
   676    i_rt n s1 << i_rt n s2 & stream_take n$s1 << stream_take n$s2"
   677 apply auto
   678  apply (insert i_th_prefix_lemma [of n n s1 s2])
   679  apply (rule i_th_i_rt_step,auto)
   680 by (drule mono_stream_take_pred,simp)
   681 
   682 lemma take_i_rt_prefix_lemma:
   683 "[| stream_take n$s1 << stream_take n$s2; i_rt n s1 << i_rt n s2 |] ==> s1 << s2"
   684 apply (case_tac "n=0",simp)
   685 apply (auto)
   686 apply (subgoal_tac "stream_take 0$s1 << stream_take 0$s2 &
   687                     i_rt 0 s1 << i_rt 0 s2")
   688  defer 1
   689  apply (rule zero_induct,blast)
   690  apply (blast dest: take_i_rt_prefix_lemma1)
   691 by simp
   692 
   693 lemma streams_prefix_lemma: "(s1 << s2) =
   694   (stream_take n$s1 << stream_take n$s2 & i_rt n s1 << i_rt n s2)"
   695 apply auto
   696   apply (simp add: monofun_cfun_arg)
   697  apply (simp add: i_rt_mono)
   698 by (erule take_i_rt_prefix_lemma,simp)
   699 
   700 lemma streams_prefix_lemma1:
   701  "[| stream_take n$s1 = stream_take n$s2; i_rt n s1 = i_rt n s2 |] ==> s1 = s2"
   702 apply (simp add: po_eq_conv,auto)
   703  apply (insert streams_prefix_lemma)
   704  by blast+
   705 
   706 
   707 (* ----------------------------------------------------------------------- *)
   708    section "sconc"
   709 (* ----------------------------------------------------------------------- *)
   710 
   711 lemma UU_sconc [simp]: " UU ooo s = s "
   712 by (simp add: sconc_def zero_inat_def)
   713 
   714 lemma scons_neq_UU: "a~=UU ==> a && s ~=UU"
   715 by auto
   716 
   717 lemma singleton_sconc [rule_format, simp]: "x~=UU --> (x && UU) ooo y = x && y"
   718 apply (simp add: sconc_def zero_inat_def iSuc_def split: inat.splits, auto)
   719 apply (rule someI2_ex,auto)
   720  apply (rule_tac x="x && y" in exI,auto)
   721 apply (simp add: i_rt_Suc_forw)
   722 apply (case_tac "xa=UU",simp)
   723 by (drule stream_exhaust_eq [THEN iffD1],auto)
   724 
   725 lemma ex_sconc [rule_format]:
   726   "ALL k y. #x = Fin k --> (EX w. stream_take k$w = x & i_rt k w = y)"
   727 apply (case_tac "#x")
   728  apply (rule stream_finite_ind [of x],auto)
   729   apply (simp add: stream.finite_def)
   730   apply (drule slen_take_lemma1,blast)
   731  apply (simp_all add: zero_inat_def iSuc_def split: inat.splits)
   732 apply (erule_tac x="y" in allE,auto)
   733 by (rule_tac x="a && w" in exI,auto)
   734 
   735 lemma rt_sconc1: "Fin n = #x ==> i_rt n (x ooo y) = y"
   736 apply (simp add: sconc_def split: inat.splits, arith?,auto)
   737 apply (rule someI2_ex,auto)
   738 by (drule ex_sconc,simp)
   739 
   740 lemma sconc_inj2: "\<lbrakk>Fin n = #x; x ooo y = x ooo z\<rbrakk> \<Longrightarrow> y = z"
   741 apply (frule_tac y=y in rt_sconc1)
   742 by (auto elim: rt_sconc1)
   743 
   744 lemma sconc_UU [simp]:"s ooo UU = s"
   745 apply (case_tac "#s")
   746  apply (simp add: sconc_def)
   747  apply (rule someI2_ex)
   748   apply (rule_tac x="s" in exI)
   749   apply auto
   750    apply (drule slen_take_lemma1,auto)
   751   apply (simp add: i_rt_lemma_slen)
   752  apply (drule slen_take_lemma1,auto)
   753  apply (simp add: i_rt_slen)
   754 by (simp add: sconc_def)
   755 
   756 lemma stream_take_sconc [simp]: "Fin n = #x ==> stream_take n$(x ooo y) = x"
   757 apply (simp add: sconc_def)
   758 apply (cases "#x")
   759 apply auto
   760 apply (rule someI2_ex, auto)
   761 by (drule ex_sconc,simp)
   762 
   763 lemma scons_sconc [rule_format,simp]: "a~=UU --> (a && x) ooo y = a && x ooo y"
   764 apply (cases "#x",auto)
   765  apply (simp add: sconc_def iSuc_Fin)
   766  apply (rule someI2_ex)
   767   apply (drule ex_sconc, simp)
   768  apply (rule someI2_ex, auto)
   769   apply (simp add: i_rt_Suc_forw)
   770   apply (rule_tac x="a && x" in exI, auto)
   771  apply (case_tac "xa=UU",auto)
   772  apply (drule stream_exhaust_eq [THEN iffD1],auto)
   773  apply (drule streams_prefix_lemma1,simp+)
   774 by (simp add: sconc_def)
   775 
   776 lemma ft_sconc: "x ~= UU ==> ft$(x ooo y) = ft$x"
   777 by (rule stream.casedist [of x],auto)
   778 
   779 lemma sconc_assoc: "(x ooo y) ooo z = x ooo y ooo z"
   780 apply (case_tac "#x")
   781  apply (rule stream_finite_ind [of x],auto simp del: scons_sconc)
   782   apply (simp add: stream.finite_def del: scons_sconc)
   783   apply (drule slen_take_lemma1,auto simp del: scons_sconc)
   784  apply (case_tac "a = UU", auto)
   785 by (simp add: sconc_def)
   786 
   787 
   788 (* ----------------------------------------------------------------------- *)
   789 
   790 lemma cont_sconc_lemma1: "stream_finite x \<Longrightarrow> cont (\<lambda>y. x ooo y)"
   791 by (erule stream_finite_ind, simp_all)
   792 
   793 lemma cont_sconc_lemma2: "\<not> stream_finite x \<Longrightarrow> cont (\<lambda>y. x ooo y)"
   794 by (simp add: sconc_def slen_def)
   795 
   796 lemma cont_sconc: "cont (\<lambda>y. x ooo y)"
   797 apply (cases "stream_finite x")
   798 apply (erule cont_sconc_lemma1)
   799 apply (erule cont_sconc_lemma2)
   800 done
   801 
   802 lemma sconc_mono: "y << y' ==> x ooo y << x ooo y'"
   803 by (rule cont_sconc [THEN cont2mono, THEN monofunE])
   804 
   805 lemma sconc_mono1 [simp]: "x << x ooo y"
   806 by (rule sconc_mono [of UU, simplified])
   807 
   808 (* ----------------------------------------------------------------------- *)
   809 
   810 lemma empty_sconc [simp]: "(x ooo y = UU) = (x = UU & y = UU)"
   811 apply (case_tac "#x",auto)
   812    apply (insert sconc_mono1 [of x y])
   813    by auto
   814 
   815 (* ----------------------------------------------------------------------- *)
   816 
   817 lemma rt_sconc [rule_format, simp]: "s~=UU --> rt$(s ooo x) = rt$s ooo x"
   818 by (rule stream.casedist,auto)
   819 
   820 lemma i_th_sconc_lemma [rule_format]:
   821   "ALL x y. Fin n < #x --> i_th n (x ooo y) = i_th n x"
   822 apply (induct_tac n, auto)
   823 apply (simp add: Fin_0 i_th_def)
   824 apply (simp add: slen_empty_eq ft_sconc)
   825 apply (simp add: i_th_def)
   826 apply (case_tac "x=UU",auto)
   827 apply (drule stream_exhaust_eq [THEN iffD1], auto)
   828 apply (erule_tac x="ya" in allE)
   829 apply (case_tac "#ya") by simp_all
   830 
   831 
   832 
   833 (* ----------------------------------------------------------------------- *)
   834 
   835 lemma sconc_lemma [rule_format, simp]: "ALL s. stream_take n$s ooo i_rt n s = s"
   836 apply (induct_tac n,auto)
   837 apply (case_tac "s=UU",auto)
   838 by (drule stream_exhaust_eq [THEN iffD1],auto)
   839 
   840 (* ----------------------------------------------------------------------- *)
   841    subsection "pointwise equality"
   842 (* ----------------------------------------------------------------------- *)
   843 
   844 lemma ex_last_stream_take_scons: "stream_take (Suc n)$s =
   845                      stream_take n$s ooo i_rt n (stream_take (Suc n)$s)"
   846 by (insert sconc_lemma [of n "stream_take (Suc n)$s"],simp)
   847 
   848 lemma i_th_stream_take_eq:
   849 "!!n. ALL n. i_th n s1 = i_th n s2 ==> stream_take n$s1 = stream_take n$s2"
   850 apply (induct_tac n,auto)
   851 apply (subgoal_tac "stream_take (Suc na)$s1 =
   852                     stream_take na$s1 ooo i_rt na (stream_take (Suc na)$s1)")
   853  apply (subgoal_tac "i_rt na (stream_take (Suc na)$s1) =
   854                     i_rt na (stream_take (Suc na)$s2)")
   855   apply (subgoal_tac "stream_take (Suc na)$s2 =
   856                     stream_take na$s2 ooo i_rt na (stream_take (Suc na)$s2)")
   857    apply (insert ex_last_stream_take_scons,simp)
   858   apply blast
   859  apply (erule_tac x="na" in allE)
   860  apply (insert i_th_last_eq [of _ s1 s2])
   861 by blast+
   862 
   863 lemma pointwise_eq_lemma[rule_format]: "ALL n. i_th n s1 = i_th n s2 ==> s1 = s2"
   864 by (insert i_th_stream_take_eq [THEN stream.take_lemmas],blast)
   865 
   866 (* ----------------------------------------------------------------------- *)
   867    subsection "finiteness"
   868 (* ----------------------------------------------------------------------- *)
   869 
   870 lemma slen_sconc_finite1:
   871   "[| #(x ooo y) = Infty; Fin n = #x |] ==> #y = Infty"
   872 apply (case_tac "#y ~= Infty",auto)
   873 apply (drule_tac y=y in rt_sconc1)
   874 apply (insert stream_finite_i_rt [of n "x ooo y"])
   875 by (simp add: slen_infinite)
   876 
   877 lemma slen_sconc_infinite1: "#x=Infty ==> #(x ooo y) = Infty"
   878 by (simp add: sconc_def)
   879 
   880 lemma slen_sconc_infinite2: "#y=Infty ==> #(x ooo y) = Infty"
   881 apply (case_tac "#x")
   882  apply (simp add: sconc_def)
   883  apply (rule someI2_ex)
   884   apply (drule ex_sconc,auto)
   885  apply (erule contrapos_pp)
   886  apply (insert stream_finite_i_rt)
   887  apply (fastsimp simp add: slen_infinite,auto)
   888 by (simp add: sconc_def)
   889 
   890 lemma sconc_finite: "(#x~=Infty & #y~=Infty) = (#(x ooo y)~=Infty)"
   891 apply auto
   892   apply (metis not_Infty_eq slen_sconc_finite1)
   893  apply (metis not_Infty_eq slen_sconc_infinite1)
   894 apply (metis not_Infty_eq slen_sconc_infinite2)
   895 done
   896 
   897 (* ----------------------------------------------------------------------- *)
   898 
   899 lemma slen_sconc_mono3: "[| Fin n = #x; Fin k = #(x ooo y) |] ==> n <= k"
   900 apply (insert slen_mono [of "x" "x ooo y"])
   901 apply (cases "#x") apply simp_all
   902 apply (cases "#(x ooo y)") apply simp_all
   903 done
   904 
   905 (* ----------------------------------------------------------------------- *)
   906    subsection "finite slen"
   907 (* ----------------------------------------------------------------------- *)
   908 
   909 lemma slen_sconc: "[| Fin n = #x; Fin m = #y |] ==> #(x ooo y) = Fin (n + m)"
   910 apply (case_tac "#(x ooo y)")
   911  apply (frule_tac y=y in rt_sconc1)
   912  apply (insert take_i_rt_len [of "THE j. Fin j = #(x ooo y)" "x ooo y" n n m],simp)
   913  apply (insert slen_sconc_mono3 [of n x _ y],simp)
   914 by (insert sconc_finite [of x y],auto)
   915 
   916 (* ----------------------------------------------------------------------- *)
   917    subsection "flat prefix"
   918 (* ----------------------------------------------------------------------- *)
   919 
   920 lemma sconc_prefix: "(s1::'a::flat stream) << s2 ==> EX t. s1 ooo t = s2"
   921 apply (case_tac "#s1")
   922  apply (subgoal_tac "stream_take nat$s1 = stream_take nat$s2")
   923   apply (rule_tac x="i_rt nat s2" in exI)
   924   apply (simp add: sconc_def)
   925   apply (rule someI2_ex)
   926    apply (drule ex_sconc)
   927    apply (simp,clarsimp,drule streams_prefix_lemma1)
   928    apply (simp+,rule slen_take_lemma3 [of _ s1 s2])
   929   apply (simp+,rule_tac x="UU" in exI)
   930 apply (insert slen_take_lemma3 [of _ s1 s2])
   931 by (rule stream.take_lemmas,simp)
   932 
   933 (* ----------------------------------------------------------------------- *)
   934    subsection "continuity"
   935 (* ----------------------------------------------------------------------- *)
   936 
   937 lemma chain_sconc: "chain S ==> chain (%i. (x ooo S i))"
   938 by (simp add: chain_def,auto simp add: sconc_mono)
   939 
   940 lemma chain_scons: "chain S ==> chain (%i. a && S i)"
   941 apply (simp add: chain_def,auto)
   942 by (rule monofun_cfun_arg,simp)
   943 
   944 lemma contlub_scons: "contlub (%x. a && x)"
   945 by (simp add: contlub_Rep_CFun2)
   946 
   947 lemma contlub_scons_lemma: "chain S ==> (LUB i. a && S i) = a && (LUB i. S i)"
   948 by (rule contlubE [OF contlub_Rep_CFun2, symmetric])
   949 
   950 lemma finite_lub_sconc: "chain Y ==> (stream_finite x) ==>
   951                         (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))"
   952 apply (rule stream_finite_ind [of x])
   953  apply (auto)
   954 apply (subgoal_tac "(LUB i. a && (s ooo Y i)) = a && (LUB i. s ooo Y i)")
   955  by (force,blast dest: contlub_scons_lemma chain_sconc)
   956 
   957 lemma contlub_sconc_lemma:
   958   "chain Y ==> (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))"
   959 apply (case_tac "#x=Infty")
   960  apply (simp add: sconc_def)
   961 apply (drule finite_lub_sconc,auto simp add: slen_infinite)
   962 done
   963 
   964 lemma contlub_sconc: "contlub (%y. x ooo y)"
   965 by (rule cont_sconc [THEN cont2contlub])
   966 
   967 lemma monofun_sconc: "monofun (%y. x ooo y)"
   968 by (simp add: monofun_def sconc_mono)
   969 
   970 
   971 (* ----------------------------------------------------------------------- *)
   972    section "constr_sconc"
   973 (* ----------------------------------------------------------------------- *)
   974 
   975 lemma constr_sconc_UUs [simp]: "constr_sconc UU s = s"
   976 by (simp add: constr_sconc_def zero_inat_def)
   977 
   978 lemma "x ooo y = constr_sconc x y"
   979 apply (case_tac "#x")
   980  apply (rule stream_finite_ind [of x],auto simp del: scons_sconc)
   981   defer 1
   982   apply (simp add: constr_sconc_def del: scons_sconc)
   983   apply (case_tac "#s")
   984    apply (simp add: iSuc_Fin)
   985    apply (case_tac "a=UU",auto simp del: scons_sconc)
   986    apply (simp)
   987   apply (simp add: sconc_def)
   988  apply (simp add: constr_sconc_def)
   989 apply (simp add: stream.finite_def)
   990 by (drule slen_take_lemma1,auto)
   991 
   992 end