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