Theory Stream_Corec_Upto3

theory Stream_Corec_Upto3
imports Stream_Lift_to_Free3
header {* Corecursion and coinduction up to *}

theory Stream_Corec_Upto3
imports Stream_Lift_to_Free3
begin


subsection{* The algebra associated to dd3 *}

definition "eval3 ≡ dtor_unfold_J (dd3 o ΣΣ3_map <id, dtor_J>)"

lemma eval3: "F_map eval3 o dd3 o ΣΣ3_map <id, dtor_J> = dtor_J o eval3"
  unfolding eval3_def dtor_unfold_J_pointfree unfolding o_assoc ..

lemma eval3_ctor_J: "ctor_J o F_map eval3 o dd3 o ΣΣ3_map <id, dtor_J> = eval3"
  unfolding o_def spec[OF eval3[unfolded o_def fun_eq_iff]] J.ctor_dtor ..

lemma eval3_leaf3: "eval3 o leaf3 = id"
proof (rule trans)
  show "eval3 o leaf3 = dtor_unfold_J dtor_J"
  apply(rule J.dtor_unfold_unique)
  unfolding o_assoc eval3[symmetric] unfolding o_assoc[symmetric] leaf3_natural
  apply(rule sym)
  unfolding F_map_comp o_assoc apply (subst o_assoc[symmetric])
  unfolding dd3_leaf3 unfolding o_assoc[symmetric] by simp
qed(metis F_map_id J.dtor_unfold_unique fun.map_id o_id)

lemma eval3_flat3: "eval3 o flat3 = eval3 o ΣΣ3_map eval3"
proof (rule trans)
  let ?K3 = "dtor_unfold_J (dd3 o ΣΣ3_map <ΣΣ3_map fst, dd3> o ΣΣ3_map (ΣΣ3_map <id, dtor_J>))"
  show "eval3 o flat3 = ?K3"
  apply(rule J.dtor_unfold_unique)
  unfolding F_map_comp o_assoc
  apply(subst o_assoc[symmetric], subst o_assoc[symmetric])
  unfolding dd3_flat3
  unfolding o_assoc[symmetric] ΣΣ3.map_comp0[symmetric] snd_convol
  unfolding flat3_natural
  unfolding o_assoc eval3 ..
  (*  *)
  have A: "<eval3, dtor_J o eval3> = <id,dtor_J> o eval3" by simp
  show "?K3 = eval3 o ΣΣ3_map eval3"
  apply(rule J.dtor_unfold_unique[symmetric])
  unfolding o_assoc[symmetric] ΣΣ3.map_comp0[symmetric] map_prod_o_convol id_o
  unfolding F_map_comp o_assoc
  apply(subst o_assoc[symmetric]) unfolding dd3_natural[symmetric]
  unfolding o_assoc[symmetric] ΣΣ3.map_comp0[symmetric]
  unfolding o_assoc unfolding map_prod_o_convol unfolding convol_o
  unfolding o_assoc[symmetric] ΣΣ3.map_comp0[symmetric] fst_convol ΣΣ3.map_id0 o_id
  unfolding o_assoc eval3 unfolding A unfolding convol_o id_o
  apply(rule sym) apply(subst eval3[symmetric])
  unfolding o_assoc[symmetric] ΣΣ3.map_comp0[symmetric] convol_o id_o ..
qed


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

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

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

lemma cΣΣ3Ocut_cutΣΣ3Oc: "cΣΣ3Ocut (cutΣΣ3Oc s) = s"
unfolding cΣΣ3Ocut_def cutΣΣ3Oc_def
unfolding o_assoc[symmetric] unfolding leaf3_natural
unfolding o_assoc apply(subst o_assoc[symmetric])
unfolding dd3_leaf3 unfolding o_assoc F_map_comp[symmetric] flat3_leaf3
unfolding F_map_id id_o by simp

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

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

