src/HOL/Corec_Examples/Stream_Processor.thy
changeset 80914 d97fdabd9e2b
parent 72184 881bd98bddee
equal deleted inserted replaced
80913:46f59511b7bb 80914:d97fdabd9e2b
    59 by(subst run.code; simp; fail)+
    59 by(subst run.code; simp; fail)+
    60 
    60 
    61 lemma copy_sel[simp]: "out copy = Get (\<lambda>a. Put a copy)"
    61 lemma copy_sel[simp]: "out copy = Get (\<lambda>a. Put a copy)"
    62   by (subst copy.code; simp)
    62   by (subst copy.code; simp)
    63 
    63 
    64 corecursive sp_comp (infixl "oo" 65) where
    64 corecursive sp_comp (infixl \<open>oo\<close> 65) where
    65   "sp oo sp' = (case (out sp, out sp') of
    65   "sp oo sp' = (case (out sp, out sp') of
    66       (Put b sp, _) \<Rightarrow> In (Put b (sp oo sp'))
    66       (Put b sp, _) \<Rightarrow> In (Put b (sp oo sp'))
    67     | (Get f, Put b sp) \<Rightarrow> In (f b) oo sp
    67     | (Get f, Put b sp) \<Rightarrow> In (f b) oo sp
    68     | (_, Get g) \<Rightarrow> get (\<lambda>a. (sp oo In (g a))))"
    68     | (_, Get g) \<Rightarrow> get (\<lambda>a. (sp oo In (g a))))"
    69   by (relation "map_prod In In ` sub <*lex*> map_prod In In ` sub")
    69   by (relation "map_prod In In ` sub <*lex*> map_prod In In ` sub")