Documentation

Mathlib.RingTheory.Bialgebra.Hom

Homomorphisms of R-bialgebras #

This file defines bundled homomorphisms of R-bialgebras. We simply mimic Mathlib/Algebra/Algebra/Hom.lean.

Main definitions #

Notations #

structure BialgHom (R : Type u_1) (A : Type u_2) (B : Type u_3) [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] extends CoalgHom :
Type (max u_2 u_3)

Given R-algebras A, B with comultiplication maps Δ_A, Δ_B and counit maps ε_A, ε_B, an R-bialgebra homomorphism A →ₐc[R] B is an R-algebra map f such that ε_B ∘ f = ε_A and (f ⊗ f) ∘ Δ_A = Δ_B ∘ f.

  • toFun : AB
  • map_add' : ∀ (x y : A), self.toFun (x + y) = self.toFun x + self.toFun y
  • map_smul' : ∀ (m : R) (x : A), self.toFun (m x) = (RingHom.id R) m self.toFun x
  • counit_comp : Coalgebra.counit ∘ₗ self.toLinearMap = Coalgebra.counit
  • map_comp_comul : TensorProduct.map self.toLinearMap self.toLinearMap ∘ₗ Coalgebra.comul = Coalgebra.comul ∘ₗ self.toLinearMap
  • map_one' : self.toFun 1 = 1

    The proposition that the function preserves 1

  • map_mul' : ∀ (x y : A), self.toFun (x * y) = self.toFun x * self.toFun y

    The proposition that the function preserves multiplication

