Documentation

Mathlib.RepresentationTheory.Action.Concrete

Constructors for Action V G for some concrete categories #

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

def Action.ofMulAction (G : Type u) (H : Type u) [Monoid G] [MulAction G H] :

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

Equations
Instances For
    @[simp]
    theorem Action.ofMulAction_apply {G : Type u} {H : Type u} [Monoid G] [MulAction G H] (g : G) (x : H) :
    (Action.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

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

      Equations
      Instances For
        @[simp]
        theorem Action.leftRegular_ρ_apply (G : Type u) [Monoid G] :
        ∀ (x1 : (MonCat.of G)) (x2 : G), (Action.leftRegular G) x1 x2 = x1 * x2
        @[simp]
        theorem Action.leftRegular_V (G : Type u) [Monoid G] :
        def Action.diagonal (G : Type u) [Monoid G] (n : ) :

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

        Equations
        Instances For
          @[simp]
          theorem Action.diagonal_ρ_apply (G : Type u) [Monoid G] (n : ) :
          ∀ (x1 : (MonCat.of G)) (x2 : Fin nG) (a : Fin n), (Action.diagonal G n) x1 x2 a = (x1 x2) a
          @[simp]
          theorem Action.diagonal_V (G : Type u) [Monoid G] (n : ) :
          (Action.diagonal G n).V = (Fin nG)

          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} {H : FintypeCat} [Monoid G] [MulAction G H] (g : G) (x : H) :

              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 : G) (h : G) :
                  ((Action.FintypeCat.toEndHom N) g).hom h = h * g⁻¹
                  def Action.FintypeCat.quotientToEndHom {G : Type u_1} [Group G] (H : Subgroup G) (N : Subgroup G) [Fintype (G N)] [N.Normal] :

                  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 : Subgroup G) (N : Subgroup G) [Fintype (G N)] [N.Normal] (x : H) (g : G) :
                    ((Action.FintypeCat.quotientToEndHom H N) x).hom g = g * x⁻¹

                    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 : Subgroup G) (N : Subgroup G) [Fintype (G N)] [Fintype (G H)] (h : N H) (x : G) :
                      (Action.FintypeCat.quotientToQuotientOfLE H N h).hom x = x
                      Equations
                      Equations
                      • X.instMulActionαMonoidFintypeVFintypeCat = X.instMulAction
                      Equations
                      • X.instMulActionαFintypeVFintypeCat = X.instMulAction