header {* Lifting of the distributive law to the free algebra *}
theory Stream_Lift_to_Free5
imports Stream_Distributive_Law5
begin
subsection{* The lifting *}
definition ddd5 :: "('a × 'a F) ΣΣ5 => 'a ΣΣ5 × 'a ΣΣ5 F" where
"ddd5 = ext5 <\<oo>\<pp>5 o Σ5_map fst, F_map flat5 o Λ5> (leaf5 ** F_map leaf5)"
definition dd5 :: "('a × 'a F) ΣΣ5 => 'a ΣΣ5 F" where
"dd5 = snd o ddd5"
lemma ddd5_transfer[transfer_rule]:
"(ΣΣ5_rel (rel_prod R (F_rel R)) ===> rel_prod (ΣΣ5_rel R) (F_rel (ΣΣ5_rel R))) ddd5 ddd5"
unfolding ddd5_def ext5_alt by transfer_prover
lemma dd5_transfer[transfer_rule]:
"(ΣΣ5_rel (rel_prod R (F_rel R)) ===> F_rel (ΣΣ5_rel R)) dd5 dd5"
unfolding dd5_def by transfer_prover
lemma F_rel_ΣΣ5_rel: "ΣΣ5_rel (rel_prod R (F_rel R)) x y ==> F_rel (ΣΣ5_rel R) (dd5 x) (dd5 y)"
by (erule rel_funD[OF dd5_transfer])
theorem dd5_leaf5: "dd5 o leaf5 = F_map leaf5 o snd"
unfolding dd5_def ddd5_def o_assoc[symmetric] ext5_comp_leaf5 snd_comp_map_prod ..
lemma ddd5_natural: "ddd5 o ΣΣ5_map (f ** F_map f) = (ΣΣ5_map f ** F_map (ΣΣ5_map f)) o ddd5"
using ddd5_transfer[of "BNF_Def.Grp UNIV f"]
unfolding F_rel_Grp prod.rel_Grp ΣΣ5.rel_Grp
unfolding Grp_def rel_fun_def by auto
theorem dd5_natural: "dd5 o ΣΣ5_map (f ** F_map f) = F_map (ΣΣ5_map f) o dd5"
using dd5_transfer[of "BNF_Def.Grp UNIV f"]
unfolding F_rel_Grp prod.rel_Grp ΣΣ5.rel_Grp
unfolding Grp_def rel_fun_def by auto
lemma Λ5_dd5: "Λ5 = dd5 o \<oo>\<pp>5 o Σ5_map leaf5"
unfolding dd5_def ddd5_def o_assoc[symmetric] Σ5.map_comp0[symmetric] ext5_commute
unfolding o_assoc snd_convol ext5_comp_leaf5
unfolding o_assoc[symmetric] Λ5_natural
unfolding o_assoc F_map_comp[symmetric] leaf5_flat5 F_map_id id_o
..
lemma fst_ddd5: "fst o ddd5 = ΣΣ5_map fst"
proof-
have "fst o ddd5 = ext5 \<oo>\<pp>5 (leaf5 o fst)"
apply(rule ext5_unique) unfolding ddd5_def o_assoc[symmetric] ext5_comp_leaf5 ext5_commute
unfolding o_assoc fst_comp_map_prod fst_convol
unfolding o_assoc[symmetric] Σ5.map_comp0 by(rule refl, rule refl)
also have "... = ΣΣ5_map fst"
apply(rule sym, rule ext5_unique)
unfolding leaf5_natural \<oo>\<pp>5_natural by(rule refl, rule refl)
finally show ?thesis .
qed
lemma ddd5_flat5: "(flat5 ** F_map flat5) o ddd5 o ΣΣ5_map ddd5 = ddd5 o flat5" (is "?L = ?R")
proof-
have "?L = ext5 <\<oo>\<pp>5 o Σ5_map fst, F_map flat5 o Λ5> ddd5"
proof(rule ext5_unique)
show "(flat5 ** F_map flat5) o ddd5 o ΣΣ5_map ddd5 o leaf5 = ddd5"
unfolding ddd5_def unfolding o_assoc[symmetric] leaf5_natural
unfolding o_assoc
apply(subst o_assoc[symmetric]) unfolding ext5_comp_leaf5
unfolding map_prod.comp F_map_comp[symmetric] flat5_leaf5 F_map_id map_prod.id id_o ..
next
have A: "<flat5 o (\<oo>\<pp>5 o Σ5_map fst) , F_map flat5 o (F_map flat5 o Λ5)> =
<\<oo>\<pp>5 o Σ5_map fst , F_map flat5 o Λ5> o Σ5_map (flat5 ** F_map flat5)"
unfolding o_assoc unfolding flat5_commute[symmetric]
apply(rule fst_snd_cong) unfolding o_assoc fst_convol snd_convol
unfolding o_assoc[symmetric] Σ5.map_comp0[symmetric] fst_comp_map_prod snd_comp_map_prod
unfolding Λ5_natural unfolding o_assoc F_map_comp[symmetric] flat5_assoc by(rule refl, rule refl)
show "(flat5 ** F_map flat5) o ddd5 o ΣΣ5_map ddd5 o \<oo>\<pp>5 =
<\<oo>\<pp>5 o Σ5_map fst , F_map flat5 o Λ5> o Σ5_map (flat5 ** F_map flat5 o ddd5 o ΣΣ5_map ddd5)"
unfolding ddd5_def unfolding o_assoc[symmetric] unfolding \<oo>\<pp>5_natural[symmetric]
unfolding o_assoc
apply(subst o_assoc[symmetric]) unfolding ext5_commute
unfolding o_assoc[symmetric] Σ5.map_comp0[symmetric]
unfolding Σ5.map_comp0
unfolding o_assoc unfolding map_prod_o_convol
unfolding ext5_ΣΣ5_map[symmetric] A ..
qed
also have "... = ?R"
proof(rule sym, rule ext5_unique)
show "ddd5 o flat5 o leaf5 = ddd5" unfolding o_assoc[symmetric] flat5_leaf5 o_id ..
next
show "ddd5 o flat5 o \<oo>\<pp>5 = <\<oo>\<pp>5 o Σ5_map fst , F_map flat5 o Λ5> o Σ5_map (ddd5 o flat5)"
unfolding ddd5_def unfolding o_assoc[symmetric] unfolding flat5_commute[symmetric]
unfolding o_assoc unfolding ext5_commute Σ5.map_comp0 unfolding o_assoc ..
qed
finally show ?thesis .
qed
theorem dd5_flat5: "F_map flat5 o dd5 o ΣΣ5_map <ΣΣ5_map fst, dd5> = dd5 o flat5"
proof-
have A: "snd o ((flat5 ** F_map flat5) o ddd5 o ΣΣ5_map ddd5) = snd o (ddd5 o flat5)"
unfolding ddd5_flat5 ..
have B: "ddd5 = <ΣΣ5_map fst , snd o ddd5>" apply(rule fst_snd_cong)
unfolding fst_ddd5 by auto
show ?thesis unfolding dd5_def
unfolding A[symmetric, unfolded o_assoc snd_comp_map_prod] o_assoc B[symmetric] ..
qed
lemma dd5_leaf52: "<ΣΣ5_map fst, dd5> o leaf5 = leaf5 ** F_map leaf5"
apply (rule fst_snd_cong) unfolding o_assoc by (simp_all add: leaf5_natural dd5_leaf5)
lemma ddd5_leaf5: "ddd5 o leaf5 = leaf5 ** F_map leaf5"
unfolding ddd5_def ext5_comp_leaf5 ..
lemma ddd5_\<oo>\<pp>5: "ddd5 o \<oo>\<pp>5 = <\<oo>\<pp>5 o Σ5_map fst , F_map flat5 o Λ5> o Σ5_map ddd5"
unfolding ddd5_def ext5_commute ..
lemma ΣΣ5_rel_induct_pointfree:
assumes leaf5: "!! x1 x2. R x1 x2 ==> phi (leaf5 x1) (leaf5 x2)"
and \<oo>\<pp>5: "!! y1 y2. [|Σ5_rel (ΣΣ5_rel R) y1 y2; Σ5_rel phi y1 y2|] ==> phi (\<oo>\<pp>5 y1) (\<oo>\<pp>5 y2)"
shows "ΣΣ5_rel R ≤ phi"
proof-
have "ΣΣ5_rel R ≤ phi \<sqinter> ΣΣ5_rel R"
apply(induct rule: ΣΣ5.ctor_rel_induct)
using assms ΣΣ5.rel_inject[of R] unfolding rel_pre_ΣΣ5_def ΣΣ5.leaf5_def ΣΣ5.\<oo>\<pp>5_def
using inf_greatest[OF Σ5.rel_mono[OF inf_le1] Σ5.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 ΣΣ5_rel_induct[case_names leaf5 \<oo>\<pp>5]:
assumes leaf5: "!! x1 x2. R x1 x2 ==> phi (leaf5 x1) (leaf5 x2)"
and \<oo>\<pp>5: "!! y1 y2. [|Σ5_rel (ΣΣ5_rel R) y1 y2; Σ5_rel phi y1 y2|] ==> phi (\<oo>\<pp>5 y1) (\<oo>\<pp>5 y2)"
shows "ΣΣ5_rel R t1 t2 --> phi t1 t2"
using ΣΣ5_rel_induct_pointfree[of R, OF assms] by auto
end