Theory Stream_Corec_Upto4

theory Stream_Corec_Upto4
imports Stream_Lift_to_Free4
header {* Corecursion and coinduction up to *}

theory Stream_Corec_Upto4
imports Stream_Lift_to_Free4
begin


subsection{* The algebra associated to dd4 *}

definition "eval4 ≡ dtor_unfold_J (dd4 o ΣΣ4_map <id, dtor_J>)"

lemma eval4: "F_map eval4 o dd4 o ΣΣ4_map <id, dtor_J> = dtor_J o eval4"
  unfolding eval4_def dtor_unfold_J_pointfree unfolding o_assoc ..

lemma eval4_ctor_J: "ctor_J o F_map eval4 o dd4 o ΣΣ4_map <id, dtor_J> = eval4"
  unfolding o_def spec[OF eval4[unfolded o_def fun_eq_iff]] J.ctor_dtor ..

lemma eval4_leaf4: "eval4 o leaf4 = id"
proof (rule trans)
  show "eval4 o leaf4 = dtor_unfold_J dtor_J"
  apply(rule J.dtor_unfold_unique)
  unfolding o_assoc eval4[symmetric] unfolding o_assoc[symmetric] leaf4_natural
  apply(rule sym)
  unfolding F_map_comp o_assoc apply (subst o_assoc[symmetric])
  unfolding dd4_leaf4 unfolding o_assoc[symmetric] by simp
qed(metis F_map_id J.dtor_unfold_unique fun.map_id o_id)

lemma eval4_flat4: "eval4 o flat4 = eval4 o ΣΣ4_map eval4"
proof (rule trans)
  let ?K4 = "dtor_unfold_J (dd4 o ΣΣ4_map <ΣΣ4_map fst, dd4> o ΣΣ4_map (ΣΣ4_map <id, dtor_J>))"
  show "eval4 o flat4 = ?K4"
  apply(rule J.dtor_unfold_unique)
  unfolding F_map_comp o_assoc
  apply(subst o_assoc[symmetric], subst o_assoc[symmetric])
  unfolding dd4_flat4
  unfolding o_assoc[symmetric] ΣΣ4.map_comp0[symmetric] snd_convol
  unfolding flat4_natural
  unfolding o_assoc eval4 ..
  (*  *)
  have A: "<eval4, dtor_J o eval4> = <id,dtor_J> o eval4" by simp
  show "?K4 = eval4 o ΣΣ4_map eval4"
  apply(rule J.dtor_unfold_unique[symmetric])
  unfolding o_assoc[symmetric] ΣΣ4.map_comp0[symmetric] map_prod_o_convol id_o
  unfolding F_map_comp o_assoc
  apply(subst o_assoc[symmetric]) unfolding dd4_natural[symmetric]
  unfolding o_assoc[symmetric] ΣΣ4.map_comp0[symmetric]
  unfolding o_assoc unfolding map_prod_o_convol unfolding convol_o
  unfolding o_assoc[symmetric] ΣΣ4.map_comp0[symmetric] fst_convol ΣΣ4.map_id0 o_id
  unfolding o_assoc eval4 unfolding A unfolding convol_o id_o
  apply(rule sym) apply(subst eval4[symmetric])
  unfolding o_assoc[symmetric] ΣΣ4.map_comp0[symmetric] convol_o id_o ..
qed


subsection{* The correspondence between coalgebras up to and coalgebras *}

(* This correspondence, and add4 the lemmas of this subsection, make
sense not only for (J,dtor_J,eval4), but also for any dd4-bialgebra *)

(* Coalgebra-up-to to coalgebra and vice versa: *)
definition cutΣΣ4Oc :: "('a => 'a ΣΣ4 F) => ('a ΣΣ4 => 'a ΣΣ4 F)"
where "cutΣΣ4Oc s ≡ F_map flat4 o dd4 o ΣΣ4_map <leaf4, s>"
definition cΣΣ4Ocut :: "('a ΣΣ4 => 'a ΣΣ4 F) => ('a => 'a ΣΣ4 F)"
where "cΣΣ4Ocut s' ≡ s' o leaf4"

