Theory Stream_Lift_to_Free4

theory Stream_Lift_to_Free4
imports Stream_Distributive_Law4
header {* Lifting of the distributive law to the free algebra *}


(* This is silimar to Lift_to_Free, but uses Λ4, ΣΣ4, dd4, SpK instead of S, etc. *)

theory Stream_Lift_to_Free4
imports Stream_Distributive_Law4
begin

subsection{* The lifting *}

(* Our aim is lift Λ4 to an (SpK,SpK,T)-distributive law dd4 compatible with the monadic structure. *)

(* In order to be able to define dd4, we need a larger codomain type: *)
definition ddd4 :: "('a × 'a F) ΣΣ4 => 'a ΣΣ4 × 'a ΣΣ4 F" where
"ddd4 = ext4 <\<oo>\<pp>4 o Σ4_map fst, F_map flat4 o Λ4> (leaf4 ** F_map leaf4)"

definition dd4 :: "('a × 'a F) ΣΣ4 => 'a ΣΣ4 F" where
"dd4 = snd o ddd4"

lemma ddd4_transfer[transfer_rule]:
  "(ΣΣ4_rel (rel_prod R (F_rel R)) ===> rel_prod (ΣΣ4_rel R) (F_rel (ΣΣ4_rel R))) ddd4 ddd4"
  unfolding ddd4_def ext4_alt by transfer_prover

lemma dd4_transfer[transfer_rule]:
  "(ΣΣ4_rel (rel_prod R (F_rel R)) ===> F_rel (ΣΣ4_rel R)) dd4 dd4"
  unfolding dd4_def by transfer_prover

lemma F_rel_ΣΣ4_rel: "ΣΣ4_rel (rel_prod R (F_rel R)) x y ==> F_rel (ΣΣ4_rel R) (dd4 x) (dd4 y)"
  by (erule rel_funD[OF dd4_transfer])

(* We verify the facts for dd4: *)
theorem dd4_leaf4: "dd4 o leaf4 = F_map leaf4 o snd"
unfolding dd4_def ddd4_def o_assoc[symmetric] ext4_comp_leaf4 snd_comp_map_prod ..

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

theorem dd4_natural: "dd4 o ΣΣ4_map (f ** F_map f) = F_map (ΣΣ4_map f) o dd4"
  using dd4_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding F_rel_Grp prod.rel_Grp ΣΣ4.rel_Grp
  unfolding Grp_def rel_fun_def by auto

lemma Λ4_dd4: "Λ4 = dd4 o \<oo>\<pp>4 o Σ4_map leaf4"
  unfolding dd4_def ddd4_def o_assoc[symmetric] Σ4.map_comp0[symmetric] ext4_commute
  unfolding o_assoc snd_convol ext4_comp_leaf4
  unfolding o_assoc[symmetric] Λ4_natural
  unfolding o_assoc F_map_comp[symmetric] leaf4_flat4 F_map_id id_o
  ..

lemma fst_ddd4: "fst o ddd4 = ΣΣ4_map fst"
proof-
  have "fst o ddd4 = ext4 \<oo>\<pp>4 (leaf4 o fst)"
  apply(rule ext4_unique) unfolding ddd4_def o_assoc[symmetric] ext4_comp_leaf4 ext4_commute
  unfolding o_assoc fst_comp_map_prod fst_convol
  unfolding o_assoc[symmetric] Σ4.map_comp0 by(rule refl, rule refl)
  also have "... = ΣΣ4_map fst"
  apply(rule sym, rule ext4_unique)
  unfolding leaf4_natural \<oo>\<pp>4_natural by(rule refl, rule refl)
  finally show ?thesis .
qed

lemma ddd4_flat4: "(flat4 ** F_map flat4) o ddd4 o ΣΣ4_map ddd4 = ddd4 o flat4" (is "?L = ?R")
proof-
  have "?L = ext4 <\<oo>\<pp>4 o Σ4_map fst, F_map flat4 o Λ4> ddd4"
  proof(rule ext4_unique)
    show "(flat4 ** F_map flat4) o ddd4 o ΣΣ4_map ddd4 o leaf4 = ddd4"
    unfolding ddd4_def unfolding o_assoc[symmetric] leaf4_natural
    unfolding o_assoc
    apply(subst o_assoc[symmetric]) unfolding ext4_comp_leaf4
    unfolding map_prod.comp F_map_comp[symmetric] flat4_leaf4 F_map_id map_prod.id id_o ..
  next
    have A: "<flat4 o (\<oo>\<pp>4 o Σ4_map fst) , F_map flat4 o (F_map flat4 o Λ4)>  =
             <\<oo>\<pp>4 o Σ4_map fst , F_map flat4 o Λ4> o Σ4_map (flat4 ** F_map flat4)"
    unfolding o_assoc unfolding flat4_commute[symmetric]
    apply(rule fst_snd_cong) unfolding o_assoc fst_convol snd_convol
    unfolding o_assoc[symmetric] Σ4.map_comp0[symmetric] fst_comp_map_prod snd_comp_map_prod
    unfolding Λ4_natural unfolding o_assoc F_map_comp[symmetric] flat4_assoc by(rule refl, rule refl)
    show "(flat4 ** F_map flat4) o ddd4 o ΣΣ4_map ddd4 o \<oo>\<pp>4 =
          <\<oo>\<pp>4 o Σ4_map fst , F_map flat4 o Λ4> o Σ4_map (flat4 ** F_map flat4 o ddd4 o ΣΣ4_map ddd4)"
    unfolding ddd4_def unfolding o_assoc[symmetric] unfolding \<oo>\<pp>4_natural[symmetric]
    unfolding o_assoc
    apply(subst o_assoc[symmetric]) unfolding ext4_commute
    unfolding o_assoc[symmetric] Σ4.map_comp0[symmetric]
    unfolding Σ4.map_comp0
    unfolding o_assoc unfolding map_prod_o_convol
    unfolding ext4_ΣΣ4_map[symmetric] A ..
  qed
  also have "... = ?R"
  proof(rule sym, rule ext4_unique)
    show "ddd4 o flat4 o leaf4 = ddd4" unfolding o_assoc[symmetric] flat4_leaf4 o_id ..
  next
    show "ddd4 o flat4 o \<oo>\<pp>4 = <\<oo>\<pp>4 o Σ4_map fst , F_map flat4 o Λ4> o Σ4_map (ddd4 o flat4)"
    unfolding ddd4_def unfolding o_assoc[symmetric] unfolding flat4_commute[symmetric]
    unfolding o_assoc unfolding ext4_commute Σ4.map_comp0 unfolding o_assoc ..
  qed
  finally show ?thesis .
