Library UniMath.SubstitutionSystems.MultiSorted_actegorical

this file is a follow-up of Multisorted_alt, where the semantic signatures Signature are replaced by functors with tensorial strength
the concept of binding signatures is inherited, as well as the reasoning about omega-cocontinuity the strength notion is based on lax lineators where endofunctors act on possibly non-endofunctors, but the signature functor generated from a multi-sorted binding signature falls into the special case of endofunctors, and the lineator notion can be transferred (through weak equivalence) to the strength notion of monoidal heterogeneous substitution systems (MHSS)
accordingly a MHSS is constructed and a monad obtained through it, cf. MultiSortedMonadConstruction_actegorical
author: Ralph Matthes, 2023
notice: this file does not correspond to Multisorted but to Multisorted_alt, even though this is not indicated in the name

Require Import UniMath.Foundations.PartD.
Require Import UniMath.Foundations.Sets.

Require Import UniMath.MoreFoundations.Tactics.

Require Import UniMath.Combinatorics.Lists.

Require Import UniMath.CategoryTheory.Core.Categories.
Require Import UniMath.CategoryTheory.Core.NaturalTransformations.
Require Import UniMath.CategoryTheory.Core.Functors.
Require Import UniMath.CategoryTheory.FunctorCategory.
Require Import UniMath.CategoryTheory.whiskering.
Require Import UniMath.CategoryTheory.Limits.Graphs.Colimits.
Require Import UniMath.CategoryTheory.Limits.BinProducts.
Require Import UniMath.CategoryTheory.Limits.Products.
Require Import UniMath.CategoryTheory.Limits.BinCoproducts.
Require Import UniMath.CategoryTheory.Limits.Coproducts.
Require Import UniMath.CategoryTheory.Limits.Terminal.
Require Import UniMath.CategoryTheory.Limits.Initial.
Require Import UniMath.CategoryTheory.FunctorAlgebras.
Require Import UniMath.CategoryTheory.Exponentials.
Require Import UniMath.CategoryTheory.Adjunctions.Core.
Require Import UniMath.CategoryTheory.Chains.All.
Require Import UniMath.CategoryTheory.Monads.Monads.
Require Import UniMath.CategoryTheory.Categories.HSET.Core.
Require Import UniMath.CategoryTheory.Categories.HSET.Colimits.
Require Import UniMath.CategoryTheory.Categories.HSET.Limits.
Require Import UniMath.CategoryTheory.Categories.HSET.Structures.
Require Import UniMath.CategoryTheory.Categories.StandardCategories.
Require Import UniMath.CategoryTheory.Groupoids.
Require UniMath.SubstitutionSystems.SortIndexing.

Require Import UniMath.SubstitutionSystems.Signatures.
Require Import UniMath.SubstitutionSystems.SumOfSignatures.
Require Import UniMath.SubstitutionSystems.BinProductOfSignatures.
Require Import UniMath.SubstitutionSystems.SignatureExamples.

for the additions in 2023
Preamble copied from Multisorted_alt

Context (sort : UU) (Hsort : isofhlevel 3 sort) (C : category).

Context (TC : Terminal C) (IC : Initial C)
          (BP : BinProducts C) (BC : BinCoproducts C)
          (PC : forall (I : UU), Products I C) (CC : forall (I : UU), isaset I Coproducts I C).

Local Notation "'1'" := (TerminalObject TC).
Local Notation "a ⊕ b" := (BinCoproductObject (BC a b)).

This represents "sort → C"
end of Preamble copied from Multisorted_alt

Construction of the lineator for the endofunctor on C^sort,C^sort