Instances For
    @[reducible]
    abbrev BialgHom.toMonoidHom {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (self : A →ₐc[R] B) :
    A →* B

    Reinterpret a BialgHom as a MonoidHom

    Equations
    • self.toMonoidHom = { toFun := self.toFun, map_one' := , map_mul' := }
    Instances For

      Given R-algebras A, B with comultiplication maps Δ_A, Δ_B and counit maps ε_A, ε_B, an R-bialgebra homomorphism A →ₐc[R] B is an R-algebra map f such that ε_B ∘ f = ε_A and (f ⊗ f) ∘ Δ_A = Δ_B ∘ f.

      Equations
      Instances For

        Given R-algebras A, B with comultiplication maps Δ_A, Δ_B and counit maps ε_A, ε_B, an R-bialgebra homomorphism A →ₐc[R] B is an R-algebra map f such that ε_B ∘ f = ε_A and (f ⊗ f) ∘ Δ_A = Δ_B ∘ f.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          class BialgHomClass (F : Type u_1) (R : outParam (Type u_2)) (A : outParam (Type u_3)) (B : outParam (Type u_4)) [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] [FunLike F A B] extends CoalgHomClass , MonoidHomClass :

          BialgHomClass F R A B asserts F is a type of bundled bialgebra homomorphisms from A to B.

            Instances
              @[instance 100]
              instance BialgHomClass.toAlgHomClass {R : Type u_1} {A : Type u_2} {B : Type u_3} {F : Type u_4} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] [FunLike F A B] [BialgHomClass F R A B] :
              AlgHomClass F R A B
              Equations
              • =
              def BialgHomClass.toBialgHom {R : Type u_1} {A : Type u_2} {B : Type u_3} {F : Type u_4} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] [FunLike F A B] [BialgHomClass F R A B] (f : F) :

              Turn an element of a type F satisfying BialgHomClass F R A B into an actual BialgHom. This is declared as the default coercion from F to A →ₐc[R] B.

              Equations
              • f = let __src := f; let __src := f; { toFun := f, map_add' := , map_smul' := , counit_comp := , map_comp_comul := , map_one' := , map_mul' := }
              Instances For
                instance BialgHomClass.instCoeToBialgHom {R : Type u_1} {A : Type u_2} {B : Type u_3} {F : Type u_4} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] [FunLike F A B] [BialgHomClass F R A B] :
                Equations
                • BialgHomClass.instCoeToBialgHom = { coe := BialgHomClass.toBialgHom }
                @[simp]
                theorem BialgHomClass.counitAlgHom_comp {R : Type u_1} {A : Type u_2} {B : Type u_3} {F : Type u_4} [CommSemiring R] [Semiring A] [Bialgebra R A] [Semiring B] [Bialgebra R B] [FunLike F A B] [BialgHomClass F R A B] (f : F) :
                @[simp]
                theorem BialgHomClass.map_comp_comulAlgHom {R : Type u_1} {A : Type u_2} {B : Type u_3} {F : Type u_4} [CommSemiring R] [Semiring A] [Bialgebra R A] [Semiring B] [Bialgebra R B] [FunLike F A B] [BialgHomClass F R A B] (f : F) :
                instance BialgHom.funLike {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] :
                FunLike (A →ₐc[R] B) A B
                Equations
                • BialgHom.funLike = { coe := fun (f : A →ₐc[R] B) => f.toFun, coe_injective' := }
                instance BialgHom.bialgHomClass {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] :
                Equations
                • =
                def BialgHom.Simps.apply {R : Type u_6} {α : Type u_7} {β : Type u_8} [CommSemiring R] [Semiring α] [Algebra R α] [Semiring β] [Algebra R β] [CoalgebraStruct R α] [CoalgebraStruct R β] (f : α →ₐc[R] β) :
                αβ

                See Note [custom simps projection]

                Equations
                Instances For
                  @[simp]
                  theorem BialgHom.coe_coe {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] {F : Type u_6} [FunLike F A B] [BialgHomClass F R A B] (f : F) :
                  f = f
                  @[simp]
                  theorem BialgHom.coe_mk {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] {f : A →ₗc[R] B} (h : f.toFun 1 = 1) (h₁ : ∀ (x y : A), f.toFun (x * y) = f.toFun x * f.toFun y) :
                  { toCoalgHom := f, map_one' := h, map_mul' := h₁ } = f
                  theorem BialgHom.coe_mks {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] {f : AB} (h₀ : ∀ (x y : A), f (x + y) = f x + f y) (h₁ : ∀ (m : R) (x : A), { toFun := f, map_add' := h₀ }.toFun (m x) = (RingHom.id R) m { toFun := f, map_add' := h₀ }.toFun x) (h₂ : Coalgebra.counit ∘ₗ { toFun := f, map_add' := h₀, map_smul' := h₁ } = Coalgebra.counit) (h₃ : TensorProduct.map { toFun := f, map_add' := h₀, map_smul' := h₁ } { toFun := f, map_add' := h₀, map_smul' := h₁ } ∘ₗ Coalgebra.comul = Coalgebra.comul ∘ₗ { toFun := f, map_add' := h₀, map_smul' := h₁ }) (h₄ : { toFun := f, map_add' := h₀, map_smul' := h₁, counit_comp := h₂, map_comp_comul := h₃ }.toFun 1 = 1) (h₅ : ∀ (x y : A), { toFun := f, map_add' := h₀, map_smul' := h₁, counit_comp := h₂, map_comp_comul := h₃ }.toFun (x * y) = { toFun := f, map_add' := h₀, map_smul' := h₁, counit_comp := h₂, map_comp_comul := h₃ }.toFun x * { toFun := f, map_add' := h₀, map_smul' := h₁, counit_comp := h₂, map_comp_comul := h₃ }.toFun y) :
                  { toFun := f, map_add' := h₀, map_smul' := h₁, counit_comp := h₂, map_comp_comul := h₃, map_one' := h₄, map_mul' := h₅ } = f
                  @[simp]
                  theorem BialgHom.coe_coalgHom_mk {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] {f : A →ₗc[R] B} (h : f.toFun 1 = 1) (h₁ : ∀ (x y : A), f.toFun (x * y) = f.toFun x * f.toFun y) :
                  { toCoalgHom := f, map_one' := h, map_mul' := h₁ } = f
                  theorem BialgHom.coe_toCoalgHom {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (f : A →ₐc[R] B) :
                  f = f
                  @[simp]
                  theorem BialgHom.coe_toLinearMap {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (f : A →ₐc[R] B) :
                  f = f
                  theorem BialgHom.coe_toAlgHom {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (f : A →ₐc[R] B) :
                  f = f
                  theorem BialgHom.toAlgHom_toLinearMap {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (f : A →ₐc[R] B) :
                  (f).toLinearMap = f
                  theorem BialgHom.coe_fn_injective {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] :
                  Function.Injective DFunLike.coe
                  theorem BialgHom.coe_fn_inj {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] {φ₁ : A →ₐc[R] B} {φ₂ : A →ₐc[R] B} :
                  φ₁ = φ₂ φ₁ = φ₂
                  theorem BialgHom.coe_coalgHom_injective {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] :
                  Function.Injective CoalgHomClass.toCoalgHom
                  theorem BialgHom.coe_algHom_injective {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] :
                  Function.Injective AlgHomClass.toAlgHom
                  theorem BialgHom.coe_linearMap_injective {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] :
                  Function.Injective fun (x : A →ₐc[R] B) => x
                  theorem BialgHom.congr_fun {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] {φ₁ : A →ₐc[R] B} {φ₂ : A →ₐc[R] B} (H : φ₁ = φ₂) (x : A) :
                  φ₁ x = φ₂ x
                  theorem BialgHom.congr_arg {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (φ : A →ₐc[R] B) {x : A} {y : A} (h : x = y) :
                  φ x = φ y
                  theorem BialgHom.ext {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] {φ₁ : A →ₐc[R] B} {φ₂ : A →ₐc[R] B} (H : ∀ (x : A), φ₁ x = φ₂ x) :
                  φ₁ = φ₂
                  theorem BialgHom.ext_iff {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] {φ₁ : A →ₐc[R] B} {φ₂ : A →ₐc[R] B} :
                  φ₁ = φ₂ ∀ (x : A), φ₁ x = φ₂ x
                  theorem BialgHom.ext_of_ring {R : Type u_1} {A : Type u_2} [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] {f : R →ₐc[R] A} {g : R →ₐc[R] A} (h : f 1 = g 1) :
                  f = g
                  @[simp]
                  theorem BialgHom.mk_coe {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] {f : A →ₐc[R] B} (h₀ : ∀ (x y : A), f (x + y) = f x + f y) (h₁ : ∀ (m : R) (x : A), { toFun := f, map_add' := h₀ }.toFun (m x) = (RingHom.id R) m { toFun := f, map_add' := h₀ }.toFun x) (h₂ : Coalgebra.counit ∘ₗ { toFun := f, map_add' := h₀, map_smul' := h₁ } = Coalgebra.counit) (h₃ : TensorProduct.map { toFun := f, map_add' := h₀, map_smul' := h₁ } { toFun := f, map_add' := h₀, map_smul' := h₁ } ∘ₗ Coalgebra.comul = Coalgebra.comul ∘ₗ { toFun := f, map_add' := h₀, map_smul' := h₁ }) (h₄ : { toFun := f, map_add' := h₀, map_smul' := h₁, counit_comp := h₂, map_comp_comul := h₃ }.toFun 1 = 1) (h₅ : ∀ (x y : A), { toFun := f, map_add' := h₀, map_smul' := h₁, counit_comp := h₂, map_comp_comul := h₃ }.toFun (x * y) = { toFun := f, map_add' := h₀, map_smul' := h₁, counit_comp := h₂, map_comp_comul := h₃ }.toFun x * { toFun := f, map_add' := h₀, map_smul' := h₁, counit_comp := h₂, map_comp_comul := h₃ }.toFun y) :
                  { toFun := f, map_add' := h₀, map_smul' := h₁, counit_comp := h₂, map_comp_comul := h₃, map_one' := h₄, map_mul' := h₅ } = f
                  def BialgHom.copy {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (f : A →ₐc[R] B) (f' : AB) (h : f' = f) :

                  Copy of a BialgHom with a new toFun equal to the old one. Useful to fix definitional equalities.

                  Equations
                  • f.copy f' h = { toCoalgHom := (f).copy f' h, map_one' := , map_mul' := }
                  Instances For
                    @[simp]
                    theorem BialgHom.coe_copy {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (f : A →ₗc[R] B) (f' : AB) (h : f' = f) :
                    (f.copy f' h) = f'
                    theorem BialgHom.copy_eq {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (f : A →ₗc[R] B) (f' : AB) (h : f' = f) :
                    f.copy f' h = f
                    @[simp]
                    theorem BialgHom.id_apply (R : Type u_1) (A : Type u_2) [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] (a : A) :
                    (BialgHom.id R A) a = a
                    def BialgHom.id (R : Type u_1) (A : Type u_2) [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] :

                    Identity map as a BialgHom.

                    Equations
                    Instances For
                      @[simp]
                      theorem BialgHom.coe_id {R : Type u_1} {A : Type u_2} [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] :
                      (BialgHom.id R A) = id
                      @[simp]
                      theorem BialgHom.id_toCoalgHom {R : Type u_1} {A : Type u_2} [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] :
                      @[simp]
                      theorem BialgHom.id_toAlgHom {R : Type u_1} {A : Type u_2} [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] :
                      (BialgHom.id R A) = AlgHom.id R A
                      @[simp]
                      theorem BialgHom.comp_apply {R : Type u_1} {A : Type u_2} {B : Type u_3} {C : Type u_4} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Semiring C] [Algebra R C] [CoalgebraStruct R A] [CoalgebraStruct R B] [CoalgebraStruct R C] (φ₁ : B →ₐc[R] C) (φ₂ : A →ₐc[R] B) :
                      ∀ (a : A), (φ₁.comp φ₂) a = φ₁ (φ₂ a)
                      def BialgHom.comp {R : Type u_1} {A : Type u_2} {B : Type u_3} {C : Type u_4} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Semiring C] [Algebra R C] [CoalgebraStruct R A] [CoalgebraStruct R B] [CoalgebraStruct R C] (φ₁ : B →ₐc[R] C) (φ₂ : A →ₐc[R] B) :

                      Composition of bialgebra homomorphisms.

                      Equations
                      • φ₁.comp φ₂ = let __src := (φ₁).comp φ₂; let __src_1 := (φ₁).comp φ₂; { toCoalgHom := __src, map_one' := , map_mul' := }
                      Instances For
                        @[simp]
                        theorem BialgHom.coe_comp {R : Type u_1} {A : Type u_2} {B : Type u_3} {C : Type u_4} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Semiring C] [Algebra R C] [CoalgebraStruct R A] [CoalgebraStruct R B] [CoalgebraStruct R C] (φ₁ : B →ₐc[R] C) (φ₂ : A →ₐc[R] B) :
                        (φ₁.comp φ₂) = φ₁ φ₂
                        @[simp]
                        theorem BialgHom.comp_toCoalgHom {R : Type u_1} {A : Type u_2} {B : Type u_3} {C : Type u_4} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Semiring C] [Algebra R C] [CoalgebraStruct R A] [CoalgebraStruct R B] [CoalgebraStruct R C] (φ₁ : B →ₐc[R] C) (φ₂ : A →ₐc[R] B) :
                        (φ₁.comp φ₂) = (φ₁).comp φ₂
                        @[simp]
                        theorem BialgHom.comp_toAlgHom {R : Type u_1} {A : Type u_2} {B : Type u_3} {C : Type u_4} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Semiring C] [Algebra R C] [CoalgebraStruct R A] [CoalgebraStruct R B] [CoalgebraStruct R C] (φ₁ : B →ₐc[R] C) (φ₂ : A →ₐc[R] B) :
                        (φ₁.comp φ₂) = (φ₁).comp φ₂
                        @[simp]
                        theorem BialgHom.comp_id {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (φ : A →ₐc[R] B) :
                        φ.comp (BialgHom.id R A) = φ
                        @[simp]
                        theorem BialgHom.id_comp {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (φ : A →ₐc[R] B) :
                        (BialgHom.id R B).comp φ = φ
                        theorem BialgHom.comp_assoc {R : Type u_1} {A : Type u_2} {B : Type u_3} {C : Type u_4} {D : Type u_5} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Semiring C] [Algebra R C] [Semiring D] [Algebra R D] [CoalgebraStruct R A] [CoalgebraStruct R B] [CoalgebraStruct R C] [CoalgebraStruct R D] (φ₁ : C →ₐc[R] D) (φ₂ : B →ₐc[R] C) (φ₃ : A →ₐc[R] B) :
                        (φ₁.comp φ₂).comp φ₃ = φ₁.comp (φ₂.comp φ₃)
                        theorem BialgHom.map_smul_of_tower {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [CoalgebraStruct R A] [CoalgebraStruct R B] (φ : A →ₐc[R] B) {R' : Type u_6} [SMul R' A] [SMul R' B] [LinearMap.CompatibleSMul A B R' R] (r : R') (x : A) :
                        φ (r x) = r φ x
                        theorem BialgHom.End_toSemigroup_toMul_mul {R : Type u_1} {A : Type u_2} [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] (φ₁ : A →ₐc[R] A) (φ₂ : A →ₐc[R] A) :
                        φ₁ * φ₂ = φ₁.comp φ₂
                        theorem BialgHom.End_toOne_one {R : Type u_1} {A : Type u_2} [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] :
                        instance BialgHom.End {R : Type u_1} {A : Type u_2} [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] :
                        Equations
                        @[simp]
                        theorem BialgHom.one_apply {R : Type u_1} {A : Type u_2} [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] (x : A) :
                        1 x = x
                        @[simp]
                        theorem BialgHom.mul_apply {R : Type u_1} {A : Type u_2} [CommSemiring R] [Semiring A] [Algebra R A] [CoalgebraStruct R A] (φ : A →ₐc[R] A) (ψ : A →ₐc[R] A) (x : A) :
                        (φ * ψ) x = φ (ψ x)

                        The counit of a bialgebra as a BialgHom.

                        Equations
                        Instances For
                          @[simp]
                          theorem Bialgebra.counitBialgHom_apply (R : Type u) (A : Type v) [CommSemiring R] [Semiring A] [Bialgebra R A] (x : A) :
                          (Bialgebra.counitBialgHom R A) x = Coalgebra.counit x
                          Equations
                          • =
                          theorem Bialgebra.ext_to_ring {R : Type u} (A : Type v) [CommSemiring R] [Semiring A] [Bialgebra R A] (f : A →ₐc[R] R) (g : A →ₐc[R] R) :
                          f = g