header{* Free algebras for an BNF *}
theory FreeAlg_step
imports Input_step
begin
declare K_step.map_transfer[transfer_rule]
composition_bnf (open) raw_Σ_step: "'a Σ_base + 'a K_step"
typedef 'a Σ_step = "UNIV :: ('a Σ_base + 'a K_step) set" by (rule UNIV_witness)
setup_lifting type_definition_Σ_step
lift_definition Σ_step_map :: "('a => 'b) => 'a Σ_step => 'b Σ_step" is "λf. map_sum (Σ_base_map f) (K_step_map f)" .
lift_definition Σ_step_set :: "'a Σ_step => 'a set"
is "λx. UNION (Basic_BNFs.setl x) Σ_base_set ∪ UNION (Basic_BNFs.setr x) K_step_set" .
lift_definition Σ_step_rel :: "('a => 'b => bool) => 'a Σ_step => 'b Σ_step => bool"
is "λR. rel_sum (Σ_base_rel R) (K_step_rel R)" .
typedef Σ_step_bd_type = "UNIV :: ((Σ_base_bd_type + bd_type_K_step) × nat) set" by (rule UNIV_witness)
definition "Σ_step_bd = dir_image ((Σ_base_bd +c bd_K_step) *c natLeq) Abs_Σ_step_bd_type"
bnf "'a Σ_step"
map: Σ_step_map
sets: Σ_step_set
bd: Σ_step_bd
rel: Σ_step_rel
unfolding Σ_step_bd_def
apply -
apply transfer apply (rule raw_Σ_step.map_id0)
apply transfer apply (rule raw_Σ_step.map_comp0)
apply transfer apply (erule raw_Σ_step.map_cong0)
apply transfer apply (rule raw_Σ_step.set_map0)
apply (rule card_order_dir_image[OF bijI raw_Σ_step.bd_card_order])
apply (metis inj_on_def Abs_Σ_step_bd_type_inverse[OF UNIV_I])
apply (metis surj_def Abs_Σ_step_bd_type_cases)
apply (rule conjunct1[OF Cinfinite_cong[OF dir_image[OF _ raw_Σ_step.bd_Card_order] raw_Σ_step.bd_Cinfinite]])
apply (metis Abs_Σ_step_bd_type_inverse[OF UNIV_I])
apply (unfold Σ_step_set_def map_fun_def id_o) [1] apply (subst o_apply)
apply (rule ordLeq_ordIso_trans[OF raw_Σ_step.set_bd dir_image[OF _ raw_Σ_step.bd_Card_order]])
apply (metis Abs_Σ_step_bd_type_inverse[OF UNIV_I])
apply (rule predicate2I) apply transfer apply (subst raw_Σ_step.rel_compp) apply assumption
apply transfer' apply (rule raw_Σ_step.rel_compp_Grp)
done
declare Σ_step.map_transfer[transfer_rule]
lemma Rep_Σ_step_transfer[transfer_rule]: "(Σ_step_rel R ===> rel_sum (Σ_base_rel R) (K_step_rel R)) Rep_Σ_step Rep_Σ_step"
unfolding rel_fun_def by transfer blast
lemma Abs_Σ_step_transfer[transfer_rule]: "(rel_sum (Σ_base_rel R) (K_step_rel R) ===> Σ_step_rel R) Abs_Σ_step Abs_Σ_step"
unfolding rel_fun_def by transfer blast
theorem Rep_Σ_step_natural: "map_sum (Σ_base_map f) (K_step_map f) o Rep_Σ_step = Rep_Σ_step o Σ_step_map f"
using Rep_Σ_step_transfer[of "BNF_Def.Grp UNIV f"]
unfolding Σ_step.rel_Grp raw_Σ_step.rel_Grp
unfolding Grp_def rel_fun_def by auto
theorem Abs_Σ_step_natural: "Σ_step_map f o Abs_Σ_step = Abs_Σ_step o map_sum (Σ_base_map f) (K_step_map f)"
using Abs_Σ_step_transfer[of "BNF_Def.Grp UNIV f"]
unfolding Σ_step.rel_Grp raw_Σ_step.rel_Grp
unfolding Grp_def rel_fun_def by auto
lemma Rep_Σ_step_o_Abs_Σ_step: "Rep_Σ_step o Abs_Σ_step = id"
apply (rule ext)
apply (rule box_equals[OF _ o_apply[symmetric] id_apply[symmetric]])
apply (rule Abs_Σ_step_inverse[OF UNIV_I])
done
lemma Σ_step_rel_Σ_step_map_Σ_step_map:
"Σ_step_rel R (Σ_step_map f x) (Σ_step_map g y) = Σ_step_rel (BNF_Def.vimage2p f g R) x y"
unfolding Σ_step.rel_Grp vimage2p_Grp Σ_step.rel_compp Σ_step.rel_conversep
unfolding relcompp.simps Grp_def by simp
subsection{* Definitions and basic setup *}
datatype_new (ΣΣ_step_set: 'x) ΣΣ_step =
\<oo>\<pp>_step "'x ΣΣ_step Σ_step" | leaf_step "'x"
for map: ΣΣ_step_map rel: ΣΣ_step_rel
declare ΣΣ_step.ctor_fold_transfer
[unfolded rel_pre_ΣΣ_step_def id_apply BNF_Comp.id_bnf_comp_def vimage2p_def, transfer_rule]
lemma \<oo>\<pp>_step_transfer[transfer_rule]:
"(Σ_step_rel (ΣΣ_step_rel R) ===> ΣΣ_step_rel R) \<oo>\<pp>_step \<oo>\<pp>_step"
by (rule rel_funI) (erule iffD2[OF ΣΣ_step.rel_inject(1)])
lemma leaf_step_transfer[transfer_rule]: "(R ===> ΣΣ_step_rel R) leaf_step leaf_step"
by (rule rel_funI) (erule iffD2[OF ΣΣ_step.rel_inject(2)])
abbreviation "ext_step s ≡ rec_ΣΣ_step (s o Σ_step_map snd)"
lemmas ext_step_\<oo>\<pp>_step = ΣΣ_step.rec(1)[of "s o Σ_step_map snd" for s,
unfolded o_apply Σ_step.map_comp snd_convol[unfolded convol_def]]
lemmas ext_step_leaf_step = ΣΣ_step.rec(2)[of "s o Σ_step_map snd" for s,
unfolded o_apply Σ_step.map_comp snd_convol[unfolded convol_def]]
lemmas leaf_step_inj = ΣΣ_step.inject(2)
lemmas \<oo>\<pp>_step_inj = ΣΣ_step.inject(1)
lemma ext_step_alt: "ext_step s f = ctor_fold_ΣΣ_step (case_sum s f)"
apply (rule ΣΣ_step.ctor_fold_unique)
apply (rule ext)
apply (rename_tac x)
apply (case_tac x)
apply (auto simp only: rec_ΣΣ_step_def ΣΣ_step.ctor_rec fun_eq_iff o_apply BNF_Comp.id_bnf_comp_def
id_def[symmetric] o_id map_pre_ΣΣ_step_def id_apply map_sum.simps sum.inject sum.distinct
Σ_step.map_comp snd_convol split: sum.splits)
done
lemma \<oo>\<pp>_step_def_pointfree: "\<oo>\<pp>_step ≡ ctor_ΣΣ_step o Inl"
unfolding \<oo>\<pp>_step_def comp_def BNF_Comp.id_bnf_comp_def .
lemma leaf_step_def_pointfree: "leaf_step ≡ ctor_ΣΣ_step o Inr"
unfolding leaf_step_def comp_def BNF_Comp.id_bnf_comp_def .
definition flat_step :: "('x ΣΣ_step) ΣΣ_step => 'x ΣΣ_step" where
flat_step_def: "flat_step ≡ ext_step \<oo>\<pp>_step id"
lemma flat_step_transfer[transfer_rule]: "(ΣΣ_step_rel (ΣΣ_step_rel R) ===> ΣΣ_step_rel R) flat_step flat_step"
unfolding flat_step_def ext_step_alt by transfer_prover
lemma ctor_fold_ΣΣ_step_pointfree:
"ctor_fold_ΣΣ_step s o ctor_ΣΣ_step = s o (map_pre_ΣΣ_step id (ctor_fold_ΣΣ_step s))"
unfolding comp_def ΣΣ_step.ctor_fold ..
lemma ΣΣ_step_map_ctor_ΣΣ_step:
"ΣΣ_step_map f o ctor_ΣΣ_step = ctor_ΣΣ_step o map_sum (Σ_step_map (ΣΣ_step_map f)) f"
unfolding comp_def
unfolding fun_eq_iff
unfolding ΣΣ_step.ctor_map
unfolding map_pre_ΣΣ_step_def
unfolding id_apply BNF_Comp.id_bnf_comp_def id_def[symmetric] o_id id_o by simp
lemma dtor_ΣΣ_step_ΣΣ_step_map:
"dtor_ΣΣ_step o ΣΣ_step_map f = map_sum (Σ_step_map (ΣΣ_step_map f)) f o dtor_ΣΣ_step"
using ΣΣ_step_map_ctor_ΣΣ_step[of f] ΣΣ_step.dtor_ctor ΣΣ_step.ctor_dtor unfolding comp_def fun_eq_iff by metis
lemma dtor_ΣΣ_step_ctor_ΣΣ_step: "dtor_ΣΣ_step o ctor_ΣΣ_step = id"
unfolding comp_def ΣΣ_step.dtor_ctor id_def ..
lemma ctor_ΣΣ_step_dtor_ΣΣ_step: "ctor_ΣΣ_step o dtor_ΣΣ_step = id"
unfolding comp_def ΣΣ_step.ctor_dtor id_def ..
lemma ΣΣ_step_rel_inf: "ΣΣ_step_rel (R \<sqinter> Σ_base) ≤ ΣΣ_step_rel R \<sqinter> ΣΣ_step_rel Σ_base"
apply (rule inf_greatest)
apply (rule ΣΣ_step.rel_mono[OF inf_sup_ord(1)])
apply (rule ΣΣ_step.rel_mono[OF inf_sup_ord(2)])
done
lemma ΣΣ_step_rel_Grp_ΣΣ_step_map: "ΣΣ_step_rel (BNF_Def.Grp UNIV f) x y <-> ΣΣ_step_map f x = y"
unfolding ΣΣ_step.rel_Grp by (auto simp: Grp_def)
lemma ΣΣ_step_rel_ΣΣ_step_map_ΣΣ_step_map: "ΣΣ_step_rel R (ΣΣ_step_map f x) (ΣΣ_step_map g y) =
ΣΣ_step_rel (BNF_Def.vimage2p f g R) x y"
unfolding ΣΣ_step.rel_Grp vimage2p_Grp apply (auto simp: ΣΣ_step.rel_compp ΣΣ_step.rel_conversep relcompp.simps)
apply (intro exI conjI)
apply (rule iffD2[OF ΣΣ_step_rel_Grp_ΣΣ_step_map refl])
apply assumption
apply (rule iffD2[OF ΣΣ_step_rel_Grp_ΣΣ_step_map refl])
unfolding ΣΣ_step_rel_Grp_ΣΣ_step_map
apply simp
done
subsection{* @{term Σ_step} extension theorems *}
theorem ext_step_commute:
"ext_step s i o \<oo>\<pp>_step = s o Σ_step_map (ext_step s i)"
unfolding ext_step_alt \<oo>\<pp>_step_def_pointfree o_assoc ctor_fold_ΣΣ_step_pointfree map_pre_ΣΣ_step_def
case_sum_o_map_sum case_sum_o_inj(1) BNF_Comp.id_bnf_comp_def id_def[symmetric] o_id ..
theorem ext_step_comp_leaf_step:
"ext_step s i o leaf_step = i"
unfolding ext_step_alt leaf_step_def_pointfree o_assoc ctor_fold_ΣΣ_step_pointfree map_pre_ΣΣ_step_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 ext_step_unique:
assumes leaf_step: "f o leaf_step = i" and com: "f o \<oo>\<pp>_step = s o Σ_step_map f"
shows "f = ext_step s i"
unfolding ext_step_alt
apply (rule ΣΣ_step.ctor_fold_unique)
apply (rule sum_comp_cases)
unfolding map_pre_ΣΣ_step_def case_sum_o_map_sum id_apply o_id case_sum_o_inj
leaf_step[unfolded leaf_step_def_pointfree o_assoc] com[unfolded \<oo>\<pp>_step_def_pointfree o_assoc]
BNF_Comp.id_bnf_comp_def id_def[symmetric] id_o
by (rule refl)+
subsection{* Customizing @{term ΣΣ_step} *}
subsection{* Injectiveness, naturality, adjunction *}
theorem leaf_step_natural: "ΣΣ_step_map f o leaf_step = leaf_step o f"
using leaf_step_transfer[of "BNF_Def.Grp UNIV f"]
unfolding ΣΣ_step.rel_Grp
unfolding Grp_def rel_fun_def by auto
lemma \<oo>\<pp>_step_natural: "\<oo>\<pp>_step o Σ_step_map (ΣΣ_step_map f) = ΣΣ_step_map f o \<oo>\<pp>_step"
using \<oo>\<pp>_step_transfer[of "BNF_Def.Grp UNIV f"]
unfolding Σ_step.rel_Grp ΣΣ_step.rel_Grp Σ_step_map_def
unfolding Grp_def rel_fun_def by auto
lemma ΣΣ_step_map_def2: "ΣΣ_step_map i = ext_step \<oo>\<pp>_step (leaf_step o i)"
by (rule ext_step_unique[OF leaf_step_natural \<oo>\<pp>_step_natural[symmetric]])
lemma ext_step_\<oo>\<pp>_step_leaf_step: "ext_step \<oo>\<pp>_step leaf_step = id"
apply (rule ext_step_unique[symmetric]) unfolding Σ_step.map_id0 o_id id_o by (rule refl)+
lemma ext_step_ΣΣ_step_map:
"ext_step s (j o f) = ext_step s j o ΣΣ_step_map f"
proof (rule ext_step_unique[symmetric])
show "ext_step s j o ΣΣ_step_map f o leaf_step = j o f"
unfolding o_assoc[symmetric] leaf_step_natural
unfolding o_assoc ext_step_comp_leaf_step ..
next
show "ext_step s j o ΣΣ_step_map f o \<oo>\<pp>_step = s o Σ_step_map (ext_step s j o ΣΣ_step_map f)"
unfolding o_assoc[symmetric] \<oo>\<pp>_step_natural[symmetric]
unfolding o_assoc ext_step_commute
unfolding o_assoc[symmetric] Σ_step.map_comp0 ..
qed
lemma ext_step_Σ_step_map:
assumes "t o Σ_step_map f = f o s"
shows "ext_step t (f o i) = f o ext_step s i"
proof (rule ext_step_unique[symmetric])
show "f o ext_step s i o leaf_step = f o i"
unfolding o_assoc[symmetric] ext_step_comp_leaf_step ..
next
show "f o ext_step s i o \<oo>\<pp>_step = t o Σ_step_map (f o ext_step s i)"
unfolding Σ_step.map_comp0 o_assoc assms
unfolding o_assoc[symmetric] ext_step_commute[symmetric] ..
qed
subsection{* Monadic laws *}
lemma flat_step_commute: "\<oo>\<pp>_step o Σ_step_map flat_step = flat_step o \<oo>\<pp>_step"
unfolding flat_step_def ext_step_commute ..
theorem flat_step_leaf_step: "flat_step o leaf_step = id"
unfolding flat_step_def ext_step_comp_leaf_step ..
theorem leaf_step_flat_step: "flat_step o ΣΣ_step_map leaf_step = id"
unfolding flat_step_def ext_step_ΣΣ_step_map[symmetric] id_o ext_step_\<oo>\<pp>_step_leaf_step ..
theorem flat_step_natural: "flat_step o ΣΣ_step_map (ΣΣ_step_map i) = ΣΣ_step_map i o flat_step"
using flat_step_transfer[of "BNF_Def.Grp UNIV i"]
unfolding prod.rel_Grp ΣΣ_step.rel_Grp
unfolding Grp_def rel_fun_def by auto
theorem flat_step_assoc: "flat_step o ΣΣ_step_map flat_step = flat_step o flat_step"
unfolding flat_step_def unfolding ext_step_ΣΣ_step_map[symmetric] id_o
proof(rule ext_step_unique[symmetric], unfold flat_step_def[symmetric])
show "flat_step o flat_step o leaf_step = flat_step"
unfolding o_assoc[symmetric] flat_step_leaf_step o_id ..
next
show "flat_step o flat_step o \<oo>\<pp>_step = \<oo>\<pp>_step o Σ_step_map (flat_step o flat_step)"
unfolding flat_step_def unfolding o_assoc[symmetric] ext_step_commute
unfolding flat_step_def[symmetric]
unfolding Σ_step.map_comp0 o_assoc unfolding flat_step_commute ..
qed
definition K_step_as_ΣΣ_step :: "'a K_step => 'a ΣΣ_step" where
"K_step_as_ΣΣ_step ≡ \<oo>\<pp>_step o Σ_step_map leaf_step o Abs_Σ_step o Inr"
lemma K_step_as_ΣΣ_step_transfer[transfer_rule]:
"(K_step_rel R ===> ΣΣ_step_rel R) K_step_as_ΣΣ_step K_step_as_ΣΣ_step"
unfolding K_step_as_ΣΣ_step_def by transfer_prover
lemma K_step_as_ΣΣ_step_natural:
"K_step_as_ΣΣ_step o K_step_map f = ΣΣ_step_map f o K_step_as_ΣΣ_step"
using K_step_as_ΣΣ_step_transfer[of "BNF_Def.Grp UNIV f"]
unfolding K_step.rel_Grp ΣΣ_step.rel_Grp
unfolding Grp_def rel_fun_def by auto
end