Reindexed basis #
This file introduces an equivalence between the set of embeddings of K
into ℂ
and the
index set of the chosen basis of the ring of integers of K
.
Tags #
house, number field, algebraic number
@[reducible, inline]
abbrev
NumberField.equivReindex
(K : Type u_1)
[Field K]
[NumberField K]
:
(K →+* ℂ) ≃ Module.Free.ChooseBasisIndex ℤ (RingOfIntegers K)
An equivalence between the set of embeddings of K
into ℂ
and the
index set of the chosen basis of the ring of integers of K
.
Equations
Instances For
@[reducible, inline]
The basis matrix for the embeddings of K
into ℂ
. This matrix is formed by
taking the lattice basis vectors of K
and reindexing them according to the
equivalence equivReindex
, then transposing the resulting matrix.
Equations
- NumberField.basisMatrix K = Matrix.of fun (i : K →+* ℂ) => (NumberField.canonicalEmbedding.latticeBasis K) ((NumberField.equivReindex K) i)
Instances For
theorem
NumberField.det_of_basisMatrix_non_zero
(K : Type u_1)
[Field K]
[NumberField K]
[DecidableEq (K →+* ℂ)]
:
(basisMatrix K).det ≠ 0
instance
NumberField.instInvertibleMatrixRingHomComplexBasisMatrix
(K : Type u_1)
[Field K]
[NumberField K]
[DecidableEq (K →+* ℂ)]
:
Equations
- NumberField.instInvertibleMatrixRingHomComplexBasisMatrix K = (NumberField.basisMatrix K).invertibleOfIsUnitDet ⋯
theorem
NumberField.canonicalEmbedding_eq_basisMatrix_mulVec
{K : Type u_1}
[Field K]
[NumberField K]
(α : K)
:
(canonicalEmbedding K) α = (basisMatrix K).transpose.mulVec fun (i : K →+* ℂ) => ↑((((integralBasis K).reindex (equivReindex K).symm).repr α) i)
theorem
NumberField.inverse_basisMatrix_mulVec_eq_repr
{K : Type u_1}
[Field K]
[NumberField K]
[DecidableEq (K →+* ℂ)]
(α : RingOfIntegers K)
(i : K →+* ℂ)
:
(basisMatrix K).transpose⁻¹.mulVec (fun (j : K →+* ℂ) => (canonicalEmbedding K) ((algebraMap (RingOfIntegers K) K) α) j)
i = ↑((((integralBasis K).reindex (equivReindex K).symm).repr ↑α) i)