Documentation

Mathlib.Topology.Category.Profinite.Extend

Extending cones in Profinite #

Let (Sᵢ)_{i : I} be a family of finite sets indexed by a cofiltered category I and let S be its limit in Profinite. Let G be a functor from Profinite to a category C and suppose that G preserves the limit described above. Suppose further that the projection maps S ⟶ Sᵢ are epimorphic for all i. Then G.obj S is isomorphic to a limit indexed by StructuredArrow S toProfinite (see Profinite.Extend.isLimitCone).

We also provide the dual result for a functor of the form G : Profiniteᵒᵖ ⥤ C.

We apply this to define Profinite.diagram', Profinite.asLimitCone', and Profinite.asLimit', analogues to their unprimed versions in Mathlib.Topology.Category.Profinite.AsLimit, in which the indexing category is StructuredArrow S toProfinite instead of DiscreteQuotient S.

A continuous map from a profinite set to a finite set factors through one of the components of the profinite set when written as a cofiltered limit of finite sets.

Given a cone in Profinite, consisting of finite sets and indexed by a cofiltered category, we obtain a functor from the indexing category to StructuredArrow c.pt toProfinite.

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

    Given a cone in Profinite, consisting of finite sets and indexed by a cofiltered category, we obtain a functor from the opposite of the indexing category to CostructuredArrow toProfinite.op ⟨c.pt⟩.

    Equations
    Instances For

      If the projection maps in the cone are epimorphic and the cone is limiting, then Profinite.Extend.functor is initial.

      TODO: investigate how to weaken the assumption ∀ i, Epi (c.π.app i) to ∀ i, ∃ j (_ : j ⟶ i), Epi (c.π.app j).

      If the projection maps in the cone are epimorphic and the cone is limiting, then Profinite.Extend.functorOp is final.

      Given a functor G from Profinite and S : Profinite, we obtain a cone on (StructuredArrow.proj S toProfinite ⋙ toProfinite ⋙ G) with cone point G.obj S.

      Whiskering this cone with Profinite.Extend.functor c gives G.mapCone c as we check in the example below.

      Equations
      Instances For

        If c and G.mapCone c are limit cones and the projection maps in c are epimorphic, then cone G c.pt is a limit cone.

        Equations
        Instances For

          Given a functor G from Profiniteᵒᵖ and S : Profinite, we obtain a cocone on (CostructuredArrow.proj toProfinite.op ⟨S⟩ ⋙ toProfinite.op ⋙ G) with cocone point G.obj ⟨S⟩.

          Whiskering this cocone with Profinite.Extend.functorOp c gives G.mapCocone c.op as we check in the example below.

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

            If c is a limit cone, G.mapCocone c.op is a colimit cone and the projection maps in c are epimorphic, then cocone G c.pt is a colimit cone.

            Equations
            Instances For
              @[reducible, inline]

              A functor StructuredArrow S toProfinite ⥤ FintypeCat whose limit in Profinite is isomorphic to S.

              Equations
              Instances For
                @[reducible, inline]

                An abbreviation for S.fintypeDiagram' ⋙ toProfinite.

                Equations
                Instances For
                  @[reducible, inline]

                  A cone over S.diagram' whose cone point is S.

                  Equations
                  Instances For
                    noncomputable def Profinite.asLimit' (S : Profinite) :

                    S.asLimitCone' is a limit cone.

                    Equations
                    Instances For
                      noncomputable def Profinite.lim' (S : Profinite) :

                      A bundled version of S.asLimitCone' and S.asLimit'.

                      Equations
                      • S.lim' = { cone := S.asLimitCone', isLimit := S.asLimit' }
                      Instances For