Theory Stream_Integrate_New_Op4

theory Stream_Integrate_New_Op4
imports Stream_Op_Input4
header{* The integration of a new operation in the up-to setting *}

theory Stream_Integrate_New_Op4
imports Stream_Op_Input4
begin


subsection{* The assumptions *}

(*
(* The operation that creates the new distributive law, since its definition splits
trough a natural transformation ll, which will be defined in More_Corec_Upto4 as follows: *)
definition algρ4 :: "J K4 => J" where
"algρ4 = unfoldU3 (ρ4 o K4_map <id, dtor_J>)"
*)

lemma ρ4_natural: "ρ4 o K4_map (f ** F_map f) = F_map (ΣΣ4_map f) o ρ4"
  using ρ4_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding F_rel_Grp ΣΣ4.rel_Grp K4.rel_Grp prod.rel_Grp
  unfolding Grp_def rel_fun_def by auto

subsection{* The integration *}

definition embL4 :: "'a ΣΣ3 => 'a ΣΣ4" where
"embL4 ≡ ext3 (\<oo>\<pp>4 o Abs_Σ4 o Inl) leaf4"

definition embR4 :: "('a K4) ΣΣ3 => 'a ΣΣ4" where
"embR4 ≡ ext3 (\<oo>\<pp>4 o Abs_Σ4 o Inl) (\<oo>\<pp>4 o Σ4_map leaf4 o Abs_Σ4 o Inr)"

definition Λ4 :: "('a × 'a F) Σ4 => 'a ΣΣ4 F" where
"Λ4 = case_sum (F_map embL4 o Λ3) ρ4 o Rep_Σ4"

lemma embL4_transfer[transfer_rule]:
  "(ΣΣ3_rel R ===> ΣΣ4_rel R) embL4 embL4"
  unfolding embL4_def ext3_alt by transfer_prover

lemma embR4_transfer[transfer_rule]:
  "(ΣΣ3_rel (K4_rel R) ===> ΣΣ4_rel R) embR4 embR4"
  unfolding embR4_def ext3_alt by transfer_prover

lemma Λ4_transfer[transfer_rule]:
  "(Σ4_rel (rel_prod R (F_rel R)) ===> F_rel (ΣΣ4_rel R)) Λ4 Λ4"
  unfolding Λ4_def by transfer_prover

lemma Λ4_natural: "Λ4 o Σ4_map (f ** F_map f) = F_map (ΣΣ4_map f) o Λ4"
  using Λ4_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding F_rel_Grp ΣΣ4.rel_Grp Σ4.rel_Grp prod.rel_Grp
  unfolding Grp_def rel_fun_def by auto

lemma embL4_natural: "embL4 o ΣΣ3_map f = ΣΣ4_map f o embL4"
  using embL4_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding ΣΣ3.rel_Grp ΣΣ4.rel_Grp
  unfolding Grp_def rel_fun_def by auto

lemma embR4_natural: "embR4 o ΣΣ3_map (K4_map f) = ΣΣ4_map f o embR4"
  using embR4_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding ΣΣ3.rel_Grp K4.rel_Grp ΣΣ4.rel_Grp
  unfolding Grp_def rel_fun_def by auto

lemma Λ4_Inl: "Λ4 o Abs_Σ4 o Inl = F_map embL4 o Λ3"
  and Λ4_Inr: "Λ4 o Abs_Σ4 o Inr = ρ4"
unfolding Λ4_def o_assoc[symmetric] Rep_Σ4_o_Abs_Σ4 o_id by auto

lemma embL4_leaf3: "embL4 o leaf3 = leaf4"
unfolding embL4_def ext3_comp_leaf3 ..

lemma embL4_\<oo>\<pp>3: "embL4 o \<oo>\<pp>3 = \<oo>\<pp>4 o Abs_Σ4 o Inl o Σ3_map embL4"
unfolding embL4_def ext3_commute ..

lemma embR4_leaf3: "embR4 o leaf3 = \<oo>\<pp>4 o Abs_Σ4 o Inr o K4_map leaf4"
unfolding embR4_def ext3_comp_leaf3
unfolding o_assoc[symmetric] Abs_Σ4_natural map_sum_Inr ..

lemma embR4_\<oo>\<pp>3: "embR4 o \<oo>\<pp>3 = \<oo>\<pp>4 o Abs_Σ4 o  Inl o Σ3_map embR4"
unfolding embR4_def ext3_commute ..

end