Theory Stream_Lift_to_Free3

theory Stream_Lift_to_Free3
imports Stream_Distributive_Law3
header {* Lifting of the distributive law to the free algebra *}


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

theory Stream_Lift_to_Free3
imports Stream_Distributive_Law3
begin

subsection{* The lifting *}

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

(* In order to be able to define dd3, we need a larger codomain type: *)
definition ddd3 :: "('a × 'a F) ΣΣ3 => 'a ΣΣ3 × 'a ΣΣ3 F" where
"ddd3 = ext3 <\<oo>\<pp>3 o Σ3_map fst, F_map flat3 o Λ3> (leaf3 ** F_map leaf3)"

definition dd3 :: "('a × 'a F) ΣΣ3 => 'a ΣΣ3 F" where
"dd3 = snd o ddd3"

lemma ddd3_transfer[transfer_rule]:
  "(ΣΣ3_rel (rel_prod R (F_rel R)) ===> rel_prod (ΣΣ3_rel R) (F_rel (ΣΣ3_rel R))) ddd3 ddd3"
  unfolding ddd3_def ext3_alt by transfer_prover

lemma dd3_transfer[transfer_rule]:
  "(ΣΣ3_rel (rel_prod R (F_rel R)) ===> F_rel (ΣΣ3_rel R)) dd3 dd3"
  unfolding dd3_def by transfer_prover

lemma F_rel_ΣΣ3_rel: "ΣΣ3_rel (rel_prod R (F_rel R)) x y ==> F_rel (ΣΣ3_rel R) (dd3 x) (dd3 y)"
  by (erule rel_funD[OF dd3_transfer])

(* We verify the facts for dd3: *)
theorem dd3_leaf3: "dd3 o leaf3 = F_map leaf3 o snd"
unfolding dd3_def ddd3_def o_assoc[symmetric] ext3_comp_leaf3 snd_comp_map_prod ..

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

theorem dd3_natural: "dd3 o ΣΣ3_map (f ** F_map f) = F_map (ΣΣ3_map f) o dd3"
  using dd3_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding F_rel_Grp prod.rel_Grp ΣΣ3.rel_Grp
  unfolding Grp_def rel_fun_def by auto

lemma Λ3_dd3: "Λ3 = dd3 o \<oo>\<pp>3 o Σ3_map leaf3"
  unfolding dd3_def ddd3_def o_assoc[symmetric] Σ3.map_comp0[symmetric] ext3_commute
  unfolding o_assoc snd_convol ext3_comp_leaf3
  unfolding o_assoc[symmetric] Λ3_natural
  unfolding o_assoc F_map_comp[symmetric] leaf3_flat3 F_map_id id_o
  ..

lemma fst_ddd3: "fst o ddd3 = ΣΣ3_map fst"
proof-
  have "fst o ddd3 = ext3 \<oo>\<pp>3 (leaf3 o fst)"
  apply(rule ext3_unique) unfolding ddd3_def o_assoc[symmetric] ext3_comp_leaf3 ext3_commute
  unfolding o_assoc fst_comp_map_prod fst_convol
  unfolding o_assoc[symmetric] Σ3.map_comp0 by(rule refl, rule refl)
  also have "... = ΣΣ3_map fst"
  apply(rule sym, rule ext3_unique)
  unfolding leaf3_natural \<oo>\<pp>3_natural by(rule refl, rule refl)
  finally show ?thesis .
qed

