Documentation

Mathlib.CategoryTheory.Action.Concrete

Constructors for Action V G for some concrete categories #

We construct Action (Type*) G from a [MulAction G X] instance and give some applications.

@[simp]
theorem Action.ρ_inv_self_apply {G : Type u} [Group G] {A : Action (Type u) G} (g : G) (x : A.V) :
A.ρ g⁻¹ (A.ρ g x) = x
@[simp]
theorem Action.ρ_self_inv_apply {G : Type u} [Group G] {A : Action (Type u) G} (g : G) (x : A.V) :
A.ρ g (A.ρ g⁻¹ x) = x
def Action.ofMulAction (G : Type u_1) (H : Type u) [Monoid G] [MulAction G H] :

Bundles a type H with a multiplicative action of G as an Action.

Equations
Instances For
    theorem Action.ofMulAction_V (G : Type u_1) (H : Type u) [Monoid G] [MulAction G H] :
    (ofMulAction G H).V = H
    @[simp]
    theorem Action.ofMulAction_apply {G : Type u_1} {H : Type u_2} [Monoid G] [MulAction G H] (g : G) (x : H) :
    (ofMulAction G H).ρ g x = g x
    def Action.ofMulActionLimitCone {ι : Type v} (G : Type (max v u)) [Monoid G] (F : ιType (max v u)) [(i : ι) → MulAction G (F i)] :

    Given a family F of types with G-actions, this is the limit cone demonstrating that the product of F as types is a product in the category of G-sets.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[reducible, inline]
      abbrev Action.leftRegular (G : Type u) [Monoid G] :

      The G-set G, acting on itself by left multiplication.

      Equations
      Instances For
        @[reducible, inline]
        abbrev Action.diagonal (G : Type u) [Monoid G] (n : ) :

        The G-set Gⁿ, acting on itself by left multiplication.

        Equations
        Instances For

          We have Fin 1 → G ≅ G as G-sets, with G acting by left multiplication.

          Equations
          Instances For

            If X is a type with [Fintype X] and G acts on X, then G also acts on FintypeCat.of X.

            Equations

            Bundles a finite type H with a multiplicative action of G as an Action.

            Equations
            Instances For
              @[simp]
              theorem Action.FintypeCat.ofMulAction_apply {G : Type u_1} {H : FintypeCat} [Monoid G] [MulAction G H.carrier] (g : G) (x : H.carrier) :
              (ofMulAction G H).ρ g x = g x

              Shorthand notation for the quotient of G by H as a finite G-set.

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

                If N is a normal subgroup of G, then this is the group homomorphism sending an element g of G to the G-endomorphism of G ⧸ₐ N given by multiplication with g⁻¹ on the right.

                Equations
                Instances For
                  @[simp]
                  theorem Action.FintypeCat.toEndHom_apply {G : Type u_1} [Group G] (N : Subgroup G) [Fintype (G N)] [N.Normal] (g h : G) :

                  If H and N are subgroups of a group G with N normal, there is a canonical group homomorphism H ⧸ N ⊓ H to the G-endomorphisms of G ⧸ N.

                  Equations
                  Instances For
                    @[simp]
                    theorem Action.FintypeCat.quotientToEndHom_mk {G : Type u_1} [Group G] (H N : Subgroup G) [Fintype (G N)] [N.Normal] (x : H) (g : G) :

                    If N and H are subgroups of a group G with N ≤ H, this is the canonical G-morphism G ⧸ N ⟶ G ⧸ H.

                    Equations
                    Instances For
                      @[simp]
                      theorem Action.FintypeCat.quotientToQuotientOfLE_hom_mk {G : Type u_1} [Group G] (H N : Subgroup G) [Fintype (G N)] [Fintype (G H)] (h : N H) (x : G) :
                      instance Action.instMulAction {V : Type (u + 1)} [CategoryTheory.LargeCategory V] {FV : VVType u_1} {CV : VType u_2} [(X Y : V) → FunLike (FV X Y) (CV X) (CV Y)] [CategoryTheory.ConcreteCategory V FV] {G : Type u_3} [Monoid G] (X : Action V G) :
                      Equations