Documentation

Mathlib.RingTheory.Coalgebra.GroupLike

Group-like elements in a coalgebra #

This file defines group-like elements in a coalgebra, i.e. elements a such that ε a = 1 and Δ a = a ⊗ₜ a.

Main declarations #

structure IsGroupLikeElem (R : Type u_2) {A : Type u_3} [CommSemiring R] [AddCommMonoid A] [Module R A] [Coalgebra R A] (a : A) :

A group-like element in a coalgebra is an element a such that ε(a) = 1 and Δ(a) = a ⊗ₜ a, where ε and Δ are the counit and comultiplication respectively.

Instances For
    @[simp]
    theorem isGroupLikeElem_self {R : Type u_2} [CommSemiring R] {r : R} :
    theorem IsGroupLikeElem.ne_zero {R : Type u_2} {A : Type u_3} [CommSemiring R] [AddCommMonoid A] [Module R A] [Coalgebra R A] {a : A} [Nontrivial R] (ha : IsGroupLikeElem R a) :
    a 0
    theorem IsGroupLikeElem.map {F : Type u_1} {R : Type u_2} {A : Type u_3} {B : Type u_4} [CommSemiring R] [AddCommMonoid A] [AddCommMonoid B] [Module R A] [Coalgebra R A] [Module R B] [Coalgebra R B] {a : A} [FunLike F A B] [CoalgHomClass F R A B] (f : F) (ha : IsGroupLikeElem R a) :

    A coalgebra homomorphism sends group-like elements to group-like elements.

    @[simp]
    theorem isGroupLikeElem_map_equiv {F : Type u_1} {R : Type u_2} {A : Type u_3} {B : Type u_4} [CommSemiring R] [AddCommMonoid A] [AddCommMonoid B] [Module R A] [Coalgebra R A] [Module R B] [Coalgebra R B] {a : A} [EquivLike F A B] [CoalgEquivClass F R A B] (f : F) :

    A coalgebra isomorphism preserves group-like elements.

    structure GroupLike (R : Type u_2) (A : Type u_3) [CommSemiring R] [AddCommMonoid A] [Module R A] [Coalgebra R A] :
    Type u_3

    The type of group-like elements in a coalgebra.

    • val : A

      The underlying element of a group-like element.

    • isGroupLikeElem_val : IsGroupLikeElem R self
    Instances For
      theorem GroupLike.ext {R : Type u_2} {A : Type u_3} {inst✝ : CommSemiring R} {inst✝¹ : AddCommMonoid A} {inst✝² : Module R A} {inst✝³ : Coalgebra R A} {x y : GroupLike R A} (val : x = y) :
      x = y
      theorem GroupLike.ext_iff {R : Type u_2} {A : Type u_3} {inst✝ : CommSemiring R} {inst✝¹ : AddCommMonoid A} {inst✝² : Module R A} {inst✝³ : Coalgebra R A} {x y : GroupLike R A} :
      x = y x = y
      instance GroupLike.instCoeOut {R : Type u_2} {A : Type u_3} [CommSemiring R] [AddCommMonoid A] [Module R A] [Coalgebra R A] :
      Equations
      @[simp]
      theorem GroupLike.val_inj {R : Type u_2} {A : Type u_3} [CommSemiring R] [AddCommMonoid A] [Module R A] [Coalgebra R A] {a b : GroupLike R A} :
      a = b a = b

      Identity equivalence between GroupLike R A and {a : A // IsGroupLikeElem R a}.

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

        Group-like elements over a domain are linearly independent.

        Group-like elements over a domain are linearly independent.