Theory StreamFusion

section ‹Stream Fusion›

theory StreamFusion
imports Stream
begin

subsection ‹Type constructors for state types›

domain Switch = S1 | S2

domain 'a Maybe = Nothing | Just 'a

hide_const (open) Left Right

domain ('a, 'b) Either = Left 'a | Right 'b

domain ('a, 'b) Both  (infixl ":!:" 25) = Both 'a 'b (infixl ":!:" 75)

domain 'a L = L (lazy 'a)


subsection ‹Map function›

fixrec
  mapStep :: "('a  'b)  ('s  ('a, 's) Step)  's  ('b, 's) Step"
where
  "mapStepfh = "
| "s    mapStepfhs = (case hs of
    Done  Done
  | Skips'  Skips'
  | Yieldxs'  Yield(fx)s')"

fixrec
  mapS :: "('a  'b)  ('a, 's) Stream  ('b, 's) Stream"
where
  "s    mapSf(Streamhs) = Stream(mapStepfh)s"

lemma unfold_mapStep:
  fixes f :: "'a  'b" and h :: "'s  ('a, 's) Step"
  assumes "s  "
  shows "unfold(mapStepfh)s = mapLf(unfoldhs)"
proof (rule below_antisym)
  show "unfold(mapStepfh)s  mapLf(unfoldhs)"
  using s  
  apply (induct arbitrary: s rule: unfold_ind [where h="mapStepfh"])
  apply (simp, simp)
  apply (case_tac "hs", simp_all add: unfold)
  done
next
  show "mapLf(unfoldhs)  unfold(mapStepfh)s"
  using s  
  apply (induct arbitrary: s rule: unfold_ind [where h="h"])
  apply (simp, simp)
  apply (case_tac "hs", simp_all add: unfold)
  done
qed

lemma unstream_mapS:
  fixes f :: "'a  'b" and a :: "('a, 's) Stream"
  shows "a    unstream(mapSfa) = mapLf(unstreama)"
by (induct a, simp, simp add: unfold_mapStep)

lemma mapS_defined: "a    mapSfa  "
by (induct a, simp_all)

lemma mapS_cong:
  fixes f :: "'a  'b"
  fixes a :: "('a, 's) Stream"
  fixes b :: "('a, 't) Stream"
  shows "f = g  a  b  mapSfa  mapSgb"
unfolding bisimilar_def
by (simp add: unstream_mapS mapS_defined)

lemma mapL_eq: "mapLfxs = unstream(mapSf(streamxs))"
by (simp add: unstream_mapS)


subsection ‹Filter function›

fixrec
  filterStep :: "('a  tr)  ('s  ('a, 's) Step)  's  ('a, 's) Step"
where
  "filterStepph = "
| "s    filterStepphs = (case hs of
    Done  Done
  | Skips'  Skips'
  | Yieldxs'  (If px then Yieldxs' else Skips'))"

fixrec
  filterS :: "('a  tr)  ('a, 's) Stream  ('a, 's) Stream"
where
  "s    filterSp(Streamhs) = Stream(filterStepph)s"

lemma unfold_filterStep:
  fixes p :: "'a  tr" and h :: "'s  ('a, 's) Step"
  assumes "s  "
  shows "unfold(filterStepph)s = filterLp(unfoldhs)"
proof (rule below_antisym)
  show "unfold(filterStepph)s  filterLp(unfoldhs)"
  using s  
  apply (induct arbitrary: s rule: unfold_ind [where h="filterStepph"])
  apply (simp, simp)
  apply (case_tac "hs", simp_all add: unfold)
  apply (case_tac "pa" rule: trE, simp_all)
  done
next
  show "filterLp(unfoldhs)  unfold(filterStepph)s"
  using s  
  apply (induct arbitrary: s rule: unfold_ind [where h="h"])
  apply (simp, simp)
  apply (case_tac "hs", simp_all add: unfold)
  apply (case_tac "pa" rule: trE, simp_all add: unfold)
  done
qed

lemma unstream_filterS:
  "a    unstream(filterSpa) = filterLp(unstreama)"