lemma ddd3_flat3: "(flat3 ** F_map flat3) o ddd3 o ΣΣ3_map ddd3 = ddd3 o flat3" (is "?L = ?R")
proof-
  have "?L = ext3 <\<oo>\<pp>3 o Σ3_map fst, F_map flat3 o Λ3> ddd3"
  proof(rule ext3_unique)
    show "(flat3 ** F_map flat3) o ddd3 o ΣΣ3_map ddd3 o leaf3 = ddd3"
    unfolding ddd3_def unfolding o_assoc[symmetric] leaf3_natural
    unfolding o_assoc
    apply(subst o_assoc[symmetric]) unfolding ext3_comp_leaf3
    unfolding map_prod.comp F_map_comp[symmetric] flat3_leaf3 F_map_id map_prod.id id_o ..
  next
    have A: "<flat3 o (\<oo>\<pp>3 o Σ3_map fst) , F_map flat3 o (F_map flat3 o Λ3)>  =
             <\<oo>\<pp>3 o Σ3_map fst , F_map flat3 o Λ3> o Σ3_map (flat3 ** F_map flat3)"
    unfolding o_assoc unfolding flat3_commute[symmetric]
    apply(rule fst_snd_cong) unfolding o_assoc fst_convol snd_convol
    unfolding o_assoc[symmetric] Σ3.map_comp0[symmetric] fst_comp_map_prod snd_comp_map_prod
    unfolding Λ3_natural unfolding o_assoc F_map_comp[symmetric] flat3_assoc by(rule refl, rule refl)
    show "(flat3 ** F_map flat3) o ddd3 o ΣΣ3_map ddd3 o \<oo>\<pp>3 =
          <\<oo>\<pp>3 o Σ3_map fst , F_map flat3 o Λ3> o Σ3_map (flat3 ** F_map flat3 o ddd3 o ΣΣ3_map ddd3)"
    unfolding ddd3_def unfolding o_assoc[symmetric] unfolding \<oo>\<pp>3_natural[symmetric]
    unfolding o_assoc
    apply(subst o_assoc[symmetric]) unfolding ext3_commute
    unfolding o_assoc[symmetric] Σ3.map_comp0[symmetric]
    unfolding Σ3.map_comp0
    unfolding o_assoc unfolding map_prod_o_convol
    unfolding ext3_ΣΣ3_map[symmetric] A ..
  qed
  also have "... = ?R"
  proof(rule sym, rule ext3_unique)
    show "ddd3 o flat3 o leaf3 = ddd3" unfolding o_assoc[symmetric] flat3_leaf3 o_id ..
  next
    show "ddd3 o flat3 o \<oo>\<pp>3 = <\<oo>\<pp>3 o Σ3_map fst , F_map flat3 o Λ3> o Σ3_map (ddd3 o flat3)"
    unfolding ddd3_def unfolding o_assoc[symmetric] unfolding flat3_commute[symmetric]
    unfolding o_assoc unfolding ext3_commute Σ3.map_comp0 unfolding o_assoc ..
  qed
  finally show ?thesis .
qed

theorem dd3_flat3: "F_map flat3 o dd3 o ΣΣ3_map <ΣΣ3_map fst, dd3> = dd3 o flat3"
proof-
  have A: "snd o ((flat3 ** F_map flat3) o ddd3 o ΣΣ3_map ddd3) = snd o (ddd3 o flat3)"
  unfolding ddd3_flat3 ..
  have B: "ddd3 = <ΣΣ3_map fst , snd o ddd3>" apply(rule fst_snd_cong)
  unfolding fst_ddd3 by auto
  show ?thesis unfolding dd3_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 dd3_leaf3 and dd3_flat3 imply the
more standard conditions for the distributive law dd3' = <ΣΣ3_map fst, dd3>
for the functors ΣΣ3 and 'a F' = 'a × 'a F_ In fact, they can be shown
equivalent to these. *)

lemma dd3_leaf32: "<ΣΣ3_map fst, dd3> o leaf3 = leaf3 ** F_map leaf3"
apply (rule fst_snd_cong) unfolding o_assoc by (simp_all add: leaf3_natural dd3_leaf3)

lemma ddd3_leaf3: "ddd3 o leaf3 = leaf3 ** F_map leaf3"
unfolding ddd3_def ext3_comp_leaf3 ..

lemma ddd3_\<oo>\<pp>3: "ddd3 o \<oo>\<pp>3 = <\<oo>\<pp>3 o Σ3_map fst , F_map flat3 o Λ3> o Σ3_map ddd3"
unfolding ddd3_def ext3_commute ..


(* More customization *)

(* TODO Jasmin: Add3 this high-level induction for the relator of datatypes:
(similarly, coinduction for codatatypes): *)
lemma ΣΣ3_rel_induct_pointfree:
assumes leaf3: "!! x1 x2. R x1 x2 ==> phi (leaf3 x1) (leaf3 x2)"
and \<oo>\<pp>3: "!! y1 y2. [|Σ3_rel (ΣΣ3_rel R) y1 y2; Σ3_rel phi y1 y2|] ==> phi (\<oo>\<pp>3 y1) (\<oo>\<pp>3 y2)"
shows "ΣΣ3_rel R ≤ phi"
proof-
  have "ΣΣ3_rel R ≤ phi \<sqinter> ΣΣ3_rel R"
  apply(induct rule: ΣΣ3.ctor_rel_induct)
  using assms ΣΣ3.rel_inject[of R] unfolding rel_pre_ΣΣ3_def ΣΣ3.leaf3_def ΣΣ3.\<oo>\<pp>3_def
  using inf_greatest[OF Σ3.rel_mono[OF inf_le1] Σ3.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 ΣΣ3_rel_induct[case_names leaf3 \<oo>\<pp>3]:
assumes leaf3: "!! x1 x2. R x1 x2 ==> phi (leaf3 x1) (leaf3 x2)"
and \<oo>\<pp>3: "!! y1 y2. [|Σ3_rel (ΣΣ3_rel R) y1 y2; Σ3_rel phi y1 y2|] ==> phi (\<oo>\<pp>3 y1) (\<oo>\<pp>3 y2)"
shows "ΣΣ3_rel R t1 t2 --> phi t1 t2"
using ΣΣ3_rel_induct_pointfree[of R, OF assms] by auto
(* end TODO *)

end