qed

theorem dd4_flat4: "F_map flat4 o dd4 o ΣΣ4_map <ΣΣ4_map fst, dd4> = dd4 o flat4"
proof-
  have A: "snd o ((flat4 ** F_map flat4) o ddd4 o ΣΣ4_map ddd4) = snd o (ddd4 o flat4)"
  unfolding ddd4_flat4 ..
  have B: "ddd4 = <ΣΣ4_map fst , snd o ddd4>" apply(rule fst_snd_cong)
  unfolding fst_ddd4 by auto
  show ?thesis unfolding dd4_def
  unfolding A[symmetric, unfolded o_assoc snd_comp_map_prod] o_assoc B[symmetric] ..
qed

(* The next two theorems are not necessary for the development.
They show that the conditions dd4_leaf4 and dd4_flat4 imply the
more standard conditions for the distributive law dd4' = <ΣΣ4_map fst, dd4>
for the functors ΣΣ4 and 'a F' = 'a × 'a F_ In fact, they can be shown
equivalent to these. *)

lemma dd4_leaf42: "<ΣΣ4_map fst, dd4> o leaf4 = leaf4 ** F_map leaf4"
apply (rule fst_snd_cong) unfolding o_assoc by (simp_all add: leaf4_natural dd4_leaf4)

lemma ddd4_leaf4: "ddd4 o leaf4 = leaf4 ** F_map leaf4"
unfolding ddd4_def ext4_comp_leaf4 ..

lemma ddd4_\<oo>\<pp>4: "ddd4 o \<oo>\<pp>4 = <\<oo>\<pp>4 o Σ4_map fst , F_map flat4 o Λ4> o Σ4_map ddd4"
unfolding ddd4_def ext4_commute ..


(* More customization *)

(* TODO Jasmin: Add4 this high-level induction for the relator of datatypes:
(similarly, coinduction for codatatypes): *)
lemma ΣΣ4_rel_induct_pointfree:
assumes leaf4: "!! x1 x2. R x1 x2 ==> phi (leaf4 x1) (leaf4 x2)"
and \<oo>\<pp>4: "!! y1 y2. [|Σ4_rel (ΣΣ4_rel R) y1 y2; Σ4_rel phi y1 y2|] ==> phi (\<oo>\<pp>4 y1) (\<oo>\<pp>4 y2)"
shows "ΣΣ4_rel R ≤ phi"
proof-
  have "ΣΣ4_rel R ≤ phi \<sqinter> ΣΣ4_rel R"
  apply(induct rule: ΣΣ4.ctor_rel_induct)
  using assms ΣΣ4.rel_inject[of R] unfolding rel_pre_ΣΣ4_def ΣΣ4.leaf4_def ΣΣ4.\<oo>\<pp>4_def
  using inf_greatest[OF Σ4.rel_mono[OF inf_le1] Σ4.rel_mono[OF inf_le2]]
  unfolding rel_sum_def BNF_Comp.id_bnf_comp_def vimage2p_def by (auto split: sum.splits) blast+
  thus ?thesis by simp
qed

lemma ΣΣ4_rel_induct[case_names leaf4 \<oo>\<pp>4]:
assumes leaf4: "!! x1 x2. R x1 x2 ==> phi (leaf4 x1) (leaf4 x2)"
and \<oo>\<pp>4: "!! y1 y2. [|Σ4_rel (ΣΣ4_rel R) y1 y2; Σ4_rel phi y1 y2|] ==> phi (\<oo>\<pp>4 y1) (\<oo>\<pp>4 y2)"
shows "ΣΣ4_rel R t1 t2 --> phi t1 t2"
using ΣΣ4_rel_induct_pointfree[of R, OF assms] by auto
(* end TODO *)

end