by (induct a, simp, simp add: unfold_filterStep)

lemma filterS_defined: "a    filterSpa  "
by (induct a, simp_all)

lemma filterS_cong:
  fixes p :: "'a  tr"
  fixes a :: "('a, 's) Stream"
  fixes b :: "('a, 't) Stream"
  shows "p = q  a  b  filterSpa  filterSqb"
unfolding bisimilar_def
by (simp add: unstream_filterS filterS_defined)

lemma filterL_eq: "filterLpxs = unstream(filterSp(streamxs))"
by (simp add: unstream_filterS)


subsection ‹Foldr function›

fixrec
  foldrS :: "('a  'b  'b)  'b  ('a, 's) Stream  'b"
where
  foldrS_Stream:
  "s    foldrSfz(Streamhs) =
    (case hs of Done  z
               | Skips'  foldrSfz(Streamhs')
               | Yieldxs'  fx(foldrSfz(Streamhs')))"

lemma unfold_foldrS:
  assumes "s  " shows "foldrSfz(Streamhs) = foldrLfz(unfoldhs)"
proof (rule below_antisym)
  show "foldrSfz(Streamhs)  foldrLfz(unfoldhs)"
  using s  
  apply (induct arbitrary: s rule: foldrS.induct)
  apply (simp, simp, simp)
  apply (case_tac "hs", simp_all add: monofun_cfun unfold)
  done
next
  show "foldrLfz(unfoldhs)  foldrSfz(Streamhs)"
  using s  
  apply (induct arbitrary: s rule: unfold_ind)
  apply (simp, simp)
  apply (case_tac "hs", simp_all add: monofun_cfun unfold)
  done
qed

lemma unstream_foldrS:
  "a    foldrSfza = foldrLfz(unstreama)"
by (induct a, simp, simp del: foldrS_Stream add: unfold_foldrS)

lemma foldrS_cong:
  fixes a :: "('a, 's) Stream"
  fixes b :: "('a, 't) Stream"
  shows "f = g  z = w  a  b  foldrSfza = foldrSgwb"
by (simp add: bisimilar_def unstream_foldrS)

lemma foldrL_eq:
  "foldrLfzxs = foldrSfz(streamxs)"
by (simp add: unstream_foldrS)


subsection ‹EnumFromTo function›

type_synonym int' = "int"

fixrec
  enumFromToStep :: "int'  (int')  (int', (int')) Step"
where
  "enumFromToStep(upy)(up(upx)) =
    (if x  y then Yield(upx)(up(up(x+1))) else Done)"

lemma enumFromToStep_strict [simp]:
  "enumFromToStepx'' = "
  "enumFromToStep(upy) = "
  "enumFromToStep(upy)(up) = "
by fixrec_simp+

lemma enumFromToStep_simps' [simp]:
  "x  y  enumFromToStep(upy)(up(upx)) =
    Yield(upx)(up(up(x+1)))"
  "¬ x  y  enumFromToStep(upy)(up(upx)) = Done"
  by simp_all

declare enumFromToStep.simps [simp del]

fixrec
  enumFromToS :: "int'  int'  (int', (int')) Stream"
where
  "enumFromToSxy = Stream(enumFromToStepy)(upx)"

declare enumFromToS.simps [simp del]

lemma unfold_enumFromToStep:
  "unfold(enumFromToStep(upy))(upn) = enumFromToLn(upy)"
proof (rule below_antisym)
  show "unfold(enumFromToStep(upy))(upn)  enumFromToLn(upy)"
    apply (induct arbitrary: n rule: unfold_ind [where h="enumFromToStep(upy)"])
    apply (simp, simp)
    apply (case_tac n, simp, simp)
    apply (case_tac "x  y", simp_all)
    done
next
  show "enumFromToLn(upy)  unfold(enumFromToStep(upy))(upn)"
    apply (induct arbitrary: n rule: enumFromToL.induct)
    apply (simp, simp)
    apply (rename_tac e n)
    apply (case_tac n, simp)
    apply (case_tac "x  y", simp_all add: unfold)
    done