(* Morphism-up-to to morphism and vice versa: *)
(* The dd3-extension of a function *)
definition extdd3 :: "('a => J) => ('a ΣΣ3 => J)"
where "extdd3 f ≡ eval3 o ΣΣ3_map f"
(* The restriction of a function *)  term eval3
definition restr :: "('a ΣΣ3 => J) => ('a => J)"
where "restr f' ≡ f' o leaf3"

(* We think of extdd3 and restr as operating
extdd3 : Hom_dd3(s,dtor_J) -> Hom(cutΣΣ3Oc s,dtor_J) and
restr : Hom(cutΣΣ3Oc s,dtor_J) -> Hom_dd3(s,dtor_J), where
Hom(cutΣΣ3Oc s,dtor_J) is the set of coalgebra morphisms betwee cutΣΣ3Oc s and dtor_J and
Hom_dd3(s,dtor_J) is the set of coalgebra up-to-dd3-morphisms between s and dtor_J  *)

(* extdd3 is wedd3-defined: *)
lemma extdd3_mor:
assumes f:  "F_map (extdd3 f) o s = dtor_J o f"
shows "F_map (extdd3 f) o cutΣΣ3Oc s = dtor_J o (extdd3 f)"
proof-
  have AA: "eval3 ** F_map eval3 o (ΣΣ3_map f ** F_map (ΣΣ3_map f) o <leaf3 , s>) =
            <f , F_map eval3 o (F_map (ΣΣ3_map f) o s)>"
  unfolding map_prod_o_convol unfolding leaf3_natural o_assoc eval3_leaf3 id_o  ..
  show ?thesis
  unfolding extdd3_def
  unfolding o_assoc eval3[symmetric]
  unfolding o_assoc[symmetric] ΣΣ3.map_comp0[symmetric]
  unfolding convol_comp[symmetric] id_o
  unfolding f[symmetric, unfolded extdd3_def]
  unfolding o_assoc
  apply(subst o_assoc[symmetric])
  unfolding F_map_comp o_assoc
  unfolding cutΣΣ3Oc_def
  unfolding o_assoc
  unfolding F_map_comp[symmetric] unfolding o_assoc[symmetric]
  unfolding flat3_natural[symmetric]
  unfolding o_assoc F_map_comp
  apply(subst o_assoc[symmetric], subst o_assoc[symmetric], subst o_assoc[symmetric])
  unfolding dd3_natural[symmetric]
  unfolding o_assoc unfolding F_map_comp[symmetric] eval3_flat3
  unfolding F_map_comp apply(subst o_assoc[symmetric])
  unfolding dd3_natural[symmetric] unfolding o_assoc[symmetric] ΣΣ3.map_comp0[symmetric]
  unfolding o_assoc[symmetric] AA[unfolded o_assoc[symmetric]] ..
qed

