Documentation

Mathlib.LinearAlgebra.SymmetricAlgebra.Basic

Symmetric Algebras #

Given a commutative semiring R, and an R-module M, we construct the symmetric algebra of M. This is the free commutative R-algebra generated (R-linearly) by the module M.

Notation #

Note #

See SymAlg R instead if you are looking for the symmetrized algebra, which gives a commutative multiplication on R by $a \circ b = \frac{1}{2}(ab + ba)$.

inductive TensorAlgebra.SymRel (R : Type u_1) (M : Type u_2) [CommSemiring R] [AddCommMonoid M] [Module R M] :

Relation on the tensor algebra which will yield the symmetric algebra when quotiented out by.

Instances For
    @[reducible, inline]
    abbrev SymmetricAlgebra (R : Type u_1) (M : Type u_2) [CommSemiring R] [AddCommMonoid M] [Module R M] :
    Type (max u_1 u_2)

    Concrete construction of the symmetric algebra of M by quotienting out the tensor algebra by the commutativity relation.

    Equations
    Instances For
      @[reducible, inline]

      Algebra homomorphism from the tensor algebra over M to the symmetric algebra over M.

      Equations
      Instances For

        Canonical inclusion of M into the symmetric algebra SymmetricAlgebra R M.

        Equations
        Instances For
          theorem SymmetricAlgebra.induction (R : Type u_1) (M : Type u_2) [CommSemiring R] [AddCommMonoid M] [Module R M] {motive : SymmetricAlgebra R MProp} (algebraMap : ∀ (r : R), motive ((algebraMap R (SymmetricAlgebra R M)) r)) (ι : ∀ (x : M), motive ((ι R M) x)) (mul : ∀ (a b : SymmetricAlgebra R M), motive amotive bmotive (a * b)) (add : ∀ (a b : SymmetricAlgebra R M), motive amotive bmotive (a + b)) (a : SymmetricAlgebra R M) :
          motive a
          Equations
          • One or more equations did not get rendered due to their size.
          def SymmetricAlgebra.lift {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] :

          For any linear map f : M →ₗ[R] A, SymmetricAlgebra.lift f lifts the linear map to an R-algebra homomorphism from SymmetricAlgebra R M to A.

          Equations
          Instances For
            @[simp]
            theorem SymmetricAlgebra.lift_ι_apply {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] (f : M →ₗ[R] A) (a : M) :
            (lift f) ((ι R M) a) = f a
            @[simp]
            theorem SymmetricAlgebra.lift_comp_ι {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] (f : M →ₗ[R] A) :
            (lift f) ∘ₗ ι R M = f
            theorem SymmetricAlgebra.algHom_ext {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {F G : SymmetricAlgebra R M →ₐ[R] A} (h : F ∘ₗ ι R M = G ∘ₗ ι R M) :
            F = G
            theorem SymmetricAlgebra.algHom_ext_iff {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {F G : SymmetricAlgebra R M →ₐ[R] A} :
            F = G F ∘ₗ ι R M = G ∘ₗ ι R M
            @[simp]
            theorem SymmetricAlgebra.lift_ι {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] :
            @[simp]
            theorem SymmetricAlgebra.algebraMap_inj {R : Type u_1} (M : Type u_2) [CommSemiring R] [AddCommMonoid M] [Module R M] (x y : R) :
            @[simp]
            theorem SymmetricAlgebra.algebraMap_eq_zero_iff {R : Type u_1} (M : Type u_2) [CommSemiring R] [AddCommMonoid M] [Module R M] (x : R) :
            (algebraMap R (SymmetricAlgebra R M)) x = 0 x = 0
            @[simp]
            theorem SymmetricAlgebra.algebraMap_eq_one_iff {R : Type u_1} (M : Type u_2) [CommSemiring R] [AddCommMonoid M] [Module R M] (x : R) :
            (algebraMap R (SymmetricAlgebra R M)) x = 1 x = 1

            A SymmetricAlgebra over a nontrivial semiring is nontrivial.

            def IsSymmetricAlgebra {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] (f : M →ₗ[R] A) :

            Given a morphism f : M →ₗ[R] A, IsSymmetricAlgebra f is a proposition saying that the algebra homomorphism from SymmetricAlgebra R M to A is bijective.

            Equations
            Instances For
              noncomputable def IsSymmetricAlgebra.equiv {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (h : IsSymmetricAlgebra f) :

              For f : M →ₗ[R] A, construst the algebra isomorphism SymmetricAlgebra R M ≃ₐ[R] A from IsSymmetricAlgebra f.

              Equations
              Instances For
                @[simp]
                theorem IsSymmetricAlgebra.equiv_apply {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (h : IsSymmetricAlgebra f) (a : SymmetricAlgebra R M) :
                @[simp]
                theorem IsSymmetricAlgebra.equiv_toAlgHom {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (h : IsSymmetricAlgebra f) :
                @[simp]
                theorem IsSymmetricAlgebra.equiv_symm_apply {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (h : IsSymmetricAlgebra f) (a : M) :
                @[simp]
                theorem IsSymmetricAlgebra.equiv_symm_comp {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (h : IsSymmetricAlgebra f) :
                theorem IsSymmetricAlgebra.of_equiv {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (e : SymmetricAlgebra R M ≃ₐ[R] A) (he : e ∘ₗ SymmetricAlgebra.ι R M = f) :
                noncomputable def IsSymmetricAlgebra.lift {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (h : IsSymmetricAlgebra f) {A' : Type u_4} [CommSemiring A'] [Algebra R A'] (g : M →ₗ[R] A') :
                A →ₐ[R] A'

                Given a morphism g : M →ₗ[R] A', lift this to a morphism of type A →ₐ[R] A' (where A satisfies the universal property of the symmetric algebra of M)

                Equations
                Instances For
                  @[simp]
                  theorem IsSymmetricAlgebra.lift_eq {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (h : IsSymmetricAlgebra f) {A' : Type u_4} [CommSemiring A'] [Algebra R A'] (g : M →ₗ[R] A') (a : M) :
                  (h.lift g) (f a) = g a
                  @[simp]
                  theorem IsSymmetricAlgebra.lift_comp_linearMap {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (h : IsSymmetricAlgebra f) {A' : Type u_4} [CommSemiring A'] [Algebra R A'] (g : M →ₗ[R] A') :
                  (h.lift g) ∘ₗ f = g
                  theorem IsSymmetricAlgebra.algHom_ext {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} {A' : Type u_4} [CommSemiring A'] [Algebra R A'] (h : IsSymmetricAlgebra f) {F G : A →ₐ[R] A'} (hFG : F ∘ₗ f = G ∘ₗ f) :
                  F = G
                  theorem IsSymmetricAlgebra.lift_unique {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (h : IsSymmetricAlgebra f) {A' : Type u_4} [CommSemiring A'] [Algebra R A'] {g : M →ₗ[R] A'} {F : A →ₐ[R] A'} (hF : F ∘ₗ f = g) :
                  F = h.lift g
                  theorem IsSymmetricAlgebra.induction {R : Type u_1} {M : Type u_2} [CommSemiring R] [AddCommMonoid M] [Module R M] {A : Type u_3} [CommSemiring A] [Algebra R A] {f : M →ₗ[R] A} (h : IsSymmetricAlgebra f) {motive : AProp} (algebraMap : ∀ (r : R), motive ((algebraMap R A) r)) (ι : ∀ (x : M), motive (f x)) (mul : ∀ (a b : A), motive amotive bmotive (a * b)) (add : ∀ (a b : A), motive amotive bmotive (a + b)) (a : A) :
                  motive a