Zulip Chat Archive
Stream: maths
Topic: semiorthogonal decompositions
Matthew Ballard (Dec 23 2021 at 21:07):
variable [pretriangulated T]
def cone {X Y: T} (f : X ⟶ Y) : T :=
Exists.some(distinguished_cocone_triangle X Y f)
structure triang_subcategory :=
(P : T → Prop)
(cone_mem : ∀ {X Y : T} (f : X ⟶ Y), P X → P Y → P (cone T f))
(shift_mem: ∀ (X : T), P X → P (X ⟦1⟧))
structure semiorthogonal_decomposition :=
(A : triang_subcategory T)
(B : triang_subcategory T)
(semiorthogonal : ∀ {X Y: T} (f : X ⟶ Y), A.P Y → B.P X → f = 0)
(triangle : ∀ (X : T), ∃ (XA XB : T) (f : XB ⟶ X) (g : X ⟶ XA) (h : XA ⟶ (XB)⟦1⟧), (A.P XA) ∧ (B.P XB) ∧ triangle.mk _ f g h ∈ dist_triang T)
A semi-orthogonal decomposition of triangulated category is a pair of triangulated subcategories such that
when and and
- for any there is a triangle
with and . A canonical reference is Bondal and Orlov's paper.
What do people think about this design choice?
Reid Barton (Dec 23 2021 at 21:31):
I assume P
is also meant to be invariant under isomorphism, and probably P 0
should hold, and maybe you want to assume something about direct sums too?
I tend to assume it's better to state things like cone_mem
in the form "for every distinguished triangle , P X -> P Y -> P Z
" rather than making some arbitrary (but useless) choice of "the" cone, and then having to invoke isomorphism invariance
Reid Barton (Dec 23 2021 at 21:35):
(e.g. compare "a pushout of a cofibration is a cofibration" https://github.com/rwbarton/lean-homotopy-theory/blob/lean-3.4.2/src/homotopy_theory/formal/cofibrations/cofibration_category.lean#L55-L56)
Matthew Ballard (Dec 23 2021 at 21:35):
Since
is a triangle, you get P 0
from P (cone T f)
.
Reid Barton said:
I assume
P
is also meant to be invariant under isomorphism,
Should it? It seems to fall in the same space as full vs strictly full subcategories.
Reid Barton said:
I tend to assume it's better to state things like
cone_mem
in the form "for every distinguished triangle ,P X -> P Y -> P Z
" rather than making some arbitrary (but useless) choice of "the" cone, and then having to invoke isomorphism invariance
Yes, this is very good advice. Thanks.
Reid Barton (Dec 23 2021 at 21:36):
Matthew Ballard said:
Since
is a triangle, you get
P 0
fromP (cone T f)
.
but P
might not hold for any object to start with!
Matthew Ballard (Dec 23 2021 at 21:36):
Ah yes :embarrassed:
Reid Barton (Dec 23 2021 at 21:37):
If P
is closed under cofibers then it is also closed under replacing objects by isomorphic ones (since is distinguished if is an iso)
Reid Barton (Dec 23 2021 at 21:38):
In other words if you rephrase cone_mem
in the style I suggested then isomorphism invariance is a consequence (once you have P 0
). If you ask for it to be closed under forming "the" cone then it's a weird concept if you did not already assume it to be isomorphism invariant.
Matthew Ballard (Dec 23 2021 at 21:38):
I agree
Matthew Ballard (Dec 23 2021 at 21:41):
About direct sums, you want to prove some results about (co)limits and left/right orthogonals. Is it valuable to package some assumptions directly?
Johan Commelin (Dec 24 2021 at 09:50):
@Matthew Ballard It will be very interesting to hear if you find the current setup of pretriangulated categories pleasant to work with. My gut feeling is that there might be too much bundling of triangles, and that this might be annoying to work with in practice.
Johan Commelin (Dec 24 2021 at 09:51):
But I haven't used it yet, so I don't have more than a gut feeling
Last updated: Dec 20 2023 at 11:08 UTC