lemma mor_cutΣΣ3Oc_flat3:
assumes f': "F_map f' o cutΣΣ3Oc s = dtor_J o f'"
shows "eval3 o ΣΣ3_map f' = f' o flat3"
proof(rule trans) (* this proof is clearly improvable: *)
  def h  "dd3 o ΣΣ3_map <id,cutΣΣ3Oc s>"
  have f'_id: "f' = f' o id" by simp
  show "eval3 o ΣΣ3_map f' = dtor_unfold_J h"
  apply(rule J.dtor_unfold_unique, rule sym)
  unfolding o_assoc eval3[symmetric]
  unfolding o_assoc[symmetric] ΣΣ3.map_comp0[symmetric]
  unfolding convol_comp_id1[symmetric] unfolding f'[symmetric]
  apply(subst f'_id)
  unfolding o_assoc ΣΣ3.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 dd3_natural[symmetric] unfolding o_assoc[symmetric]
  unfolding ΣΣ3.map_comp0[symmetric] map_prod_o_convol ..
  (*  *)
  have AA: "<id , cutΣΣ3Oc s> = (flat3 ** F_map flat3) o (id ** dd3) o <leaf3, ΣΣ3_map <leaf3 , s>>"
  unfolding map_prod_o_convol o_assoc map_prod.comp cutΣΣ3Oc_def o_id flat3_leaf3 ..
  have BB: "flat3 ** F_map flat3 o id ** dd3 o <leaf3 , ΣΣ3_map <leaf3 , s>> = flat3 ** F_map flat3 o id ** dd3 o <ΣΣ3_map leaf3 , ΣΣ3_map <leaf3 , s>>"
  unfolding map_prod.comp unfolding map_prod_o_convol unfolding o_id unfolding flat3_leaf3 leaf3_flat3 ..
  show "dtor_unfold_J h = f' o flat3"
  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 ΣΣ3.map_comp0 apply(rule sym)
  unfolding o_assoc apply(subst o_assoc[symmetric])
  unfolding dd3_natural
  unfolding o_assoc F_map_comp[symmetric]
  unfolding flat3_assoc unfolding F_map_comp
  unfolding cutΣΣ3Oc_def o_assoc[symmetric] apply(rule arg_cong[of _ _ "op o (F_map flat3)"])
  unfolding o_assoc
  unfolding o_assoc[symmetric] unfolding ΣΣ3.map_comp0[symmetric] unfolding map_prod_o_convol id_o
  unfolding flat3_natural[symmetric] unfolding o_assoc
  unfolding dd3_flat3[symmetric] unfolding o_assoc[symmetric] unfolding ΣΣ3.map_comp0[symmetric]
  unfolding convol_o unfolding ΣΣ3.map_comp0[symmetric] unfolding fst_convol ..
qed

(* restr is wedd3-defined:  *)
lemma restr_mor:
assumes f': "F_map f' o cutΣΣ3Oc s = dtor_J o f'"
shows "F_map (extdd3 (restr f')) o s = dtor_J o restr f'"
unfolding extdd3_def restr_def ΣΣ3.map_comp0
unfolding o_assoc mor_cutΣΣ3Oc_flat3[OF f']
unfolding o_assoc[symmetric] leaf3_flat3 o_id
unfolding o_assoc f'[symmetric]
unfolding o_assoc[symmetric] cΣΣ3Ocut_cutΣΣ3Oc[unfolded cΣΣ3Ocut_def] ..

lemma extdd3_restr:
assumes f': "F_map f' o cutΣΣ3Oc s = dtor_J o f'"
shows "extdd3 (restr f') = f'"
proof-
  have "f' = eval3 o ΣΣ3_map f' o leaf3"
  unfolding o_assoc[symmetric] leaf3_natural
  unfolding o_assoc eval3_leaf3 by simp
  also have "... = eval3 o ΣΣ3_map (f' o leaf3)"
  unfolding ΣΣ3.map_comp0 o_assoc
  unfolding mor_cutΣΣ3Oc_flat3[OF f'] unfolding o_assoc[symmetric] flat3_leaf3 leaf3_flat3 ..
  finally have A: "f' = eval3 o ΣΣ3_map (f' o leaf3)" .
  show ?thesis unfolding extdd3_def restr_def A[symmetric] ..
qed

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

lemma extdd3_surj:
assumes f': "F_map f' o cutΣΣ3Oc s = dtor_J o f'"
shows "∃ f. extdd3 f = f'"
using extdd3_restr[OF f'] by(rule exI[of _ "restr f'"])

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

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

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


subsection{* Coiteration up-to *}

definition "unfoldU3 s ≡ restr (dtor_unfold_J (cutΣΣ3Oc s))"

theorem unfoldU3_pointfree:
"F_map (extdd3 (unfoldU3 s)) o s = dtor_J o unfoldU3 s"
unfolding unfoldU3_def apply(rule restr_mor)
unfolding dtor_unfold_J_pointfree ..

theorem unfoldU3: "F_map (extdd3 (unfoldU3 s)) (s a) = dtor_J (unfoldU3 s a)"
using unfoldU3_pointfree unfolding o_def fun_eq_iff by(rule allE)

theorem unfoldU3_ctor_J:
"ctor_J (F_map (extdd3 (unfoldU3 s)) (s a)) = unfoldU3 s a"
using unfoldU3 by (metis J.ctor_dtor)

theorem unfoldU3_unique:
assumes "F_map (extdd3 f) o s = dtor_J o f"
shows "f = unfoldU3 s"
proof-
  note f = extdd3_mor[OF assms]  note co = extdd3_mor[OF unfoldU3_pointfree]
  have A: "extdd3 f = extdd3 (unfoldU3 s)"
  proof(rule trans)
    show "extdd3 f = dtor_unfold_J (cutΣΣ3Oc s)" apply(rule J.dtor_unfold_unique) using f .
    show "dtor_unfold_J (cutΣΣ3Oc s) = extdd3 (unfoldU3 s)"
     apply(rule J.dtor_unfold_unique[symmetric]) using co .
  qed
  show ?thesis using A unfolding extdd3_inj[OF assms unfoldU3_pointfree] .
qed

lemma unfoldU3_ctor_J_pointfree:
"ctor_J o F_map (extdd3 (unfoldU3 s)) o s = unfoldU3 s"
unfolding o_def fun_eq_iff by (subst unfoldU3_ctor_J[symmetric]) (rule allI, rule refl)

(* Corecursion up to: *)
definition corecU3 :: "('a => (J + 'a) ΣΣ3 F) => 'a => J" where
"corecU3 s = unfoldU3 (case_sum (dd3 o leaf3 o <Inl, F_map Inl o dtor_J>) s) o Inr"

definition extddRec3 where
"extddRec3 f ≡ eval3 o ΣΣ3_map (case_sum id f)"

lemma unfoldU3_Inl:
"unfoldU3 (case_sum (dd3 o leaf3 o <Inl , F_map Inl o dtor_J>) s) o Inl = id"
(is "?L = ?R")
proof-
  have "?L = unfoldU3 (dd3 o leaf3 o <id, dtor_J>)"
  apply(rule unfoldU3_unique)
  unfolding o_assoc unfoldU3_pointfree[symmetric]
  unfolding o_assoc[symmetric] case_sum_o_inj extdd3_def F_map_comp ΣΣ3.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 dd3_natural[symmetric]
  apply(subst o_assoc[symmetric]) unfolding leaf3_natural
  unfolding o_assoc[symmetric] map_prod_o_convol o_id ..
  also have "... = ?R"
  apply(rule sym, rule unfoldU3_unique)
  unfolding extdd3_def ΣΣ3.map_id0 o_id dd3_leaf3
  unfolding o_assoc[symmetric] snd_convol
  unfolding o_assoc F_map_comp[symmetric] eval3_leaf3 F_map_id id_o ..
  finally show ?thesis .
qed

theorem corecU3_pointfree:
"F_map (extddRec3 (corecU3 s)) o s = dtor_J o corecU3 s" (is "?L = ?R")
unfolding corecU3_def
unfolding o_assoc unfoldU3_pointfree[symmetric] extddRec3_def
unfolding o_assoc[symmetric] case_sum_o_inj
apply(subst unfoldU3_Inl[symmetric, of s])
unfolding o_assoc case_sum_Inl_Inr_L extdd3_def ..

theorem corecU3:
"F_map (extddRec3 (corecU3 s)) (s a) = dtor_J (corecU3 s a)"
using corecU3_pointfree unfolding o_def fun_eq_iff by(rule allE)


subsection{* Coinduction up-to *}

definition "cptdd3 R ≡ (ΣΣ3_rel R ===> R) eval3 eval3"

definition "cngdd3 R ≡ equivp R ∧ cptdd3 R"

lemma cngdd3_Retr: "cngdd3 R ==> cngdd3 (R \<sqinter> Retr R)"
  unfolding cngdd3_def cptdd3_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 ΣΣ3_rel_inf])
  apply (erule inf2E)
  apply (rule inf2I)
  apply (erule rel_funE)
  apply assumption
  apply assumption
  apply (subst Retr_def)
  apply (subst eval3_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 eval3_def[symmetric]
  apply (rule predicate2D[OF F_rel_mono])
  apply (rule predicate2I)
  apply (erule rel_funD)
  apply assumption
  apply (rule F_rel_ΣΣ3_rel)
  unfolding ΣΣ3_rel_ΣΣ3_map_ΣΣ3_map vimage2p_rel_prod vimage2p_id
  unfolding vimage2p_def Retr_def[symmetric]
  apply assumption
  done

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

lemma cngdd3_genCngdd3: "cngdd3 (genCngdd3 R)"
unfolding cngdd3_def proof safe
  show "cptdd3 (genCngdd3 R)"
  unfolding cptdd3_def rel_fun_def proof safe
    fix x y assume A: "ΣΣ3_rel (genCngdd3 R) x y"
    show "genCngdd3 R (eval3 x) (eval3 y)"
    unfolding genCngdd3_def[abs_def] proof safe
      fix R' assume "R ≤ R'" and 2: "cngdd3 R'"
      hence "ΣΣ3_rel R' x y" by (metis A ΣΣ3.rel_mono_strong genCngdd3_def)
      thus "R' (eval3 x) (eval3 y)" using 2 unfolding cngdd3_def cptdd3_def rel_fun_def by auto
    qed
  qed
qed(rule equivpI, unfold reflp_def symp_def transp_def genCngdd3_def cngdd3_def equivp_def, auto)

lemma
    genCngdd3_refl[intro,simp]: "genCngdd3 R j j"
and genCngdd3_sym[intro]: "genCngdd3 R j1 j2 ==> genCngdd3 R j2 j1"
and genCngdd3_trans[intro]: "[|genCngdd3 R j1 j2; genCngdd3 R j2 j3|] ==> genCngdd3 R j1 j3"
using cngdd3_genCngdd3 unfolding cngdd3_def equivp_def by auto

lemma genCngdd3_eval3_rel_fun: "(ΣΣ3_rel (genCngdd3 R) ===> genCngdd3 R) eval3 eval3"
using cngdd3_genCngdd3 unfolding cngdd3_def cptdd3_def by auto

lemma genCngdd3_eval3: "ΣΣ3_rel (genCngdd3 R) x y ==> genCngdd3 R (eval3 x) (eval3 y)"
using genCngdd3_eval3_rel_fun unfolding rel_fun_def by auto

lemma leq_genCngdd3: "R ≤ genCngdd3 R"
and imp_genCngdd3[intro]: "R j1 j2 ==> genCngdd3 R j1 j2"
unfolding genCngdd3_def[abs_def] by auto

lemma genCngdd3_minimal: "[|R ≤ R'; cngdd3 R'|] ==> genCngdd3 R ≤ R'"
unfolding genCngdd3_def[abs_def] by (metis (lifting, no_types) predicate2I)

theorem coinductionU_genCngdd3:
assumes "∀ a b. R a b --> F_rel (genCngdd3 R) (dtor_J a) (dtor_J b)"
shows "R a b --> a = b"
proof-
  let ?R' = "genCngdd3 R"
  have "R ≤ Retr ?R'" using assms unfolding Retr_def[abs_def] by auto
  hence "R ≤ ?R' \<sqinter> Retr ?R'" using leq_genCngdd3 by auto
  moreover have "cngdd3 (?R' \<sqinter> Retr ?R')" using cngdd3_Retr[OF cngdd3_genCngdd3[of R]] .
  ultimately have "?R' ≤ ?R' \<sqinter> Retr ?R'" using genCngdd3_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_genCngdd3 by auto
qed


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

(* Since (J ΣΣ3, eval3) 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 Λ3, not to dd3): *)
definition algΛ3 :: "J Σ3 => J" where "algΛ3 = eval3 o \<oo>\<pp>3 o Σ3_map leaf3"

theorem eval3_comp_\<oo>\<pp>3: "eval3 o \<oo>\<pp>3 = algΛ3 o Σ3_map eval3"
unfolding algΛ3_def
unfolding o_assoc[symmetric] Σ3.map_comp0[symmetric]
unfolding leaf3_natural[symmetric] unfolding Σ3.map_comp0
apply(rule sym) unfolding o_assoc apply(subst o_assoc[symmetric])
unfolding \<oo>\<pp>3_natural
unfolding o_assoc eval3_flat3[symmetric]
apply(subst o_assoc[symmetric]) unfolding flat3_commute[symmetric]
unfolding o_assoc[symmetric] Σ3.map_comp0[symmetric] flat3_leaf3 Σ3.map_id0 o_id ..

theorem eval3_\<oo>\<pp>3: "eval3 (\<oo>\<pp>3 t) = algΛ3 (Σ3_map eval3 t)"
using eval3_comp_\<oo>\<pp>3 unfolding o_def fun_eq_iff by (rule allE)

theorem eval3_leaf3': "eval3 (leaf3 j) = j"
using eval3_leaf3 unfolding o_def fun_eq_iff id_def by (rule allE)

theorem dtor_J_algΛ3: "dtor_J o algΛ3 = F_map eval3 o Λ3 o Σ3_map <id, dtor_J>"
  unfolding algΛ3_def eval3_def o_assoc dtor_unfold_J_pointfree Λ3_dd3
  unfolding o_assoc[symmetric] \<oo>\<pp>3_natural[symmetric] Σ3.map_comp0[symmetric] leaf3_natural
  ..

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

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

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

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

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

lemma cptdd3_iff_cptΛ3: "cptdd3 R <-> cptΛ3 R"
apply (rule iffI)
apply (unfold cptΛ3_def cptdd3_def algΛ3_def o_assoc[symmetric]) [1]
apply (erule rel_funD[OF rel_funD[OF comp_transfer]])
apply transfer_prover

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

(* This is the definition of genCngdd3 we need to work with: *)
theorem genCngdd3_def2: "genCngdd3 R j1 j2 <-> (∀ R'. R ≤ R' ∧ cngΛ3 R' --> R' j1 j2)"
unfolding genCngdd3_def cngdd3_def cngΛ3_def cptdd3_iff_cptΛ3 ..


subsection {* Incremental coinduction *}

interpretation I3: Incremental where Retr = Retr and Cl = genCngdd3
proof
  show "mono Retr" by (rule mono_retr)
next
  show "mono genCngdd3" unfolding mono_def
  unfolding genCngdd3_def[abs_def] by (smt order.trans predicate2I)
next
  fix r show "genCngdd3 (genCngdd3 r) = genCngdd3 r"
  by (metis cngdd3_genCngdd3 genCngdd3_minimal leq_genCngdd3 order_class.order.antisym)
next
  fix r show "r ≤ genCngdd3 r" by (metis leq_genCngdd3)
next
  fix r assume "genCngdd3 r = r" thus "genCngdd3 (r \<sqinter> Retr r) = r \<sqinter> Retr r"
  by (metis antisym cngdd3_Retr cngdd3_genCngdd3 genCngdd3_minimal
            inf.orderI inf_idem leq_genCngdd3)
qed

definition ded3 where "ded3 r s ≡ I3.ded r s"

theorems Ax = I3.Ax'[unfolded ded3_def[symmetric]]
theorems Split = I3.Split[unfolded ded3_def[symmetric]]
theorems Coind = I3.Coind[unfolded ded3_def[symmetric]]

theorem soundness_ded:
assumes "ded3 (op =) s"  shows "s ≤ (op =)"
unfolding gfp_Retr_eq[symmetric] apply(rule I3.soundness_ded[unfolded ded3_def[symmetric]] )
using assms unfolding gfp_Retr_eq[symmetric] ded3_def .

unused_thms

end