Theory Tree_FreeAlg0

theory Tree_FreeAlg0
imports Tree_Behavior_BNF
header{* The initial free algebra for bootstrapping, for a copy of the behavior BNF *}

theory Tree_FreeAlg0
imports Tree_Behavior_BNF
begin

(* Initially, the signature functor and the behavior functor coincide: Σ0 = (a copy of) F
This corresponds to corecursion with any > 0 number of guards. *)

typedef 'a Σ0 = "UNIV :: 'a F set" by (rule UNIV_witness)

setup_lifting type_definition_Σ0

lift_definition Σ0_map :: "('a => 'b) => 'a Σ0 => 'b Σ0" is F_map .
lift_definition Σ0_set :: "'a Σ0 => 'a set" is F_set .
lift_definition Σ0_rel :: "('a => 'b => bool) => 'a Σ0 => 'b Σ0 => bool" is F_rel .

typedef Σ0_bd_type = "UNIV :: bd_type_F set" by (rule UNIV_witness)
definition "Σ0_bd = dir_image F_bd Abs_Σ0_bd_type"

bnf "'a Σ0"
  map: Σ0_map
  sets: Σ0_set
  bd: Σ0_bd
  rel: Σ0_rel
unfolding Σ0_bd_def
apply -
apply transfer apply (rule pre_J.map_id0)
apply transfer apply (rule pre_J.map_comp0)
apply transfer apply (erule pre_J.map_cong0)
apply transfer apply (rule pre_J.set_map0)
apply (rule card_order_dir_image[OF bijI pre_J.bd_card_order])
apply (metis inj_on_def Abs_Σ0_bd_type_inverse[OF UNIV_I])
apply (metis surj_def Abs_Σ0_bd_type_cases)
apply (rule conjunct1[OF Cinfinite_cong[OF dir_image[OF _ pre_J.bd_Card_order] pre_J.bd_Cinfinite]])
apply (metis Abs_Σ0_bd_type_inverse[OF UNIV_I])
apply (unfold Σ0_set_def map_fun_def id_o) [1] apply (subst o_apply)
apply (rule ordLeq_ordIso_trans[OF pre_J.set_bd dir_image[OF _ pre_J.bd_Card_order]])
apply (metis Abs_Σ0_bd_type_inverse[OF UNIV_I])
apply (rule predicate2I) apply transfer apply (subst pre_J.rel_compp) apply assumption
apply transfer' apply (rule pre_J.rel_compp_Grp)
done

declare Σ0.map_transfer[transfer_rule]

lemma Rep_Σ0_transfer[transfer_rule]: "(Σ0_rel R ===> F_rel R) Rep_Σ0 Rep_Σ0"
  unfolding rel_fun_def by transfer blast

lemma Abs_Σ0_transfer[transfer_rule]: "(F_rel R ===> Σ0_rel R) Abs_Σ0 Abs_Σ0"
  unfolding rel_fun_def by transfer blast

theorem Rep_Σ0_natural: "F_map f o Rep_Σ0 = Rep_Σ0 o Σ0_map f"
  using Rep_Σ0_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding Σ0.rel_Grp F_rel_Grp
  unfolding Grp_def rel_fun_def by auto

theorem Abs_Σ0_natural: "Σ0_map f o Abs_Σ0 = Abs_Σ0 o F_map f"
  using Abs_Σ0_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding Σ0.rel_Grp F_rel_Grp
  unfolding Grp_def rel_fun_def by auto

lemma Rep_Σ0_o_Abs_Σ0: "Rep_Σ0 o Abs_Σ0 = id"
  apply (rule ext)
  apply (rule box_equals[OF _ o_apply[symmetric] id_apply[symmetric]])
  apply (rule Abs_Σ0_inverse[OF UNIV_I])
  done

lemma Σ0_rel_Σ0_map_Σ0_map:
  "Σ0_rel R (Σ0_map f x) (Σ0_map g y) = Σ0_rel (BNF_Def.vimage2p f g R) x y"
  unfolding Σ0.rel_Grp vimage2p_Grp Σ0.rel_compp Σ0.rel_conversep
  unfolding relcompp.simps Grp_def by blast

(* ΣΣ0 is a copy of FreeAlg, but imports Behavior_BNF instead of Signature_BNF,
menaing that it uses a specific Σ0 which is a copy of F, instead of S.  *)

