Skip to content

Conversation

tillrampe
Copy link
Contributor

This pull request constructs the bicategory of monads and bimodules (cf. Shulman - Framed Bicategories and Monoidal Fibrations, Ch. 11). This is the result that the previous pull requests have been leading up to.

tillrampe added 18 commits May 30, 2025 20:35
Bring the changes to cat-of-bimods made during the revision process to bicat-of-bimods
…Construction.Bimodules.Properties.agda as they have been renamed
@tillrampe
Copy link
Contributor Author

As of now, this pull request type-checks, but it will need revision. Among the things to improve are:

  • Don't import in where-clauses.
  • Use using-commands wherever feasible.
  • Don't open multiple instances of a module. Use projections instead. (I want to rename many variables in Categories.Bicategory.Construction.Bimodules.Tensorproduct and refer to the involved bimodules as L and R for 'left' and 'right'.)
  • Split Categories.Bicategory.Construction.Bimodules.Tensorproduct into two files: TensorproductOfBimodules.agda and TensorproductOfHomomorphisms.agda
  • Use Categories.Bicategory.Monad.Bimodule.bimodHelper to define the tensorproduct of bimodules.
  • eta-reduce where possible

These should be the major style requests made during the previous revision processes. Let me know if I forgot any.
I want to continue writing the tensorproduct of bimodules in function composition order. Although unusual in the algebra literature, it is more import to follow the order of composition in the bicategory 𝒞.

Copy link
Collaborator

@JacquesCarette JacquesCarette left a comment

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Right about the changes that need done. I suggest a few more.

Overall: heroic work! Very impressive.