qed

lemma unstream_enumFromToS:
  "unstream(enumFromToSxy) = enumFromToLxy"
apply (simp add: enumFromToS.simps)
apply (induct y, simp add: unfold)
apply (induct x, simp add: unfold)
apply (simp add: unfold_enumFromToStep)
done

lemma enumFromToS_defined: "enumFromToSxy  "
  by (simp add: enumFromToS.simps)

lemma enumFromToS_cong:
  "x = x'  y = y'  enumFromToSxy  enumFromToSx'y'"
unfolding bisimilar_def by (simp add: enumFromToS_defined)

lemma enumFromToL_eq: "enumFromToLxy = unstream(enumFromToSxy)"
by (simp add: unstream_enumFromToS)


subsection ‹Append function›

fixrec
  appendStep ::
    "('s  ('a, 's) Step) 
     ('t  ('a, 't) Step) 
     't  ('s, 't) Either  ('a, ('s, 't) Either) Step"
where
  "sa    appendStephahbsb0(Leftsa) =
    (case hasa of
      Done  Skip(Rightsb0)
    | Skipsa'  Skip(Leftsa')
    | Yieldxsa'  Yieldx(Leftsa'))"
| "sb    appendStephahbsb0(Rightsb) =
    (case hbsb of
      Done  Done
    | Skipsb'  Skip(Rightsb')
    | Yieldxsb'  Yieldx(Rightsb'))"

lemma appendStep_strict [simp]: "appendStephahbsb0 = "
by fixrec_simp

fixrec
  appendS ::
    "('a, 's) Stream  ('a, 't) Stream  ('a, ('s, 't) Either) Stream"
where
  "sa0    sb0   
    appendS(Streamhasa0)(Streamhbsb0) =
      Stream(appendStephahbsb0)(Leftsa0)"

lemma unfold_appendStep:
  fixes ha :: "'s  ('a, 's) Step"
  fixes hb :: "'t  ('a, 't) Step"
  assumes sb0 [simp]: "sb0  "
  shows
  "(sa. sa    unfold(appendStephahbsb0)(Leftsa) =
         appendL(unfoldhasa)(unfoldhbsb0)) 
   (sb. sb    unfold(appendStephahbsb0)(Rightsb) =
         unfoldhbsb)"