derived from a multisorted signature
we are only interested in E to have value either sortToC or C

  Local Definition pointedstrengthfromselfaction_CAT :=
    lineator_lax Mon_ptdendo_CAT ActPtd_CAT_FromSelf ActPtd_CAT_FromSelf.

  Let pointedlaxcommutator_CAT (G : sortToC2) : UU :=
        BindingSigToMonad_actegorical.pointedlaxcommutator_CAT G.

  Local Definition δCCCATEndo (M : MultiSortedSig sort) :
    actegory_coprod_distributor Mon_ptdendo_CAT (CoproductsMultiSortedSig M) ActPtd_CAT_Endo.
  Show proof.

  Local Definition δCCCATfromSelf (M : MultiSortedSig sort) :
    actegory_coprod_distributor Mon_ptdendo_CAT (CoproductsMultiSortedSig M) ActPtd_CAT_FromSelf.
  Show proof.

  Definition ptdlaxcommutatorCAT_option_functor (s : sort) :
    pointedlaxcommutator_CAT (sorted_option_functor s).
  Show proof.

  Definition ptdlaxcommutatorCAT_option_list (xs : list sort) :
    pointedlaxcommutator_CAT (option_list xs).
  Show proof.
    induction xs as [[|n] xs].
    + induction xs.
      apply unit_relativelaxcommutator.
    + induction n as [|n IH].
      * induction xs as [m []].
        apply ptdlaxcommutatorCAT_option_functor.
      * induction xs as [m [k xs]].
        use composedrelativelaxcommutator.
        -- exact (ptdlaxcommutatorCAT_option_functor m).
        -- exact (IH (k,,xs)).

  Definition StrengthCAT_exp_functor (lt : list sort × sort) :
    pointedstrengthfromprecomp_CAT C (exp_functor lt).
  Show proof.
    induction lt as [l t]; revert l.
    use list_ind.
    - cbn.       use reindexed_lax_lineator.
      exact (lax_lineator_postcomp_actegories_from_precomp_CAT _ _ _ (projSortToC t)).
    - intros x xs H; simpl.
      use comp_lineator_lax.
      3: { use reindexed_lax_lineator.
           2: { exact (lax_lineator_postcomp_actegories_from_precomp_CAT _ _ _ (projSortToC t)). }
      }
      use reindexedstrength_from_commutator.
      exact (ptdlaxcommutatorCAT_option_list (cons x xs)).

  Definition StrengthCAT_exp_functor_list (xs : list (list sort × sort)) :
    pointedstrengthfromprecomp_CAT C (exp_functor_list xs).
  Show proof.
    induction xs as [[|n] xs].
    - induction xs.
      use reindexed_lax_lineator.
      apply constconst_functor_lax_lineator.
    - induction n as [|n IH].
      + induction xs as [m []].
        exact (StrengthCAT_exp_functor m).
      + induction xs as [m [k xs]].
        apply (lax_lineator_binprod Mon_ptdendo_CAT ActPtd_CAT_Endo (ActPtd_CAT C)).
        * apply StrengthCAT_exp_functor.
        * exact (IH (k,,xs)).

  Definition StrengthCAT_hat_exp_functor_list (xst : list (list sort × sort) × sort) :
    pointedstrengthfromprecomp_CAT sortToC (hat_exp_functor_list xst).
  Show proof.
    use comp_lineator_lax.
    - exact (ActPtd_CAT C).
    - apply StrengthCAT_exp_functor_list.
    - use reindexed_lax_lineator.
      apply lax_lineator_postcomp_actegories_from_precomp_CAT.

  Definition MultiSortedSigToStrengthCAT (M : MultiSortedSig sort) :
    pointedstrengthfromprecomp_CAT sortToC (MultiSortedSigToFunctor M).
  Show proof.



we now adapt the definitions to MultiSortedSigToFunctor'

  Local Definition MultiSortedSigToFunctor' : MultiSortedSig sort -> sortToC3 := MultiSortedSigToFunctor' sort Hsort C TC BP BC CC.
  Local Definition hat_exp_functor_list'_optimized : list (list sort × sort) × sort -> sortToC3
    := hat_exp_functor_list'_optimized sort Hsort C TC BP BC CC.
  Local Definition hat_exp_functor_list'_piece : (list sort × sort) × sort -> sortToC3
    := hat_exp_functor_list'_piece sort Hsort C TC BC CC.

  Definition StrengthCAT_hat_exp_functor_list'_piece (xt : (list sort × sort) × sort) :
    pointedstrengthfromselfaction_CAT (hat_exp_functor_list'_piece xt).
  Show proof.

  Definition StrengthCAT_hat_exp_functor_list'_optimized (xst : list (list sort × sort) × sort) :
    pointedstrengthfromselfaction_CAT (hat_exp_functor_list'_optimized xst).
  Show proof.
    induction xst as [xs t].
    induction xs as [[|n] xs].
    - induction xs.
      use reindexed_lax_lineator.
      use comp_lineator_lax.       2: { apply constconst_functor_lax_lineator. }
      apply lax_lineator_postcomp_SelfActCAT_alt.
    - induction n as [|n IH].
      + induction xs as [m []].
        apply StrengthCAT_hat_exp_functor_list'_piece.
      + induction xs as [m [k xs]].
        apply (lax_lineator_binprod Mon_ptdendo_CAT ActPtd_CAT_FromSelf ActPtd_CAT_FromSelf).
        * apply StrengthCAT_hat_exp_functor_list'_piece.
        * exact (IH (k,,xs)).

  Definition MultiSortedSigToStrength' (M : MultiSortedSig sort) :
    pointedstrengthfromselfaction_CAT (MultiSortedSigToFunctor' M).
  Show proof.

End strength_through_actegories.

End MBindingSig.