Documentation

Mathlib.Condensed.Discrete.Colimit

The condensed set given by left Kan extension from FintypeCat to Profinite. #

This file provides the necessary API to prove that a condensed set X is discrete if and only if for every profinite set S = limᵢSᵢ, X(S) ≅ colimᵢX(Sᵢ), and the analogous result for light condensed sets.

@[reducible, inline]

The presheaf on Profinite of locally constant functions to X.

Equations
Instances For

    The functor locallyConstantPresheaf takes cofiltered limits of finite sets with surjective projection maps to colimits.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem Condensed.isColimitLocallyConstantPresheafDiagram_desc_apply (X : Type (u + 1)) (S : Profinite) (s : CategoryTheory.Limits.Cocone (S.diagram.op.comp (Condensed.locallyConstantPresheaf X))) (i : DiscreteQuotient S.toTop) (f : LocallyConstant (↑(S.diagram.obj i).toTop) X) :
      (Condensed.isColimitLocallyConstantPresheafDiagram X S).desc s (LocallyConstant.comap (S.asLimitCone.app i) f) = s.app (Opposite.op i) f
      @[reducible, inline]

      Given a presheaf F on Profinite, lanPresheaf F is the left Kan extension of its restriction to finite sets along the inclusion functor of finite sets into Profinite.

      Equations
      Instances For

        To presheaves on Profinite whose restrictions to finite sets are isomorphic have isomorphic left Kan extensions.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For

          A presheaf, which takes a profinite set written as a cofiltered limit to the corresponding colimit, agrees with the left Kan extension of its restriction.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For

            lanPresheaf (locallyConstantPresheaf X) is a sheaf for the coherent topology on Profinite.

            Equations
            Instances For

              lanPresheaf (locallyConstantPresheaf X) as a condensed set.

              Equations
              Instances For

                The functor which takes a finite set to the set of maps into F(*) for a presheaf F on Profinite.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[simp]
                  theorem Condensed.finYoneda_map (F : CategoryTheory.Functor Profiniteᵒᵖ (Type (u + 1))) :
                  ∀ {X Y : FintypeCatᵒᵖ} (f : X Y) (g : (Opposite.unop X)F.obj (FintypeCat.toProfinite.op.obj (Opposite.op (FintypeCat.of PUnit.{u + 1} )))) (a : (Opposite.unop Y)), (Condensed.finYoneda F).map f g a = (g f.unop) a

                  locallyConstantPresheaf restricted to finite sets is isomorphic to finYoneda F.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For

                    A finite set as a coproduct cocone in Profinite over itself.

                    Equations
                    Instances For

                      A finite set is the coproduct of its points in Profinite.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For

                        Auxiliary definition for isoFinYoneda.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For

                          The restriction of a finite product preserving presheaf F on Profinite to the category of finite sets is isomorphic to finYoneda F.

                          Equations
                          Instances For

                            A presheaf F, which takes a profinite set written as a cofiltered limit to the corresponding colimit, is isomorphic to the presheaf LocallyConstant - F(*).

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For
                              @[reducible, inline]

                              The presheaf on LightProfinite of locally constant functions to X.

                              Equations
                              Instances For

                                The functor locallyConstantPresheaf takes sequential limits of finite sets with surjective projection maps to colimits.

                                Equations
                                • One or more equations did not get rendered due to their size.
                                Instances For

                                  isColimitLocallyConstantPresheaf in the case of S.asLimit.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    @[reducible, inline]

                                    Given a presheaf F on LightProfinite, lanPresheaf F is the left Kan extension of its restriction to finite sets along the inclusion functor of finite sets into Profinite.

                                    Equations
                                    Instances For

                                      To presheaves on LightProfinite whose restrictions to finite sets are isomorphic have isomorphic left Kan extensions.

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For

                                        A presheaf, which takes a light profinite set written as a sequential limit to the corresponding colimit, agrees with the left Kan extension of its restriction.

                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        Instances For

                                          lanPresheaf (locallyConstantPresheaf X) as a light condensed set.

                                          Equations
                                          Instances For

                                            The functor which takes a finite set to the set of maps into F(*) for a presheaf F on LightProfinite.

                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For

                                              locallyConstantPresheaf restricted to finite sets is isomorphic to finYoneda F.

                                              Equations
                                              • One or more equations did not get rendered due to their size.
                                              Instances For

                                                A finite set as a coproduct cocone in LightProfinite over itself.

                                                Equations
                                                Instances For

                                                  A finite set is the coproduct of its points in LightProfinite.

                                                  Equations
                                                  • One or more equations did not get rendered due to their size.
                                                  Instances For

                                                    Auxiliary definition for isoFinYoneda.

                                                    Equations
                                                    • One or more equations did not get rendered due to their size.
                                                    Instances For

                                                      The restriction of a finite product preserving presheaf F on Profinite to the category of finite sets is isomorphic to finYoneda F.

                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For

                                                        A presheaf F, which takes a light profinite set written as a sequential limit to the corresponding colimit, is isomorphic to the presheaf LocallyConstant - F(*).

                                                        Equations
                                                        • One or more equations did not get rendered due to their size.
                                                        Instances For