proof -
  note unfold [simp]
  let ?h = "appendStephahbsb0"

  have 1:
  "(sa. sa   
         unfold?h(Leftsa) 
         appendL(unfoldhasa)(unfoldhbsb0))
   
   (sb. sb    unfold?h(Rightsb)  unfoldhbsb)"
    apply (rule unfold_ind [where h="?h"])
      apply simp
     apply simp
    apply (intro conjI allI impI)
     apply (case_tac "hasa", simp, simp, simp, simp)
    apply (case_tac "hbsb", simp, simp, simp, simp)
    done

  let ?P = "λua ub. sa. sa   
        appendL(uasa)(ubsb0)  unfold?h(Leftsa)"

  let ?Q = "λub. sb. sb    ubsb  unfold?h(Rightsb)"

  have P_base: "ub. ?P  ub"
    by simp

  have Q_base: "?Q "
    by simp

  have P_step: "ua ub. ?P ua ub  ?Q ub  ?P (unfoldFhaua) ub"
    apply (intro allI impI)
    apply (case_tac "hasa", simp, simp, simp, simp)
    done

  have Q_step: "ua ub. ?Q ub  ?Q (unfoldFhbub)"
    apply (intro allI impI)
    apply (case_tac "hbsb", simp, simp, simp, simp)
    done

  have Q: "?Q (unfoldhb)"
    apply (rule unfold_ind [where h="hb"], simp)
     apply (rule Q_base)
    apply (erule Q_step)
    done

  have P: "?P (unfoldha) (unfoldhb)"
    apply (rule unfold_ind [where h="ha"], simp)
     apply (rule P_base)
    apply (erule P_step)
    apply (rule Q)
    done

  have 2: "?P (unfoldha) (unfoldhb)  ?Q (unfoldhb)"
    using P Q by (rule conjI)

  from 1 2 show ?thesis
    by (simp add: po_eq_conv [where 'a="'a LList"])
qed

lemma appendS_defined: "xs    ys    appendSxsys  "
by (cases xs, simp, cases ys, simp, simp)

lemma unstream_appendS:
  "a    b   
    unstream(appendSab) = appendL(unstreama)(unstreamb)"
apply (cases a, simp, cases b, simp)
apply (simp add: unfold_appendStep)
done

lemma appendS_cong:
  fixes f :: "'a  'b"
  fixes a :: "('a, 's) Stream"
  fixes b :: "('a, 't) Stream"
  shows "a  a'  b  b'  appendSab  appendSa'b'"
unfolding bisimilar_def
by (simp add: unstream_appendS appendS_defined)

lemma appendL_eq: "appendLxsys = unstream(appendS(streamxs)(streamys))"
by (simp add: unstream_appendS)


subsection ‹ZipWith function›

fixrec
  zipWithStep ::
    "('a  'b  'c) 
     ('s  ('a, 's) Step) 
     ('t  ('b, 't) Step) 
      's :!: 't :!: 'a L Maybe  ('c, 's :!: 't :!: 'a L Maybe) Step"
where
  "sa    sb   
   zipWithStepfhahb(sa :!: sb :!: Nothing) =
   (case hasa of
     Done  Done
   | Skipsa'  Skip(sa' :!: sb :!: Nothing)
   | Yieldasa'  Skip(sa' :!: sb :!: Just(La)))"
| "sa    sb   
   zipWithStepfhahb(sa :!: sb :!: Just(La)) =
   (case hbsb of
     Done  Done
   | Skipsb'  Skip(sa :!: sb' :!: Just(La))
   | Yieldbsb'  Yield(fab)(sa :!: sb' :!: Nothing))"

lemma zipWithStep_strict [simp]: "zipWithStepfhahb = "
by fixrec_simp

fixrec
  zipWithS :: "('a  'b  'c) 
      ('a, 's) Stream  ('b, 't) Stream  ('c, 's :!: 't :!: 'a L Maybe) Stream"
where
  "sa0    sb0    zipWithSf(Streamhasa0)(Streamhbsb0) =
    Stream(zipWithStepfhahb)(sa0 :!: sb0 :!: Nothing)"

lemma zipWithS_fix_ind_lemma:
  fixes P Q :: "nat  nat  bool"
  assumes P_0: "j. P 0 j" and P_Suc: "i j. P i j  Q i j  P (Suc i) j"
  assumes Q_0: "i. Q i 0" and Q_Suc: "i j. P i j  Q i j  Q i (Suc j)"
  shows "P i j  Q i j"
apply (induct n  "i + j" arbitrary: i j)
apply (simp add: P_0 Q_0)
apply (rule conjI)
apply (case_tac i, simp add: P_0, simp add: P_Suc)
apply (case_tac j, simp add: Q_0, simp add: Q_Suc)
done

lemma zipWithS_fix_ind:
  assumes x: "x = fixf" and y: "y = fixg"
  assumes adm_P: "adm (λx. P (fst x) (snd x))"
  assumes adm_Q: "adm (λx. Q (fst x) (snd x))"
  assumes P_0: "b. P  b" and P_Suc: "a b. P a b  Q a b  P (fa) b"
  assumes Q_0: "a. Q a " and Q_Suc: "a b. P a b  Q a b  Q a (gb)"
  shows "P x y  Q x y"