@@ -0,0 +1,82 @@
{-# OPTIONS --without-K --safe --lossy-unification #-}
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

revisit --lossy-unification (just in case it is no longer needed)


open import Categories.Bicategory.Monad.Bimodule
import Categories.Category.Construction.Bimodules
open Categories.Category.Construction.Bimodules {o} {ℓ} {e} {t} {𝒞} renaming (Bimodules to Bimodules₁)
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Renaming to a name with a subscript is rarely clear. Please try to rename to something more conceptual.

Copy link
Contributor Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

The 1 in Bimodules₁ stands for 1-category. Isn't that conceptual already? I could rename it to 1CatOfBimodules or similar, but I'm not sure that improves readability. Further, the file Categories.Bicategory.Construction.Span.agda follows the same naming convention.

Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Sometimes things slip through, and I don't notice that it's a sub-optimal convention.

What about 1-Bimodules, which is closer to what people in category theory do?

open import Categories.Bicategory
open import Categories.Bicategory.LocalCoequalizers

module Categories.Bicategory.Construction.Bimodules {o ℓ e t} {𝒞 : Bicategory o ℓ e t} {localCoeq : LocalCoequalizers 𝒞} where
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Surprised that the last argument is implicit. If that's actually fine, that's probably worth a comment!


Bimodules : Bicategory (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e (o ⊔ ℓ ⊔ e ⊔ t)
Bimodules = record
{ enriched = record
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I think it would make sense to pull the 'enriching base' out as its own definition. It might even belong elsewhere?

sq₁ : CommutativeSquare (F₂∘₁T₂▷actionˡ₁) ((act-to-the-left) ◁ T₁) (act-to-the-left) (F₂▷actionˡ₁)
sq₁ = begin
act-to-the-left ∘ᵥ F₂∘₁T₂▷actionˡ₁ ≈⟨ refl⟩∘⟨ sym-assoc₂ ⟩
F₂ ▷ actionʳ₁ ∘ᵥ (associator.from ∘ᵥ (F₂ ∘₁ T₂) ▷ actionˡ₁)
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

like in previous files, the shorthands for the associators make a big different on length.

identityˡ-triangle : F₂▷actionˡ₁ ∘ᵥ (F₂ ∘₁ F₁) ▷ η₁ ∘ᵥ unitorʳ.to ≈ id₂
identityˡ-triangle = begin
F₂▷actionˡ₁ ∘ᵥ (F₂ ∘₁ F₁) ▷ η₁ ∘ᵥ unitorʳ.to ≈⟨ assoc₂ ⟩
F₂ ▷ actionˡ₁ ∘ᵥ associator.from ∘ᵥ (F₂ ∘₁ F₁) ▷ η₁ ∘ᵥ unitorʳ.to ≈⟨ refl⟩∘⟨ sym-assoc₂ ⟩
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

line up the justifications


identityˡ∘arr : (actionˡ ∘ᵥ F ▷ η₁ ∘ᵥ unitorʳ.to) ∘ᵥ Coequalizer.arr F₂⊗F₁ ≈ id₂ ∘ᵥ Coequalizer.arr F₂⊗F₁
identityˡ∘arr = begin
(actionˡ ∘ᵥ F ▷ η₁ ∘ᵥ unitorʳ.to) ∘ᵥ Coequalizer.arr F₂⊗F₁ ≈⟨ assoc₂ ⟩
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

looks like there is a lot of 'refocusing' steps that can often be dealt with with push/pull. Definitely worth trying, so that the 'heart' of the proof becomes easier to see.

open Bimodule B₁ using () renaming (actionˡ to actionˡ₁)

abstract
linearˡ∘arr∘arr : ((Bimodule.actionˡ (B₃ ⊗₀ B₂ ⊗₀ B₁)
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Is the 'shape' of this lemma forced? It's always odd when the first N steps (and last K steps) are all assoc-shuffling.

open TensorproductOfHomomorphisms (f₃ ⊗₁ f₂) f₁ renaming (αSq to αSq[f₃⊗f₂]⊗f₁)
open TensorproductOfHomomorphisms f₃ (f₂ ⊗₁ f₁) renaming (αSq to αSqf₃⊗[f₂⊗f₁])

α⇒⊗-natural∘arr : (α⇒⊗ {B₃'} {B₂'} {B₁'}
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

consider formatting like

foo :
  term1
  ≈
  term2

and similarly for equalities, i.e. don't indent by 20+ but instead by 2. And for this stuff, a width limit of 120 is probably fine (instead of just 80).

∘ᵥ Coequalizer.arr [[F₄⊗F₃]⊗F₂]⊗F₁
∘ᵥ Coequalizer.arr [F₄⊗F₃]⊗F₂ ◁ F₁
∘ᵥ Coequalizer.arr F₄⊗F₃ ◁ F₂ ◁ F₁
≈⟨ sym-assoc₂ ⟩
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

have you experimented with indenting the justification lines a lot more? Since these are all quite short, you could create a 2-column view (but keep the first column as you have it.)

Copy link
Contributor Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I've just tried it. It looks decent.

@tillrampe
Copy link
Contributor Author

I am considering changing the directional indicators ˡ and ʳ as in actionʳ and actionˡ from diagrammatic order to algebraic order. Algebraic order seems to be the common choice in this library. However, the definition of monads stands out as having diagrammatic ordering in the naming convention (see indentityˡ and identityʳ). What do you advise?

@JacquesCarette
Copy link
Collaborator

We were too busy just creating the library to pay (enough) attention to being systematic about right/left, algebraic vs diagrammatic order, etc. (Same happened with stdlib and we're suffering because of it now.)

So I wouldn't worry too much about you being consistent with the library, since there is no explicit design decision on that front (yet!).

I think a lot of the choices in the library were non-choices, in the sense that they followed older / naive categorical writing. And that often used algebraic order. Whereas I think that more modern approaches are shifting to diagrammatic order!

If you're on the category theory Zulip, that might be an interesting question to ask there.

…results about the tensorproduct but is now redundant
…ble results about the tensorproduct of homomoprhisms but is now redundant
…rename 1-cell variables A, B, C, D to X, Y, Z, W
…s from the construction in TensorproductOfBimodules
@tillrampe
Copy link
Contributor Author

I will be rather busy until beginning of October. So, I might not be able to make changes to this pull request. I will finish this, however, as soon as my schedule allows.

@JacquesCarette
Copy link
Collaborator

Thanks for letting me know. Would it be worth my re-reviewing this in the mean time?

@tillrampe
Copy link
Contributor Author

Don't think so. I've mainly executed the changes we agreed were necessary before; like using push, pull, etc. from Morphism.Reasoning to shorten proofs, and not renaming fields of Bimodule and Monad instances.

Sign up for free to join this conversation on GitHub. Already have an account? Sign in to comment
Labels
None yet
Projects
None yet
Development

Successfully merging this pull request may close these issues.

2 participants