Theory Stream_Integrate_New_Op5

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

theory Stream_Integrate_New_Op5
imports Stream_Op_Input5
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_Upto5 as follows: *)
definition algρ5 :: "J K5 => J" where
"algρ5 = unfoldU4 (ρ5 o K5_map <id, dtor_J>)"
*)

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

subsection{* The integration *}

definition embL5 :: "'a ΣΣ4 => 'a ΣΣ5" where
"embL5 ≡ ext4 (\<oo>\<pp>5 o Abs_Σ5 o Inl) leaf5"

definition embR5 :: "('a K5) ΣΣ4 => 'a ΣΣ5" where
"embR5 ≡ ext4 (\<oo>\<pp>5 o Abs_Σ5 o Inl) (\<oo>\<pp>5 o Σ5_map leaf5 o Abs_Σ5 o Inr)"

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

lemma embL5_transfer[transfer_rule]:
  "(ΣΣ4_rel R ===> ΣΣ5_rel R) embL5 embL5"
  unfolding embL5_def ext4_alt by transfer_prover

lemma embR5_transfer[transfer_rule]:
  "(ΣΣ4_rel (K5_rel R) ===> ΣΣ5_rel R) embR5 embR5"
  unfolding embR5_def ext4_alt by transfer_prover

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

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

lemma embL5_natural: "embL5 o ΣΣ4_map f = ΣΣ5_map f o embL5"
  using embL5_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding ΣΣ4.rel_Grp ΣΣ5.rel_Grp
  unfolding Grp_def rel_fun_def by auto

lemma embR5_natural: "embR5 o ΣΣ4_map (K5_map f) = ΣΣ5_map f o embR5"
  using embR5_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding ΣΣ4.rel_Grp K5.rel_Grp ΣΣ5.rel_Grp
  unfolding Grp_def rel_fun_def by auto

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

lemma embL5_leaf4: "embL5 o leaf4 = leaf5"
unfolding embL5_def ext4_comp_leaf4 ..

lemma embL5_\<oo>\<pp>4: "embL5 o \<oo>\<pp>4 = \<oo>\<pp>5 o Abs_Σ5 o Inl o Σ4_map embL5"
unfolding embL5_def ext4_commute ..

lemma embR5_leaf4: "embR5 o leaf4 = \<oo>\<pp>5 o Abs_Σ5 o Inr o K5_map leaf5"
unfolding embR5_def ext4_comp_leaf4
unfolding o_assoc[symmetric] Abs_Σ5_natural map_sum_Inr ..

lemma embR5_\<oo>\<pp>4: "embR5 o \<oo>\<pp>4 = \<oo>\<pp>5 o Abs_Σ5 o  Inl o Σ4_map embR5"
unfolding embR5_def ext4_commute ..

end