lemma cΣΣ4Ocut_cutΣΣ4Oc: "cΣΣ4Ocut (cutΣΣ4Oc s) = s"
unfolding cΣΣ4Ocut_def cutΣΣ4Oc_def
unfolding o_assoc[symmetric] unfolding leaf4_natural
unfolding o_assoc apply(subst o_assoc[symmetric])
unfolding dd4_leaf4 unfolding o_assoc F_map_comp[symmetric] flat4_leaf4
unfolding F_map_id id_o by simp

lemma cutΣΣ4Oc_inj: "cutΣΣ4Oc s4 = cutΣΣ4Oc s2 <-> s4 = s2"
by (metis cΣΣ4Ocut_cutΣΣ4Oc)

lemma cΣΣ4Ocut_surj: "∃ s'. cΣΣ4Ocut s' = s"
using cΣΣ4Ocut_cutΣΣ4Oc by(rule exI[of _ "cutΣΣ4Oc s"])

(* Morphism-up-to to morphism and vice versa: *)
(* The dd4-extension of a function *)
definition extdd4 :: "('a => J) => ('a ΣΣ4 => J)"
where "extdd4 f ≡ eval4 o ΣΣ4_map f"
(* The restriction of a function *)  term eval4
definition restr :: "('a ΣΣ4 => J) => ('a => J)"
where "restr f' ≡ f' o leaf4"

(* We think of extdd4 and restr as operating
extdd4 : Hom_dd4(s,dtor_J) -> Hom(cutΣΣ4Oc s,dtor_J) and
restr : Hom(cutΣΣ4Oc s,dtor_J) -> Hom_dd4(s,dtor_J), where
Hom(cutΣΣ4Oc s,dtor_J) is the set of coalgebra morphisms betwee cutΣΣ4Oc s and dtor_J and
Hom_dd4(s,dtor_J) is the set of coalgebra up-to-dd4-morphisms between s and dtor_J  *)

(* extdd4 is wedd4-defined: *)
lemma extdd4_mor:
assumes f:  "F_map (extdd4 f) o s = dtor_J o f"
shows "F_map (extdd4 f) o cutΣΣ4Oc s = dtor_J o (extdd4 f)"
proof-
  have AA: "eval4 ** F_map eval4 o (ΣΣ4_map f ** F_map (ΣΣ4_map f) o <leaf4 , s>) =
            <f , F_map eval4 o (F_map (ΣΣ4_map f) o s)>"
  unfolding map_prod_o_convol unfolding leaf4_natural o_assoc eval4_leaf4 id_o  ..
  show ?thesis
  unfolding extdd4_def
  unfolding o_assoc eval4[symmetric]
  unfolding o_assoc[symmetric] ΣΣ4.map_comp0[symmetric]
  unfolding convol_comp[symmetric] id_o
  unfolding f[symmetric, unfolded extdd4_def]
  unfolding o_assoc
  apply(subst o_assoc[symmetric])
  unfolding F_map_comp o_assoc
  unfolding cutΣΣ4Oc_def
  unfolding o_assoc
  unfolding F_map_comp[symmetric] unfolding o_assoc[symmetric]
  unfolding flat4_natural[symmetric]
  unfolding o_assoc F_map_comp
  apply(subst o_assoc[symmetric], subst o_assoc[symmetric], subst o_assoc[symmetric])
  unfolding dd4_natural[symmetric]
  unfolding o_assoc unfolding F_map_comp[symmetric] eval4_flat4
  unfolding F_map_comp apply(subst o_assoc[symmetric])
  unfolding dd4_natural[symmetric] unfolding o_assoc[symmetric] ΣΣ4.map_comp0[symmetric]
  unfolding o_assoc[symmetric] AA[unfolded o_assoc[symmetric]] ..
qed

