Polar coordinate change of variables for the mixed space of a number field #
We define two polar coordinate changes of variables for the mixed space ℝ^r₁ × ℂ^r₂
associated
to a number field K
of signature (r₁, r₂)
. The first one is mixedEmbedding.polarCoord
and has
value in realMixedSpace K
defined as ℝ^r₁ × (ℝ ⨯ ℝ)^r₂
, the second is
mixedEmbedding.polarSpaceCoord
and has value in polarSpace K
defined as ℝ^(r₁+r₂) × ℝ^r₂
.
The change of variables with the polarSpace
is useful to compute the volume of subsets of the
mixed space with enough symmetries, see volume_eq_two_pi_pow_mul_integral
and
volume_eq_two_pow_mul_two_pi_pow_mul_integral
Main definitions and results #
mixedEmbedding.polarCoord
: the polar coordinate change of variables between the mixed spaceℝ^r₁ × ℂ^r₂
andℝ^r₁ × (ℝ × ℝ)^r₂
defined as the identity on the first component and mapping(zᵢ)ᵢ
to(‖zᵢ‖, Arg zᵢ)ᵢ
on the second component.mixedEmbedding.integral_comp_polarCoord_symm
: the change of variables formula formixedEmbedding.polarCoord
mixedEmbedding.polarSpaceCoord
: the polar coordinate change of variables between the mixed spaceℝ^r₁ × ℂ^r₂
and the polar spaceℝ^(r₁ + r₂) × ℝ^r₂
defined by sendingx
tox w
or‖x w‖
depending on wetherw
is real or complex for the first component, and toArg (x w)
,w
complex, for the second component.mixedEmbedding.integral_comp_polarSpaceCoord_symm
: the change of variables formula formixedEmbedding.polarSpaceCoord
mixedEmbedding.volume_eq_two_pi_pow_mul_integral
: if the measurable setA
of the mixed space is norm-stable at complex places in the sense thatnormAtComplexPlaces⁻¹ (normAtComplexPlaces '' A) = A
, then its volume can be computed via an integral overnormAtComplexPlaces '' A
.mixedEmbedding.volume_eq_two_pow_mul_two_pi_pow_mul_integral
: if the measurable setA
of the mixed space is norm-stable in the sense thatnormAtAllPlaces⁻¹ (normAtAllPlaces '' A) = A
, then its volume can be computed via an integral overnormAtAllPlaces '' A
.
The real mixed space ℝ^r₁ × (ℝ × ℝ)^r₂
with (r₁, r₂)
the signature of K
.
Equations
Instances For
The natural homeomorphism between the mixed space ℝ^r₁ × ℂ^r₂
and the real mixed space
ℝ^r₁ × (ℝ × ℝ)^r₂
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The polar coordinate partial homeomorphism of ℝ^r₁ × (ℝ × ℝ)^r₂
defined as the identity on
the first component and mapping (rᵢ cos θᵢ, rᵢ sin θᵢ)ᵢ
to (rᵢ, θᵢ)ᵢ
on the second component.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The derivative of polarCoordReal.symm
, see hasFDerivAt_polarCoordReal_symm
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The polar coordinate partial homeomorphism between the mixed space ℝ^r₁ × ℂ^r₂
and
ℝ^r₁ × (ℝ × ℝ)^r₂
defined as the identity on the first component and mapping (zᵢ)ᵢ
to
(‖zᵢ‖, Arg zᵢ)ᵢ
on the second component.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The space ℝ^(r₁+r₂) × ℝ^r₂
, it is homeomorph to the realMixedSpace
, see
homeoRealMixedSpacePolarSpace
.
Equations
- NumberField.mixedEmbedding.polarSpace K = ((NumberField.InfinitePlace K → ℝ) × ({ w : NumberField.InfinitePlace K // w.IsComplex } → ℝ))
Instances For
The measurable equivalence between the realMixedSpace
and the polarSpace
. It is actually an
homeomorphism, see homeoRealMixedSpacePolarSpace
, but defining it in this way makes it easier
to prove that it is volume preserving, see volume_preserving_homeoRealMixedSpacePolarSpace
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The homeomorphism between the realMixedSpace
and the polarSpace
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The polar coordinate partial homeomorphism between the mixed space ℝ^r₁ × ℂ^r₂
and the polar
space ℝ^(r₁ + r₂) × ℝ^r₂
defined by sending x
to x w
or ‖x w‖
depending on wether w
is
real or complex for the first component, and to Arg (x w)
, w
complex, for the second component.
Equations
Instances For
If the measurable set A
is norm-stable at complex places in the sense that
normAtComplexPlaces⁻¹ (normAtComplexPlaces '' A) = A
, then its volume can be computed via an
integral over normAtComplexPlaces '' A
.
If the measurable set A
is norm-stable in the sense that
normAtAllPlaces⁻¹ (normAtAllPlaces '' A) = A
, then its volume can be computed via an integral
over normAtAllPlaces '' A
.