Theory Stream_Integrate_New_Op3

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

theory Stream_Integrate_New_Op3
imports Stream_Op_Input3
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_Upto3 as follows: *)
definition algρ3 :: "J K3 => J" where
"algρ3 = unfoldU2 (ρ3 o K3_map <id, dtor_J>)"
*)

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

subsection{* The integration *}

definition embL3 :: "'a ΣΣ2 => 'a ΣΣ3" where
"embL3 ≡ ext2 (\<oo>\<pp>3 o Abs_Σ3 o Inl) leaf3"

definition embR3 :: "('a K3) ΣΣ2 => 'a ΣΣ3" where
"embR3 ≡ ext2 (\<oo>\<pp>3 o Abs_Σ3 o Inl) (\<oo>\<pp>3 o Σ3_map leaf3 o Abs_Σ3 o Inr)"

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

lemma embL3_transfer[transfer_rule]:
  "(ΣΣ2_rel R ===> ΣΣ3_rel R) embL3 embL3"
  unfolding embL3_def ext2_alt by transfer_prover

lemma embR3_transfer[transfer_rule]:
  "(ΣΣ2_rel (K3_rel R) ===> ΣΣ3_rel R) embR3 embR3"
  unfolding embR3_def ext2_alt by transfer_prover

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

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

lemma embL3_natural: "embL3 o ΣΣ2_map f = ΣΣ3_map f o embL3"
  using embL3_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding ΣΣ2.rel_Grp ΣΣ3.rel_Grp
  unfolding Grp_def rel_fun_def by auto

lemma embR3_natural: "embR3 o ΣΣ2_map (K3_map f) = ΣΣ3_map f o embR3"
  using embR3_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding ΣΣ2.rel_Grp K3.rel_Grp ΣΣ3.rel_Grp
  unfolding Grp_def rel_fun_def by auto

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

lemma embL3_leaf2: "embL3 o leaf2 = leaf3"
unfolding embL3_def ext2_comp_leaf2 ..

lemma embL3_\<oo>\<pp>2: "embL3 o \<oo>\<pp>2 = \<oo>\<pp>3 o Abs_Σ3 o Inl o Σ2_map embL3"
unfolding embL3_def ext2_commute ..

lemma embR3_leaf2: "embR3 o leaf2 = \<oo>\<pp>3 o Abs_Σ3 o Inr o K3_map leaf3"
unfolding embR3_def ext2_comp_leaf2
unfolding o_assoc[symmetric] Abs_Σ3_natural map_sum_Inr ..

lemma embR3_\<oo>\<pp>2: "embR3 o \<oo>\<pp>2 = \<oo>\<pp>3 o Abs_Σ3 o  Inl o Σ2_map embR3"
unfolding embR3_def ext2_commute ..

end