Documentation

Mathlib.RingTheory.Norm.Defs

Norm for (finite) ring extensions #

Suppose we have an R-algebra S with a finite basis. For each s : S, the determinant of the linear map given by multiplying by s gives information about the roots of the minimal polynomial of s over R.

Implementation notes #

Typically, the norm is defined specifically for finite field extensions. The current definition is as general as possible and the assumption that we have fields or that the extension is finite is added to the lemmas as needed.

We only define the norm for left multiplication (Algebra.leftMulMatrix, i.e. LinearMap.mulLeft). For now, the definitions assume S is commutative, so the choice doesn't matter anyway.

See also Algebra.trace, which is defined similarly as the trace of Algebra.leftMulMatrix.

References #

noncomputable def Algebra.norm (R : Type u_1) {S : Type u_2} [CommRing R] [Ring S] [Algebra R S] :
S →* R

The norm of an element s of an R-algebra is the determinant of (*) s.

Equations
Instances For
    theorem Algebra.norm_apply (R : Type u_1) {S : Type u_2} [CommRing R] [Ring S] [Algebra R S] (x : S) :
    (Algebra.norm R) x = LinearMap.det ((Algebra.lmul R S) x)
    theorem Algebra.norm_eq_one_of_not_exists_basis (R : Type u_1) {S : Type u_2} [CommRing R] [Ring S] [Algebra R S] (h : ¬∃ (s : Finset S), Nonempty (Basis { x : S // x s } R S)) (x : S) :
    (Algebra.norm R) x = 1
    theorem Algebra.norm_eq_one_of_not_module_finite {R : Type u_1} {S : Type u_2} [CommRing R] [Ring S] [Algebra R S] (h : ¬Module.Finite R S) (x : S) :
    (Algebra.norm R) x = 1
    theorem Algebra.norm_eq_matrix_det {R : Type u_1} {S : Type u_2} [CommRing R] [Ring S] [Algebra R S] {ι : Type w} [Fintype ι] [DecidableEq ι] (b : Basis ι R S) (s : S) :
    theorem Algebra.norm_algebraMap_of_basis {R : Type u_1} {S : Type u_2} [CommRing R] [Ring S] [Algebra R S] {ι : Type w} [Fintype ι] (b : Basis ι R S) (x : R) :

    If x is in the base ring K, then the norm is x ^ [L : K].

    @[simp]
    theorem Algebra.norm_algebraMap {K : Type u_4} [Field K] {L : Type u_7} [Ring L] [Algebra K L] (x : K) :

    If x is in the base field K, then the norm is x ^ [L : K].

    (If L is not finite-dimensional over K, then norm = 1 = x ^ 0 = x ^ (finrank L K).)