Theory Stream

section ‹Stream Iterators›

theory Stream
imports LazyList
begin

subsection ‹Type definitions for streams›

text ‹Note that everything is strict in the state type.›

domain ('a,'s) Step = Done | Skip 's | Yield (lazy 'a) 's

type_synonym ('a, 's) Stepper = "'s  ('a, 's) Step"

domain ('a,'s) Stream = Stream (lazy "('a, 's) Stepper") 's


subsection ‹Converting from streams to lists›

fixrec
  unfold :: "('a, 's) Stepper -> ('s -> 'a LList)"
where
  "unfoldh = "
| "s   
    unfoldhs =
      (case hs of
        Done  LNil
      | Skips'  unfoldhs'
      | Yieldxs'  LConsx(unfoldhs'))"

fixrec
  unfoldF :: "('a, 's) Stepper  ('s  'a LList)  ('s  'a LList)"
where
  "unfoldFhu = "
| "s   
    unfoldFhus =
      (case hs of
        Done  LNil
      | Skips'  us'
      | Yieldxs'  LConsx(us'))"

lemma unfold_eq_fix: "unfoldh = fix(unfoldFh)"
proof (rule below_antisym)
  show "unfoldh  fix(unfoldFh)"
    apply (rule unfold.induct, simp, simp)
    apply (subst fix_eq)
    apply (rule cfun_belowI, rename_tac s)
    apply (case_tac "s = ", simp, simp)
    apply (intro monofun_cfun monofun_LAM below_refl, simp_all)
    done
  show "fix(unfoldFh)  unfoldh"
    apply (rule fix_ind, simp, simp)
    apply (subst unfold.unfold)
    apply (rule cfun_belowI, rename_tac s)
    apply (case_tac "s = ", simp, simp)
    apply (intro monofun_cfun monofun_LAM below_refl, simp_all)
    done
qed

lemma unfold_ind:
    fixes P :: "('s  'a LList)  bool"
    assumes "adm P" and "P " and "u. P u  P (unfoldFhu)"
    shows "P (unfoldh)"
unfolding unfold_eq_fix by (rule fix_ind [of P, OF assms])

fixrec
  unfold2 :: "('s  'a LList)  ('a, 's) Step  'a LList"
where
  "unfold2uDone = LNil"
| "s    unfold2u(Skips) = us"
| "s    unfold2u(Yieldxs) = LConsx(us)"

lemma unfold2_strict [simp]: "unfold2u = "
by fixrec_simp

lemma unfold: "s    unfoldhs = unfold2(unfoldh)(hs)"
by (case_tac "hs", simp_all)

lemma unfoldF: "s    unfoldFhus = unfold2u(hs)"
by (case_tac "hs", simp_all)

declare unfold.simps(2) [simp del]
declare unfoldF.simps(2) [simp del]
declare unfoldF [simp]

fixrec
  unstream :: "('a, 's) Stream  'a LList"
where
  "s    unstream(Streamhs) = unfoldhs"

lemma unstream_strict [simp]: "unstream = "
by fixrec_simp


subsection ‹Converting from lists to streams›

fixrec
  streamStep :: "('a LList)  ('a, ('a LList)) Step"
where
  "streamStep(upLNil) = Done"
| "streamStep(up(LConsxxs)) = Yieldx(upxs)"

lemma streamStep_strict [simp]: "streamStep(up) = "
by fixrec_simp

fixrec
  stream :: "'a LList  ('a, ('a LList)) Stream"
where
  "streamxs = StreamstreamStep(upxs)"

lemma stream_defined [simp]: "streamxs  "
  by simp

lemma unstream_stream [simp]:
  fixes xs :: "'a LList"
  shows "unstream(streamxs) = xs"
by (induct xs, simp_all add: unfold)

declare stream.simps [simp del]


subsection ‹Bisimilarity relation on streams›

definition
  bisimilar :: "('a, 's) Stream  ('a, 't) Stream  bool" (infix "" 50)
where
  "a  b  unstreama = unstreamb  a    b  "

lemma unstream_cong:
  "a  b  unstreama = unstreamb"
    unfolding bisimilar_def by simp

lemma stream_cong:
  "xs = ys  streamxs  streamys"
    unfolding bisimilar_def by simp

lemma stream_unstream_cong:
  "a  b  stream(unstreama)  b"
    unfolding bisimilar_def by simp

end