theory Stream_Lib
imports "../Stream/Stream_Behavior_BNF" "../Stream/Stream_More_Corec_Upto5"
begin
type_synonym stream = J
abbreviation "head xs ≡ fst (dtor_J xs)"
abbreviation "tail xs ≡ snd (dtor_J xs)"
abbreviation "SCons n xs ≡ ctor_J (n, xs)"
code_datatype ctor_J
declare J.dtor_ctor[code]
definition smap where
"smap f = corec_J (λxs. (f (head xs), Inr (tail xs)))"
lemma head_smap[simp]: "head (smap f xs) = f (head xs)"
unfolding smap_def corec_J_def J.dtor_corec BNF_Comp.id_bnf_comp_def map_pre_J_def by simp
lemma tail_smap[simp]: "tail (smap f xs) = smap f (tail xs)"
unfolding smap_def corec_J_def J.dtor_corec BNF_Comp.id_bnf_comp_def map_pre_J_def by simp
lemma smap_code[code]: "smap f xs = SCons (f (head xs)) (smap f (tail xs))"
by (metis J.ctor_dtor prod.collapse head_smap tail_smap)
definition same where
"same x = corec_J (λx. (x, Inr x)) x"
lemma head_same[simp]: "head (same x) = x"
unfolding same_def corec_J_def J.dtor_corec BNF_Comp.id_bnf_comp_def map_pre_J_def by simp
lemma tail_same[simp]: "tail (same x) = same x"
unfolding same_def corec_J_def J.dtor_corec BNF_Comp.id_bnf_comp_def map_pre_J_def by simp
definition "sconst x = ctor_J (x, same 0)"
lemma head_sconst[simp]: "head (sconst x) = x"
unfolding sconst_def by (simp add: J.dtor_ctor)
lemma tail_sconst[simp]: "tail (sconst x) = same 0"
unfolding sconst_def by (simp add: J.dtor_ctor)
section {* Abbreviations *}
abbreviation "SCONS0 ≡ gg0 :: 'a ΣΣ0 F => 'a ΣΣ0"
abbreviation "LEAF0 ≡ leaf0"
abbreviation GUARD0 :: "'a F => 'a F ΣΣ0" where "GUARD0 ≡ LEAF0"
abbreviation END0 :: "J => (J + 'a) ΣΣ0" where "END0 xs ≡ LEAF0 (Inl xs)"
abbreviation CONT0 :: "'a => (J + 'a) ΣΣ0" where "CONT0 a ≡ LEAF0 (Inr a)"
abbreviation "SCONS1 ≡ gg1 :: 'a ΣΣ1 F => 'a ΣΣ1"
abbreviation "LEAF1 ≡ leaf1"
abbreviation GUARD1 :: "'a F => 'a F ΣΣ1" where "GUARD1 ≡ LEAF1"
abbreviation END1 :: "J => (J + 'a) ΣΣ1" where "END1 xs ≡ LEAF1 (Inl xs)"
abbreviation CONT1 :: "'a => (J + 'a) ΣΣ1" where "CONT1 a ≡ LEAF1 (Inr a)"
abbreviation "SCONS2 ≡ gg2 :: 'a ΣΣ2 F => 'a ΣΣ2"
abbreviation "LEAF2 ≡ leaf2"
abbreviation GUARD2 :: "'a F => 'a F ΣΣ2" where "GUARD2 ≡ LEAF2"
abbreviation END2 :: "J => (J + 'a) ΣΣ2" where "END2 xs ≡ LEAF2 (Inl xs)"
abbreviation CONT2 :: "'a => (J + 'a) ΣΣ2" where "CONT2 a ≡ LEAF2 (Inr a)"
abbreviation "SCONS3 ≡ gg3 :: 'a ΣΣ3 F => 'a ΣΣ3"
abbreviation "LEAF3 ≡ leaf3"
abbreviation GUARD3 :: "'a F => 'a F ΣΣ3" where "GUARD3 ≡ LEAF3"
abbreviation END3 :: "J => (J + 'a) ΣΣ3" where "END3 xs ≡ LEAF3 (Inl xs)"
abbreviation CONT3 :: "'a => (J + 'a) ΣΣ3" where "CONT3 a ≡ LEAF3 (Inr a)"
abbreviation "SCONS4 ≡ gg4 :: 'a ΣΣ4 F => 'a ΣΣ4"
abbreviation "LEAF4 ≡ leaf4"
abbreviation GUARD4 :: "'a F => 'a F ΣΣ4" where "GUARD4 ≡ LEAF4"
abbreviation END4 :: "J => (J + 'a) ΣΣ4" where "END4 xs ≡ LEAF4 (Inl xs)"
abbreviation CONT4 :: "'a => (J + 'a) ΣΣ4" where "CONT4 a ≡ LEAF4 (Inr a)"
abbreviation "SCONS5 ≡ gg5 :: 'a ΣΣ5 F => 'a ΣΣ5"
abbreviation "LEAF5 ≡ leaf5"
abbreviation GUARD5 :: "'a F => 'a F ΣΣ5" where "GUARD5 ≡ LEAF5"
abbreviation END5 :: "J => (J + 'a) ΣΣ5" where "END5 xs ≡ LEAF5 (Inl xs)"
abbreviation CONT5 :: "'a => (J + 'a) ΣΣ5" where "CONT5 a ≡ LEAF5 (Inr a)"
end