subsection{* Definitions and basic setup *}

(* 'x ΣΣ0 consist of Σ0-terms with variables in 'x: *)
datatype_new (ΣΣ0_set: 'x) ΣΣ0 =
  \<oo>\<pp>0 "'x ΣΣ0 Σ0" | leaf0 "'x"
  for map: ΣΣ0_map rel: ΣΣ0_rel

declare ΣΣ0.ctor_fold_transfer
 [unfolded rel_pre_ΣΣ0_def id_apply BNF_Comp.id_bnf_comp_def vimage2p_def, transfer_rule]

lemma \<oo>\<pp>0_transfer[transfer_rule]: "(Σ0_rel (ΣΣ0_rel R) ===> ΣΣ0_rel R) \<oo>\<pp>0 \<oo>\<pp>0"
  by (rule rel_funI) (erule iffD2[OF ΣΣ0.rel_inject(1)])

lemma leaf0_transfer[transfer_rule]: "(R ===> ΣΣ0_rel R) leaf0 leaf0"
  by (rule rel_funI) (erule iffD2[OF ΣΣ0.rel_inject(2)])

(* The ΣΣ0 free extension, for an algebra s, of an inteprleaftion i of the
variables into the algebra carrier:  *)
abbreviation "ext0 s ≡ rec_ΣΣ0 (s o Σ0_map snd)"
lemmas ext0_\<oo>\<pp>0 = ΣΣ0.rec(1)[of "s o Σ0_map snd" for s,
  unfolded o_apply Σ0.map_comp snd_convol[unfolded convol_def]]
lemmas ext0_leaf0 = ΣΣ0.rec(2)[of "s o Σ0_map snd" for s,
  unfolded o_apply Σ0.map_comp snd_convol[unfolded convol_def]]
lemmas leaf0_inj = ΣΣ0.inject(2)
lemmas \<oo>\<pp>0_inj = ΣΣ0.inject(1)

lemma ext0_alt: "ext0 s f = ctor_fold_ΣΣ0 (case_sum s f)"
  apply (rule ΣΣ0.ctor_fold_unique)
  apply (rule ext)
  apply (rename_tac x)
  apply (case_tac x)
  apply (auto simp only: rec_ΣΣ0_def ΣΣ0.ctor_rec fun_eq_iff o_apply BNF_Comp.id_bnf_comp_def
    id_def[symmetric] o_id map_pre_ΣΣ0_def id_apply map_sum.simps sum.inject sum.distinct
    Σ0.map_comp snd_convol split: sum.splits)
  done

(* The term algebra: *)

lemma \<oo>\<pp>0_def_pointfree: "\<oo>\<pp>0 = ctor_ΣΣ0 o Inl"
unfolding \<oo>\<pp>0_def BNF_Comp.id_bnf_comp_def comp_def ..

(* Monadic structure *)
lemma leaf0_def_pointfree: "leaf0 = ctor_ΣΣ0 o Inr"
unfolding leaf0_def BNF_Comp.id_bnf_comp_def comp_def ..

definition flat0 :: "('x ΣΣ0) ΣΣ0 => 'x ΣΣ0" where
  flat0_def: "flat0 ≡ ext0 \<oo>\<pp>0 id"

lemma flat0_transfer[transfer_rule]: "(ΣΣ0_rel (ΣΣ0_rel R) ===> ΣΣ0_rel R) flat0 flat0"
  unfolding flat0_def ext0_alt by transfer_prover

(* facts about ΣΣ0: *)
lemma ctor_fold_ΣΣ0_pointfree:
"ctor_fold_ΣΣ0 s o ctor_ΣΣ0 = s o (map_pre_ΣΣ0 id (ctor_fold_ΣΣ0 s))"
unfolding comp_def ΣΣ0.ctor_fold ..

lemma ΣΣ0_map_ctor_ΣΣ0:
"ΣΣ0_map f o ctor_ΣΣ0 = ctor_ΣΣ0 o map_sum (Σ0_map (ΣΣ0_map f)) f"
unfolding comp_def fun_eq_iff ΣΣ0.ctor_map map_pre_ΣΣ0_def BNF_Comp.id_bnf_comp_def id_apply by simp

lemma dtor_ΣΣ0_ΣΣ0_map:
"dtor_ΣΣ0 o ΣΣ0_map f = map_sum (Σ0_map (ΣΣ0_map f)) f o dtor_ΣΣ0"
using ΣΣ0_map_ctor_ΣΣ0[of f] ΣΣ0.dtor_ctor ΣΣ0.ctor_dtor unfolding comp_def fun_eq_iff by metis

lemma dtor_ΣΣ0_ctor_ΣΣ0: "dtor_ΣΣ0 o ctor_ΣΣ0 = id"
unfolding comp_def ΣΣ0.dtor_ctor id_def ..

lemma ctor_ΣΣ0_dtor_ΣΣ0: "ctor_ΣΣ0 o dtor_ΣΣ0 = id"
unfolding comp_def ΣΣ0.ctor_dtor id_def ..

lemma ΣΣ0_rel_inf: "ΣΣ0_rel (R \<sqinter> S) ≤ ΣΣ0_rel R \<sqinter> ΣΣ0_rel S"
  apply (rule inf_greatest)
  apply (rule ΣΣ0.rel_mono[OF inf_sup_ord(1)])
  apply (rule ΣΣ0.rel_mono[OF inf_sup_ord(2)])
  done

lemma ΣΣ0_rel_Grp_ΣΣ0_map: "ΣΣ0_rel (BNF_Def.Grp UNIV f) x y <-> ΣΣ0_map f x = y"
  unfolding ΣΣ0.rel_Grp by (auto simp: Grp_def)

lemma ΣΣ0_rel_ΣΣ0_map_ΣΣ0_map: "ΣΣ0_rel R (ΣΣ0_map f x) (ΣΣ0_map g y) =
  ΣΣ0_rel (BNF_Def.vimage2p f g R) x y"
  unfolding ΣΣ0.rel_Grp vimage2p_Grp apply (auto simp: ΣΣ0.rel_compp ΣΣ0.rel_conversep relcompp.simps)
  apply (intro exI conjI)
  apply (rule iffD2[OF ΣΣ0_rel_Grp_ΣΣ0_map refl])
  apply assumption
  apply (rule iffD2[OF ΣΣ0_rel_Grp_ΣΣ0_map refl])
  unfolding ΣΣ0_rel_Grp_ΣΣ0_map
  apply simp
  done

subsection{* @{term Σ0} extension theorems *}

theorem ext0_commute:
"ext0 s i o \<oo>\<pp>0 = s o Σ0_map (ext0 s i)"
unfolding ext0_alt \<oo>\<pp>0_def_pointfree o_assoc ctor_fold_ΣΣ0_pointfree map_pre_ΣΣ0_def
  case_sum_o_map_sum case_sum_o_inj(1) BNF_Comp.id_bnf_comp_def id_def[symmetric] o_id ..

theorem ext0_comp_leaf0:
"ext0 s i o leaf0 = i"
unfolding ext0_alt leaf0_def_pointfree o_assoc ctor_fold_ΣΣ0_pointfree map_pre_ΣΣ0_def
  case_sum_o_map_sum case_sum_o_inj(2) id_apply BNF_Comp.id_bnf_comp_def id_def[symmetric] o_id ..

theorem ext0_unique:
assumes leaf0: "f o leaf0 = i" and com: "f o \<oo>\<pp>0 = s o Σ0_map f"
shows "f = ext0 s i"
unfolding ext0_alt
apply (rule ΣΣ0.ctor_fold_unique)
apply (rule sum_comp_cases)
unfolding map_pre_ΣΣ0_def case_sum_o_map_sum id_apply o_id case_sum_o_inj
  leaf0[unfolded leaf0_def_pointfree o_assoc] com[unfolded \<oo>\<pp>0_def_pointfree o_assoc]
  BNF_Comp.id_bnf_comp_def id_def[symmetric] id_o
by (rule refl)+


(* We have shown that the functions
ext0 s and λ f. f o leaf0 form flattually inverse bijections
between functions i :: 'x => 'a (where 'a is the carrier of the algebra s)
and algebra morphisms f : ('x ΣΣ0, \<oo>\<pp>0) -> ('a, s)
*)


subsection{* Customizing @{term ΣΣ0} *}

subsection{* Injectiveness, naturality, adjunction *}

theorem leaf0_natural: "ΣΣ0_map f o leaf0 = leaf0 o f"
  using leaf0_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding ΣΣ0.rel_Grp
  unfolding Grp_def rel_fun_def by auto

lemma \<oo>\<pp>0_natural: "\<oo>\<pp>0 o Σ0_map (ΣΣ0_map f) = ΣΣ0_map f o \<oo>\<pp>0"
  using \<oo>\<pp>0_transfer[of "BNF_Def.Grp UNIV f"]
  unfolding Σ0.rel_Grp ΣΣ0.rel_Grp
  unfolding Grp_def rel_fun_def by auto

lemma ΣΣ0_map_def2: "ΣΣ0_map i = ext0 \<oo>\<pp>0 (leaf0 o i)"
by (rule ext0_unique[OF leaf0_natural \<oo>\<pp>0_natural[symmetric]])

lemma ext0_\<oo>\<pp>0_leaf0: "ext0 \<oo>\<pp>0 leaf0 = id"
apply (rule ext0_unique[symmetric]) unfolding Σ0.map_id0 o_id id_o by (rule refl)+

lemma ext0_ΣΣ0_map:
"ext0 s (j o f) = ext0 s j o ΣΣ0_map f"
proof (rule ext0_unique[symmetric])
  show "ext0 s j o ΣΣ0_map f o leaf0 = j o f"
  unfolding o_assoc[symmetric] leaf0_natural
  unfolding o_assoc ext0_comp_leaf0 ..
next
  show "ext0 s j o ΣΣ0_map f o \<oo>\<pp>0 = s o Σ0_map (ext0 s j o ΣΣ0_map f)"
  unfolding o_assoc[symmetric] \<oo>\<pp>0_natural[symmetric]
  unfolding o_assoc ext0_commute
  unfolding o_assoc[symmetric] Σ0.map_comp0 ..
qed

lemma ext0_Σ0_map:
assumes "t o Σ0_map f = f o s"
shows "ext0 t (f o i) = f o ext0 s i"
proof (rule ext0_unique[symmetric])
  show "f o ext0 s i o leaf0 = f o i"
  unfolding o_assoc[symmetric] ext0_comp_leaf0 ..
next
  show "f o ext0 s i o \<oo>\<pp>0 = t o Σ0_map (f o ext0 s i)"
  unfolding Σ0.map_comp0 o_assoc assms
  unfolding o_assoc[symmetric] ext0_commute[symmetric] ..
qed


subsection{* Monadic laws *}

lemma flat0_commute: "\<oo>\<pp>0 o Σ0_map flat0 = flat0 o \<oo>\<pp>0"
unfolding flat0_def ext0_commute ..

(* ΣΣ0 2 identity laws*)
theorem flat0_leaf0: "flat0 o leaf0 = id"
unfolding flat0_def ext0_comp_leaf0 ..

theorem leaf0_flat0: "flat0 o ΣΣ0_map leaf0 = id"
unfolding flat0_def ext0_ΣΣ0_map[symmetric] id_o ext0_\<oo>\<pp>0_leaf0 ..

theorem flat0_natural: "flat0 o ΣΣ0_map (ΣΣ0_map i) = ΣΣ0_map i o flat0"
  using flat0_transfer[of "BNF_Def.Grp UNIV i"]
  unfolding ΣΣ0.rel_Grp
  unfolding Grp_def rel_fun_def by auto

(* Associativity *)
theorem flat0_assoc: "flat0 o ΣΣ0_map flat0 = flat0 o flat0"
unfolding flat0_def unfolding ext0_ΣΣ0_map[symmetric] id_o
proof(rule ext0_unique[symmetric], unfold flat0_def[symmetric])
  show "flat0 o flat0 o leaf0 = flat0"
  unfolding o_assoc[symmetric] flat0_leaf0 by simp
next
  show "flat0 o flat0 o \<oo>\<pp>0 = \<oo>\<pp>0 o Σ0_map (flat0 o flat0)"
  unfolding flat0_def unfolding o_assoc[symmetric] ext0_commute
  unfolding flat0_def[symmetric]
  unfolding Σ0.map_comp0 o_assoc unfolding flat0_commute ..
qed

end