Roots of unity and primitive roots of unity #
We define roots of unity in the context of an arbitrary commutative monoid,
as a subgroup of the group of units. We also define a predicate IsPrimitiveRoot on commutative
monoids, expressing that an element is a primitive root of unity.
Main definitions #
- rootsOfUnity n M, for- n : ℕ+is the subgroup of the units of a commutative monoid- Mconsisting of elements- xthat satisfy- x ^ n = 1.
- IsPrimitiveRoot ζ k: an element- ζis a primitive- k-th root of unity if- ζ ^ k = 1, and if- lsatisfies- ζ ^ l = 1then- k ∣ l.
- primitiveRoots k R: the finset of primitive- k-th roots of unity in an integral domain- R.
- IsPrimitiveRoot.autToPow: the monoid hom that takes an automorphism of a ring to the power it sends that specific primitive root, as a member of- (ZMod n)ˣ.
Main results #
- rootsOfUnity.isCyclic: the roots of unity in an integral domain form a cyclic group.
- IsPrimitiveRoot.zmodEquivZPowers:- ZMod kis equivalent to the subgroup generated by a primitive- k-th root of unity.
- IsPrimitiveRoot.zpowers_eq: in an integral domain, the subgroup generated by a primitive- k-th root of unity is equal to the- k-th roots of unity.
- IsPrimitiveRoot.card_primitiveRoots: if an integral domain has a primitive- k-th root of unity, then it has- φ kof them.
Implementation details #
It is desirable that rootsOfUnity is a subgroup,
and it will mainly be applied to rings (e.g. the ring of integers in a number field) and fields.
We therefore implement it as a subgroup of the units of a commutative monoid.
We have chosen to define rootsOfUnity n for n : ℕ+, instead of n : ℕ,
because almost all lemmas need the positivity assumption,
and in particular the type class instances for Fintype and IsCyclic.
On the other hand, for primitive roots of unity, it is desirable to have a predicate
not just on units, but directly on elements of the ring/field.
For example, we want to say that exp (2 * pi * I / n) is a primitive n-th root of unity
in the complex numbers, without having to turn that number into a unit first.
This creates a little bit of friction, but lemmas like IsPrimitiveRoot.isUnit and
IsPrimitiveRoot.coe_units_iff should provide the necessary glue.
rootsOfUnity k M is the subgroup of elements m : Mˣ that satisfy m ^ k = 1.
Equations
- rootsOfUnity k M = { carrier := {ζ : Mˣ | ζ ^ ↑k = 1}, mul_mem' := ⋯, one_mem' := ⋯, inv_mem' := ⋯ }
Instances For
Make an element of rootsOfUnity from a member of the base ring, and a proof that it has
a positive power equal to one.
Equations
- rootsOfUnity.mkOfPowEq ζ h = ⟨Units.ofPowEqOne ζ (↑n) h ⋯, ⋯⟩
Instances For
Restrict a ring homomorphism to the nth roots of unity.
Equations
- restrictRootsOfUnity σ n = let h := ⋯; { toFun := fun (ξ : ↥(rootsOfUnity n R)) => ⟨unitOfInvertible (σ ↑↑ξ), ⋯⟩, map_one' := ⋯, map_mul' := ⋯ }
Instances For
Restrict a monoid isomorphism to the nth roots of unity.
Equations
- σ.restrictRootsOfUnity n = { toFun := ⇑(restrictRootsOfUnity σ n), invFun := ⇑(restrictRootsOfUnity σ.symm n), left_inv := ⋯, right_inv := ⋯, map_mul' := ⋯ }
Instances For
Equivalence between the k-th roots of unity in R and the k-th roots of 1.
This is implemented as equivalence of subtypes,
because rootsOfUnity is a subgroup of the group of units,
whereas nthRoots is a multiset.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- rootsOfUnity.fintype R k = Fintype.ofEquiv { x : R // x ∈ Polynomial.nthRoots (↑k) 1 } (rootsOfUnityEquivNthRoots R k).symm
Equations
- ⋯ = ⋯
Turn a primitive root μ into a member of the rootsOfUnity subgroup.
Equations
- h.toRootsOfUnity = rootsOfUnity.mkOfPowEq μ ⋯
Instances For
primitiveRoots k R is the finset of primitive k-th roots of unity
in the integral domain R.
Equations
- primitiveRoots k R = Finset.filter (fun (ζ : R) => IsPrimitiveRoot ζ k) (Polynomial.nthRoots k 1).toFinset
Instances For
If there is an n-th primitive root of unity in R and b divides n,
then there is a b-th primitive root of unity in R.
If 1 < k then (∑ i ∈ range k, ζ ^ i) = 0.
If 1 < k, then ζ ^ k.pred = -(∑ i ∈ range k.pred, ζ ^ i).
The (additive) monoid equivalence between ZMod k
and the powers of a primitive root of unity ζ.
Equations
- One or more equations did not get rendered due to their size.
Instances For
If R contains a n-th primitive root, and S/R is a ring extension,
then the n-th roots of unity in R and S are isomorphic.
Also see IsPrimitiveRoot.map_rootsOfUnity for the equality as Subgroup Sˣ.
Equations
- rootsOfUnityEquivOfPrimitiveRoots hf hζ = ((rootsOfUnity n R).equivMapOfInjective (Units.map ↑f) ⋯).trans (MulEquiv.subgroupCongr ⋯)
Instances For
A version of IsPrimitiveRoot.eq_pow_of_mem_rootsOfUnity that takes a natural number k
as argument instead of a PNat (and ζ : R instead of ζ : Rˣ).
The cardinality of the multiset nthRoots ↑n (1 : R) is n
if there is a primitive root of unity in R.
The multiset nthRoots ↑n (1 : R) has no repeated elements
if there is a primitive root of unity in R.
If an integral domain has a primitive k-th root of unity, then it has φ k of them.
The sets primitiveRoots k R are pairwise disjoint.
nthRoots n as a Finset is equal to the union of primitiveRoots i R for i ∣ n
if there is a primitive root of unity in R.
This holds for any Nat, not just PNat, see nthRoots_one_eq_bUnion_primitive_roots.
nthRoots n as a Finset is equal to the union of primitiveRoots i R for i ∣ n
if there is a primitive root of unity in R.
The MonoidHom that takes an automorphism to the power of μ that μ gets mapped to under it.
Equations
- IsPrimitiveRoot.autToPow R hμ = let μ' := hμ.toRootsOfUnity; let_fun ho := ⋯; { toFun := fun (σ : S ≃ₐ[R] S) => ↑⋯.choose, map_one' := ⋯, map_mul' := ⋯ }.toHomUnits