proof -
  have 1: "i j. P (iterate if) (iterate jg)  Q (iterate if) (iterate jg)"
    apply (rule_tac i=i and j=j in zipWithS_fix_ind_lemma)
    apply (simp add: P_0)
    apply (simp add: P_Suc)
    apply (simp add: Q_0)
    apply (simp add: Q_Suc)
    done
  have "case_prod P (i. (iterate if, iterate ig))"
    apply (rule admD)
    apply (simp add: split_def adm_P)
    apply simp
    apply (simp add: 1)
    done
  then have P: "P x y"
    unfolding x y fix_def2
    by (simp add: lub_prod)
  have "case_prod Q (i. (iterate if, iterate ig))"
    apply (rule admD)
    apply (simp add: split_def adm_Q)
    apply simp
    apply (simp add: 1)
    done
  then have Q: "Q x y"
    unfolding x y fix_def2
    by (simp add: lub_prod)
  from P Q show ?thesis by simp
qed

lemma unfold_zipWithStep:
  fixes f :: "'a  'b  'c"
  fixes ha :: "'s  ('a, 's) Step"
  fixes hb :: "'t  ('b, 't) Step"
  defines h_def: "h  zipWithStepfhahb"
  shows
  "(sa sb. sa    sb   
    unfoldh(sa :!: sb :!: Nothing) =
      zipWithLf(unfoldhasa)(unfoldhbsb)) 
   (sa sb a. sa    sb   
    unfoldh(sa :!: sb :!: Just(La)) =
      zipWithLf(LConsa(unfoldhasa))(unfoldhbsb))"
proof -
  note unfold [simp]
  have h_simps [simp]:
    "sa sb. sa    sb    h(sa :!: sb :!: Nothing) =
      (case hasa of
        Done  Done
      | Skipsa'  Skip(sa' :!: sb :!: Nothing)
      | Yieldasa'  Skip(sa' :!: sb :!: Just(La)))"
    "sa sb a. sa    sb    h(sa :!: sb :!: Just(La)) =
      (case hbsb of
        Done  Done
      | Skipsb'  Skip(sa :!: sb' :!: Just(La))
      | Yieldbsb'  Yield(fab)(sa :!: sb' :!: Nothing))"
    "h = "
    unfolding h_def by simp_all

  have 1:
  "(sa sb. sa    sb   
    unfoldh(sa :!: sb :!: Nothing) 
      zipWithLf(unfoldhasa)(unfoldhbsb))
   
   (sa sb a. sa    sb   
    unfoldh(sa :!: sb :!: Just(La)) 
      zipWithLf(LConsa(unfoldhasa))(unfoldhbsb))"
    apply (rule unfold_ind [where h="h"], simp)
     apply simp
    apply (intro conjI allI impI)
     apply (case_tac "hasa", simp, simp, simp, simp)
    apply (case_tac "hbsb", simp, simp, simp, simp)
    done

  let ?P = "λua ub. sa sb. sa    sb   
        zipWithLf(uasa)(ubsb)  unfoldh(sa :!: sb :!: Nothing)"

  let ?Q = "λua ub. sa sb a. sa    sb   
        zipWithLf(LConsa(uasa))(ubsb) 
          unfoldh(sa :!: sb :!: Just(La))"

  have P_base: "ub. ?P  ub"
    by simp

  have Q_base: "ua. ?Q ua "
    by simp

  have P_step: "ua ub. ?P ua ub  ?Q ua ub  ?P (unfoldFhaua) ub"
    by (clarsimp, case_tac "hasa", simp_all)

  have Q_step: "ua ub. ?P ua ub  ?Q ua ub  ?Q ua (unfoldFhbub)"
    by (clarsimp, case_tac "hbsb", simp_all)

  have 2: "?P (unfoldha) (unfoldhb)  ?Q (unfoldha) (unfoldhb)"
    apply (rule zipWithS_fix_ind [OF unfold_eq_fix [of ha] unfold_eq_fix [of hb]])
    apply (simp, simp) (* admissibility *)
    apply (rule P_base)
    apply (erule (1) P_step)
    apply (rule Q_base)
    apply (erule (1) Q_step)
    done

  from 1 2 show ?thesis
    by (simp_all add: po_eq_conv [where 'a="'c LList"])
qed

lemma zipWithS_defined: "a    b    zipWithSfab  "
by (cases a, simp, cases b, simp, simp)

lemma unstream_zipWithS:
  "a    b   
    unstream(zipWithSfab) = zipWithLf(unstreama)(unstreamb)"
apply (cases a, simp, cases b, simp)
apply (simp add: unfold_zipWithStep)
done

lemma zipWithS_cong:
  "f = f'  a  a'  b  b' 
    zipWithSfab  zipWithSfa'b'"
unfolding bisimilar_def
by (simp add: unstream_zipWithS zipWithS_defined)

lemma zipWithL_eq:
  "zipWithLfxsys = unstream(zipWithSf(streamxs)(streamys))"
by (simp add: unstream_zipWithS)


subsection ‹ConcatMap function›

fixrec
  concatMapStep ::
    "('a  ('b, 't) Stream) 
     ('s  ('a, 's) Step) 
     's :!: ('b, 't) Stream Maybe 
     ('b, 's :!: ('b, 't) Stream Maybe) Step"
where
  "sa    concatMapStepfha(sa :!: Nothing) =
    (case hasa of
      Done  Done
    | Skipsa'  Skip(sa' :!: Nothing)
    | Yieldasa'  Skip(sa' :!: Just(fa)))"
| "sa    sb   
    concatMapStepfha(sa :!: Just(Streamhbsb)) =
    (case hbsb of
      Done  Skip(sa :!: Nothing)
    | Skipsb'  Skip(sa :!: Just(Streamhbsb'))
    | Yieldbsb'  Yieldb(sa :!: Just(Streamhbsb')))"