lemma mor_cutΣΣ4Oc_flat4:
assumes f': "F_map f' o cutΣΣ4Oc s = dtor_J o f'"
shows "eval4 o ΣΣ4_map f' = f' o flat4"
proof(rule trans) (* this proof is clearly improvable: *)
  def h  "dd4 o ΣΣ4_map <id,cutΣΣ4Oc s>"
  have f'_id: "f' = f' o id" by simp
  show "eval4 o ΣΣ4_map f' = dtor_unfold_J h"
  apply(rule J.dtor_unfold_unique, rule sym)
  unfolding o_assoc eval4[symmetric]
  unfolding o_assoc[symmetric] ΣΣ4.map_comp0[symmetric]
  unfolding convol_comp_id1[symmetric] unfolding f'[symmetric]
  apply(subst f'_id)
  unfolding o_assoc ΣΣ4.map_comp0
  apply(subst o_assoc[symmetric])
  unfolding o_assoc[symmetric] F_map_comp
  unfolding h_def apply(rule sym)
  unfolding o_assoc apply(subst o_assoc[symmetric])
  unfolding dd4_natural[symmetric] unfolding o_assoc[symmetric]
  unfolding ΣΣ4.map_comp0[symmetric] map_prod_o_convol ..
  (*  *)
  have AA: "<id , cutΣΣ4Oc s> = (flat4 ** F_map flat4) o (id ** dd4) o <leaf4, ΣΣ4_map <leaf4 , s>>"
  unfolding map_prod_o_convol o_assoc map_prod.comp cutΣΣ4Oc_def o_id flat4_leaf4 ..
  have BB: "flat4 ** F_map flat4 o id ** dd4 o <leaf4 , ΣΣ4_map <leaf4 , s>> = flat4 ** F_map flat4 o id ** dd4 o <ΣΣ4_map leaf4 , ΣΣ4_map <leaf4 , s>>"
  unfolding map_prod.comp unfolding map_prod_o_convol unfolding o_id unfolding flat4_leaf4 leaf4_flat4 ..
  show "dtor_unfold_J h = f' o flat4"
  apply(rule J.dtor_unfold_unique[symmetric], rule sym)
  unfolding o_assoc f'[symmetric]
  unfolding F_map_comp o_assoc[symmetric]
  apply(rule arg_cong[of _ _ "op o (F_map f')"])
  unfolding h_def
  unfolding AA BB
  unfolding ΣΣ4.map_comp0 apply(rule sym)
  unfolding o_assoc apply(subst o_assoc[symmetric])
  unfolding dd4_natural
  unfolding o_assoc F_map_comp[symmetric]
  unfolding flat4_assoc unfolding F_map_comp
  unfolding cutΣΣ4Oc_def o_assoc[symmetric] apply(rule arg_cong[of _ _ "op o (F_map flat4)"])
  unfolding o_assoc
  unfolding o_assoc[symmetric] unfolding ΣΣ4.map_comp0[symmetric] unfolding map_prod_o_convol id_o
  unfolding flat4_natural[symmetric] unfolding o_assoc
  unfolding dd4_flat4[symmetric] unfolding o_assoc[symmetric] unfolding ΣΣ4.map_comp0[symmetric]
  unfolding convol_o unfolding ΣΣ4.map_comp0[symmetric] unfolding fst_convol ..
qed

(* restr is wedd4-defined:  *)
lemma restr_mor:
assumes f': "F_map f' o cutΣΣ4Oc s = dtor_J o f'"
shows "F_map (extdd4 (restr f')) o s = dtor_J o restr f'"
unfolding extdd4_def restr_def ΣΣ4.map_comp0
unfolding o_assoc mor_cutΣΣ4Oc_flat4[OF f']
unfolding o_assoc[symmetric] leaf4_flat4 o_id
unfolding o_assoc f'[symmetric]
unfolding o_assoc[symmetric] cΣΣ4Ocut_cutΣΣ4Oc[unfolded cΣΣ4Ocut_def] ..

lemma extdd4_restr:
assumes f': "F_map f' o cutΣΣ4Oc s = dtor_J o f'"
shows "extdd4 (restr f') = f'"
proof-
  have "f' = eval4 o ΣΣ4_map f' o leaf4"
  unfolding o_assoc[symmetric] leaf4_natural
  unfolding o_assoc eval4_leaf4 by simp
  also have "... = eval4 o ΣΣ4_map (f' o leaf4)"
  unfolding ΣΣ4.map_comp0 o_assoc
  unfolding mor_cutΣΣ4Oc_flat4[OF f'] unfolding o_assoc[symmetric] flat4_leaf4 leaf4_flat4 ..
  finally have A: "f' = eval4 o ΣΣ4_map (f' o leaf4)" .
  show ?thesis unfolding extdd4_def restr_def A[symmetric] ..
qed

lemma restr_inj:
assumes f4': "F_map f4' o cutΣΣ4Oc s = dtor_J o f4'"
and f2': "F_map f2' o cutΣΣ4Oc s = dtor_J o f2'"
shows "restr f4' = restr f2' <-> f4' = f2'"
using extdd4_restr[OF f4'] extdd4_restr[OF f2'] by metis

lemma extdd4_surj:
assumes f': "F_map f' o cutΣΣ4Oc s = dtor_J o f'"
shows "∃ f. extdd4 f = f'"
using extdd4_restr[OF f'] by(rule exI[of _ "restr f'"])

lemma restr_extdd4:
assumes f: "F_map (extdd4 f) o s = dtor_J o f"
shows "restr (extdd4 f) = f"
proof-
  have "dtor_J o f = F_map (extdd4 f) o s" using assms unfolding extdd4_def by (rule sym)
  also have "... = dtor_J o restr (extdd4 f)"
  unfolding restr_def unfolding o_assoc extdd4_mor[OF f, symmetric]
  unfolding o_assoc[symmetric] cΣΣ4Ocut_cutΣΣ4Oc[unfolded cΣΣ4Ocut_def] ..
  finally have "dtor_J o f = dtor_J o restr (extdd4 f)" .
  thus ?thesis unfolding dtor_J_o_inj by (rule sym)
qed

lemma extdd4_inj:
assumes f1: "F_map (extdd4 f1) o s = dtor_J o f1"
and f2: "F_map (extdd4 f2) o s = dtor_J o f2"
shows "extdd4 f1 = extdd4 f2 <-> f1 = f2"
using restr_extdd4[OF f1] restr_extdd4[OF f2] by metis

lemma restr_surj:
assumes f: "F_map (extdd4 f) o s = dtor_J o f"
shows "∃ f'. restr f' = f"
using restr_extdd4[OF f] by(rule exI[of _ "extdd4 f"])


subsection{* Coiteration up-to *}

definition "unfoldU4 s ≡ restr (dtor_unfold_J (cutΣΣ4Oc s))"

theorem unfoldU4_pointfree:
"F_map (extdd4 (unfoldU4 s)) o s = dtor_J o unfoldU4 s"
unfolding unfoldU4_def apply(rule restr_mor)
unfolding dtor_unfold_J_pointfree ..

theorem unfoldU4: "F_map (extdd4 (unfoldU4 s)) (s a) = dtor_J (unfoldU4 s a)"
using unfoldU4_pointfree unfolding o_def fun_eq_iff by(rule allE)

theorem unfoldU4_ctor_J:
"ctor_J (F_map (extdd4 (unfoldU4 s)) (s a)) = unfoldU4 s a"
using unfoldU4 by (metis J.ctor_dtor)

theorem unfoldU4_unique:
assumes "F_map (extdd4 f) o s = dtor_J o f"
shows "f = unfoldU4 s"
proof-
  note f = extdd4_mor[OF assms]  note co = extdd4_mor[OF unfoldU4_pointfree]
  have A: "extdd4 f = extdd4 (unfoldU4 s)"
  proof(rule trans)
    show "extdd4 f = dtor_unfold_J (cutΣΣ4Oc s)" apply(rule J.dtor_unfold_unique) using f .
    show "dtor_unfold_J (cutΣΣ4Oc s) = extdd4 (unfoldU4 s)"
     apply(rule J.dtor_unfold_unique[symmetric]) using co .
  qed
  show ?thesis using A unfolding extdd4_inj[OF assms unfoldU4_pointfree] .
qed

lemma unfoldU4_ctor_J_pointfree:
"ctor_J o F_map (extdd4 (unfoldU4 s)) o s = unfoldU4 s"
unfolding o_def fun_eq_iff by (subst unfoldU4_ctor_J[symmetric]) (rule allI, rule refl)

(* Corecursion up to: *)
definition corecU4 :: "('a => (J + 'a) ΣΣ4 F) => 'a => J" where
"corecU4 s = unfoldU4 (case_sum (dd4 o leaf4 o <Inl, F_map Inl o dtor_J>) s) o Inr"

definition extddRec4 where
"extddRec4 f ≡ eval4 o ΣΣ4_map (case_sum id f)"

lemma unfoldU4_Inl:
"unfoldU4 (case_sum (dd4 o leaf4 o <Inl , F_map Inl o dtor_J>) s) o Inl = id"
(is "?L = ?R")
proof-
  have "?L = unfoldU4 (dd4 o leaf4 o <id, dtor_J>)"
  apply(rule unfoldU4_unique)
  unfolding o_assoc unfoldU4_pointfree[symmetric]
  unfolding o_assoc[symmetric] case_sum_o_inj extdd4_def F_map_comp ΣΣ4.map_comp0
  unfolding o_assoc apply(subst o_assoc[symmetric], subst o_assoc[symmetric], subst o_assoc[symmetric],
                          subst o_assoc[symmetric], subst o_assoc[symmetric], subst o_assoc[symmetric])
  unfolding dd4_natural[symmetric]
  apply(subst o_assoc[symmetric]) unfolding leaf4_natural
  unfolding o_assoc[symmetric] map_prod_o_convol o_id ..
  also have "... = ?R"
  apply(rule sym, rule unfoldU4_unique)
  unfolding extdd4_def ΣΣ4.map_id0 o_id dd4_leaf4
  unfolding o_assoc[symmetric] snd_convol
  unfolding o_assoc F_map_comp[symmetric] eval4_leaf4 F_map_id id_o ..
  finally show ?thesis .
qed

theorem corecU4_pointfree:
"F_map (extddRec4 (corecU4 s)) o s = dtor_J o corecU4 s" (is "?L = ?R")
unfolding corecU4_def
unfolding o_assoc unfoldU4_pointfree[symmetric] extddRec4_def
unfolding o_assoc[symmetric] case_sum_o_inj
apply(subst unfoldU4_Inl[symmetric, of s])
unfolding o_assoc case_sum_Inl_Inr_L extdd4_def ..

theorem corecU4:
"F_map (extddRec4 (corecU4 s)) (s a) = dtor_J (corecU4 s a)"
using corecU4_pointfree unfolding o_def fun_eq_iff by(rule allE)


subsection{* Coinduction up-to *}

definition "cptdd4 R ≡ (ΣΣ4_rel R ===> R) eval4 eval4"

definition "cngdd4 R ≡ equivp R ∧ cptdd4 R"

lemma cngdd4_Retr: "cngdd4 R ==> cngdd4 (R \<sqinter> Retr R)"
  unfolding cngdd4_def cptdd4_def
  apply (erule conjE)
  apply (rule conjI[OF equivp_inf[OF _ equivp_retr]])
  apply assumption
  apply assumption
  apply (rule rel_funI)
  apply (frule predicate2D[OF ΣΣ4_rel_inf])
  apply (erule inf2E)
  apply (rule inf2I)
  apply (erule rel_funE)
  apply assumption
  apply assumption
  apply (subst Retr_def)
  apply (subst eval4_def)+
  apply (subst J.dtor_unfold)+
  unfolding F_rel_F_map_F_map Grp_def relcompp.simps[abs_def] conversep.simps[abs_def]
  apply auto
  unfolding eval4_def[symmetric]
  apply (rule predicate2D[OF F_rel_mono])
  apply (rule predicate2I)
  apply (erule rel_funD)
  apply assumption
  apply (rule F_rel_ΣΣ4_rel)
  unfolding ΣΣ4_rel_ΣΣ4_map_ΣΣ4_map vimage2p_rel_prod vimage2p_id
  unfolding vimage2p_def Retr_def[symmetric]
  apply assumption
  done

(* The generated congruence: *)
definition "genCngdd4 R j1 j2 ≡ ∀ R'. R ≤ R' ∧ cngdd4 R' --> R' j1 j2"

lemma cngdd4_genCngdd4: "cngdd4 (genCngdd4 R)"
unfolding cngdd4_def proof safe
  show "cptdd4 (genCngdd4 R)"
  unfolding cptdd4_def rel_fun_def proof safe
    fix x y assume A: "ΣΣ4_rel (genCngdd4 R) x y"
    show "genCngdd4 R (eval4 x) (eval4 y)"
    unfolding genCngdd4_def[abs_def] proof safe
      fix R' assume "R ≤ R'" and 2: "cngdd4 R'"
      hence "ΣΣ4_rel R' x y" by (metis A ΣΣ4.rel_mono_strong genCngdd4_def)
      thus "R' (eval4 x) (eval4 y)" using 2 unfolding cngdd4_def cptdd4_def rel_fun_def by auto
    qed
  qed
qed(rule equivpI, unfold reflp_def symp_def transp_def genCngdd4_def cngdd4_def equivp_def, auto)

lemma
    genCngdd4_refl[intro,simp]: "genCngdd4 R j j"
and genCngdd4_sym[intro]: "genCngdd4 R j1 j2 ==> genCngdd4 R j2 j1"
and genCngdd4_trans[intro]: "[|genCngdd4 R j1 j2; genCngdd4 R j2 j3|] ==> genCngdd4 R j1 j3"
using cngdd4_genCngdd4 unfolding cngdd4_def equivp_def by auto

lemma genCngdd4_eval4_rel_fun: "(ΣΣ4_rel (genCngdd4 R) ===> genCngdd4 R) eval4 eval4"
using cngdd4_genCngdd4 unfolding cngdd4_def cptdd4_def by auto

lemma genCngdd4_eval4: "ΣΣ4_rel (genCngdd4 R) x y ==> genCngdd4 R (eval4 x) (eval4 y)"
using genCngdd4_eval4_rel_fun unfolding rel_fun_def by auto

lemma leq_genCngdd4: "R ≤ genCngdd4 R"
and imp_genCngdd4[intro]: "R j1 j2 ==> genCngdd4 R j1 j2"
unfolding genCngdd4_def[abs_def] by auto

lemma genCngdd4_minimal: "[|R ≤ R'; cngdd4 R'|] ==> genCngdd4 R ≤ R'"
unfolding genCngdd4_def[abs_def] by (metis (lifting, no_types) predicate2I)

theorem coinductionU_genCngdd4:
assumes "∀ a b. R a b --> F_rel (genCngdd4 R) (dtor_J a) (dtor_J b)"
shows "R a b --> a = b"
proof-
  let ?R' = "genCngdd4 R"
  have "R ≤ Retr ?R'" using assms unfolding Retr_def[abs_def] by auto
  hence "R ≤ ?R' \<sqinter> Retr ?R'" using leq_genCngdd4 by auto
  moreover have "cngdd4 (?R' \<sqinter> Retr ?R')" using cngdd4_Retr[OF cngdd4_genCngdd4[of R]] .
  ultimately have "?R' ≤ ?R' \<sqinter> Retr ?R'" using genCngdd4_minimal by metis
  hence "?R' ≤ Retr ?R'" by auto
  hence "?R' a b --> a = b"  unfolding Retr_def[abs_def] by (intro J.dtor_coinduct) auto
  thus ?thesis using leq_genCngdd4 by auto
qed


subsection{* Flattening from term algebra to "one-step" algebra *}

(* Since (J ΣΣ4, eval4) is an Eilenberg-Moore algebra (i.e., is compatible with
the monadic structure, it yields an algebra on the signature functor.
This is crucial for having suitable simplification rules. *)

(* The one-step algebra (naturally associated to Λ4, not to dd4): *)
definition algΛ4 :: "J Σ4 => J" where "algΛ4 = eval4 o \<oo>\<pp>4 o Σ4_map leaf4"

theorem eval4_comp_\<oo>\<pp>4: "eval4 o \<oo>\<pp>4 = algΛ4 o Σ4_map eval4"
unfolding algΛ4_def
unfolding o_assoc[symmetric] Σ4.map_comp0[symmetric]
unfolding leaf4_natural[symmetric] unfolding Σ4.map_comp0
apply(rule sym) unfolding o_assoc apply(subst o_assoc[symmetric])
unfolding \<oo>\<pp>4_natural
unfolding o_assoc eval4_flat4[symmetric]
apply(subst o_assoc[symmetric]) unfolding flat4_commute[symmetric]
unfolding o_assoc[symmetric] Σ4.map_comp0[symmetric] flat4_leaf4 Σ4.map_id0 o_id ..

theorem eval4_\<oo>\<pp>4: "eval4 (\<oo>\<pp>4 t) = algΛ4 (Σ4_map eval4 t)"
using eval4_comp_\<oo>\<pp>4 unfolding o_def fun_eq_iff by (rule allE)

theorem eval4_leaf4': "eval4 (leaf4 j) = j"
using eval4_leaf4 unfolding o_def fun_eq_iff id_def by (rule allE)

theorem dtor_J_algΛ4: "dtor_J o algΛ4 = F_map eval4 o Λ4 o Σ4_map <id, dtor_J>"
  unfolding algΛ4_def eval4_def o_assoc dtor_unfold_J_pointfree Λ4_dd4
  unfolding o_assoc[symmetric] \<oo>\<pp>4_natural[symmetric] Σ4.map_comp0[symmetric] leaf4_natural
  ..

theorem dtor_J_algΛ4': "dtor_J (algΛ4 t) = F_map eval4 (Λ4 (Σ4_map (<id, dtor_J>) t))"
  by (rule trans[OF o_eq_dest[OF dtor_J_algΛ4] o_apply])

theorem ctor_J_algΛ4: "algΛ4 t = ctor_J (F_map eval4 (Λ4 (Σ4_map (<id, dtor_J>) t)))"
  by (rule iffD1[OF J.dtor_inject trans[OF dtor_J_algΛ4' J.dtor_ctor[symmetric]]])

(* Note: The user wiΛ4 indicate an "equation" equ :: J Σ4 => (J ΣΣ4) F
in order to define the operation algΛ4 :: J Σ4 => J on the final coalgebra
according to dtor_J_algΛ4.
The package wiΛ4 identify the polymorphic function Λ4 :: ('a × 'a F) Σ4 => 'a ΣΣ4 F
such that equ = Λ4 o Σ4 <id, dtor_J>. Then the package wiΛ4 attempt to prove
automaticaΛ4y that Λ4 is natural.  If the proof fails, the user wiΛ4 be asked to prove it.
Upon success, the distributive law is being registered.
*)

(* Customizing coinduction up-to: *)
definition "cptΛ4 R ≡ (Σ4_rel R ===> R) algΛ4 algΛ4"

definition "cngΛ4 R ≡ equivp R ∧ cptΛ4 R"

lemma cptdd4_iff_cptΛ4: "cptdd4 R <-> cptΛ4 R"
apply (rule iffI)
apply (unfold cptΛ4_def cptdd4_def algΛ4_def o_assoc[symmetric]) [1]
apply (erule rel_funD[OF rel_funD[OF comp_transfer]])
apply transfer_prover

apply (unfold cptΛ4_def cptdd4_def) [1]
unfolding rel_fun_def
apply (rule allI)+
apply (rule ΣΣ4_rel_induct)
apply (simp only: eval4_leaf4')
unfolding eval4_\<oo>\<pp>4
apply (drule spec2)
apply (erule mp)
apply (rule iffD2[OF Σ4_rel_Σ4_map_Σ4_map])
apply (subst vimage2p_def)
apply assumption
done

(* This is the definition of genCngdd4 we need to work with: *)
theorem genCngdd4_def2: "genCngdd4 R j1 j2 <-> (∀ R'. R ≤ R' ∧ cngΛ4 R' --> R' j1 j2)"
unfolding genCngdd4_def cngdd4_def cngΛ4_def cptdd4_iff_cptΛ4 ..


subsection {* Incremental coinduction *}

interpretation I4: Incremental where Retr = Retr and Cl = genCngdd4
proof
  show "mono Retr" by (rule mono_retr)
next
  show "mono genCngdd4" unfolding mono_def
  unfolding genCngdd4_def[abs_def] by (smt order.trans predicate2I)
next
  fix r show "genCngdd4 (genCngdd4 r) = genCngdd4 r"
  by (metis cngdd4_genCngdd4 genCngdd4_minimal leq_genCngdd4 order_class.order.antisym)
next
  fix r show "r ≤ genCngdd4 r" by (metis leq_genCngdd4)
next
  fix r assume "genCngdd4 r = r" thus "genCngdd4 (r \<sqinter> Retr r) = r \<sqinter> Retr r"
  by (metis antisym cngdd4_Retr cngdd4_genCngdd4 genCngdd4_minimal
            inf.orderI inf_idem leq_genCngdd4)
qed

definition ded4 where "ded4 r s ≡ I4.ded r s"

theorems Ax = I4.Ax'[unfolded ded4_def[symmetric]]
theorems Split = I4.Split[unfolded ded4_def[symmetric]]
theorems Coind = I4.Coind[unfolded ded4_def[symmetric]]

theorem soundness_ded:
assumes "ded4 (op =) s"  shows "s ≤ (op =)"
unfolding gfp_Retr_eq[symmetric] apply(rule I4.soundness_ded[unfolded ded4_def[symmetric]] )
using assms unfolding gfp_Retr_eq[symmetric] ded4_def .

unused_thms

end