lemma concatMapStep_strict [simp]: "concatMapStepfha = "
by fixrec_simp

fixrec
  concatMapS ::
    "('a  ('b, 't) Stream)  ('a, 's) Stream 
     ('b, 's :!: ('b, 't) Stream Maybe) Stream"
where
  "s    concatMapSf(Streamhs) = Stream(concatMapStepfh)(s :!: Nothing)"

lemma concatMapS_strict [simp]: "concatMapSf = "
by fixrec_simp

lemma unfold_concatMapStep:
  fixes ha :: "'s  ('a, 's) Step"
  fixes f :: "'a  ('b, 't) Stream"
  defines h_def: "h  concatMapStepfha"
  defines f'_def: "f'  unstream oo f"
  shows
  "(sa. sa   
      unfoldh(sa :!: Nothing) = concatMapLf'(unfoldhasa)) 
   (sa hb sb. sa    sb   
      unfoldh(sa :!: Just(Streamhbsb)) =
        appendL(unfoldhbsb)(concatMapLf'(unfoldhasa)))"
proof -
  note unfold [simp]
  have h_simps [simp]:
    "sa. sa    h(sa :!: Nothing) =
      (case hasa of Done  Done
      | Skipsa'  Skip(sa' :!: Nothing)
      | Yieldasa'  Skip(sa' :!: Just(fa)))"
    "sa hb sb. sa    sb    h(sa :!: Just(Streamhbsb)) =
      (case hbsb of Done  Skip(sa :!: Nothing)
      | Skipsb'  Skip(sa :!: Just(Streamhbsb'))
      | Yieldbsb'  Yieldb(sa :!: Just(Streamhbsb')))"
    "h = "
    unfolding h_def by simp_all

  have f'_beta [simp]: "a. f'a = unstream(fa)"
    unfolding f'_def by simp

  have 1:
  "(sa. sa   
      unfoldh(sa :!: Nothing)  concatMapLf'(unfoldhasa))
   
   (sa hb sb. sa    sb   
      unfoldh(sa :!: Just(Streamhbsb)) 
        appendL(unfoldhbsb)(concatMapLf'(unfoldhasa)))"
    apply (rule unfold_ind [where h="h"], simp)
     apply simp
    apply (intro conjI allI impI)
     apply (case_tac "hasa", simp, simp, simp)
     apply (rename_tac a sa')
     apply (case_tac "fa", simp, simp)
    apply (case_tac "hbsb", simp, simp, simp, simp)
    done

  let ?P = "λua. sa. sa   
        concatMapLf'(uasa)  unfoldh(sa :!: Nothing)"

  let ?Q = "λhb ua ub. sa sb. sa    sb   
        appendL(ubsb)(concatMapLf'(uasa)) 
            unfoldh(sa :!: Just(Streamhbsb))"

  have P_base: "?P "
    by simp

  have P_step: "ua. ?P ua  hb. ?Q hb ua (unfoldhb)  ?P (unfoldFhaua)"
    apply (intro allI impI)
    apply (case_tac "hasa", simp, simp, simp)
    apply (rename_tac a sa')
    apply (case_tac "fa", simp, simp)
    done

  have Q_base: "ua hb. ?Q hb ua "
    by simp

  have Q_step: "hb ua ub. ?P ua  ?Q hb ua ub  ?Q hb ua (unfoldFhbub)"
    apply (intro allI impI)
    apply (case_tac "hbsb", simp, simp, simp, simp)
    done

  have 2: "?P (unfoldha)  (hb. ?Q hb (unfoldha) (unfoldhb))"
    apply (rule unfold_ind [where h="ha"], simp)
     apply (rule conjI)
      apply (rule P_base)
     apply (rule allI, rule_tac h=hb in unfold_ind, simp)
      apply (rule Q_base)
     apply (erule Q_step [OF P_base])
    apply (erule conjE)
    apply (rule conjI)
     apply (erule (1) P_step)
    apply (rule allI, rule_tac h=hb in unfold_ind, simp)
     apply (rule Q_base)
    apply (erule (2) Q_step [OF P_step])
    done

  from 1 2 show ?thesis
    by (simp_all add: po_eq_conv [where 'a="'b LList"])
qed

lemma unstream_concatMapS:
  "unstream(concatMapSfa) = concatMapL(unstream oo f)(unstreama)"
by (cases a, simp, simp add: unfold_concatMapStep)

lemma concatMapS_defined: "a    concatMapSfa  "
by (induct a, simp_all)

lemma concatMapS_cong:
  fixes f :: "'a  ('b, 's) Stream"
  fixes g :: "'a  ('b, 't) Stream"
  fixes a :: "('a, 'u) Stream"
  fixes b :: "('a, 'v) Stream"
  shows "(x. f x  g x)  a  b  cont f  cont g 
    concatMapS(Λ x. f x)a  concatMapS(Λ x. g x)b"
unfolding bisimilar_def
by (simp add: unstream_concatMapS oo_def concatMapS_defined)

lemma concatMapL_eq:
  "concatMapLfxs = unstream(concatMapS(stream oo f)(streamxs))"
by (simp add: unstream_concatMapS oo_def eta_cfun)


subsection ‹Examples›

lemmas stream_eqs =
  mapL_eq
  filterL_eq
  foldrL_eq
  enumFromToL_eq
  appendL_eq
  zipWithL_eq
  concatMapL_eq

lemmas stream_congs =
  unstream_cong
  stream_cong
  stream_unstream_cong
  mapS_cong
  filterS_cong
  foldrS_cong
  enumFromToS_cong
  appendS_cong
  zipWithS_cong
  concatMapS_cong

lemma
  "mapLf oo filterLp oo mapLg =
   unstream oo mapSf oo filterSp oo mapSg oo stream"
apply (rule cfun_eqI, simp)
apply (unfold stream_eqs)
apply (intro stream_congs refl)
done

lemma
  "foldrLfz(mapLg(filterLp(enumFromToLxy))) =
    foldrSfz(mapSg(filterSp(enumFromToSxy)))"
apply (unfold stream_eqs)
apply (intro stream_congs refl)
done

lemma oo_LAM [simp]: "cont g  f oo (Λ x. g x) = (Λ x. f(g x))"
  unfolding oo_def by simp

lemma
  "concatMapL(Λ k.
    mapL(Λ m. fkm)(enumFromToLonek))(enumFromToLonen) =
   unstream(concatMapS(Λ k.
    mapS(Λ m. fkm)(enumFromToSonek))(enumFromToSonen))"
unfolding stream_eqs
apply simp
apply (simp add: stream_congs)
done

end