Fractional ideals #
This file defines fractional ideals of an integral domain and proves basic facts about them.
Main definitions #
Let S
be a submonoid of an integral domain R
and P
the localization of R
at S
.
IsFractional
defines whichR
-submodules ofP
are fractional idealsFractionalIdeal S P
is the type of fractional ideals inP
- a coercion
coeIdeal : Ideal R → FractionalIdeal S P
CommSemiring (FractionalIdeal S P)
instance: the typical ideal operations generalized to fractional idealsLattice (FractionalIdeal S P)
instance
Main statements #
mul_left_mono
andmul_right_mono
state that ideal multiplication is monotonemul_div_self_cancel_iff
states that1 / I
is the inverse ofI
if one exists
Implementation notes #
Fractional ideals are considered equal when they contain the same elements,
independent of the denominator a : R
such that a I ⊆ R
.
Thus, we define FractionalIdeal
to be the subtype of the predicate IsFractional
,
instead of having FractionalIdeal
be a structure of which a
is a field.
Most definitions in this file specialize operations from submodules to fractional ideals,
proving that the result of this operation is fractional if the input is fractional.
Exceptions to this rule are defining (+) := (⊔)
and ⊥ := 0
,
in order to re-use their respective proof terms.
We can still use simp
to show ↑I + ↑J = ↑(I + J)
and ↑⊥ = ↑0
.
Many results in fact do not need that P
is a localization, only that P
is an
R
-algebra. We omit the IsLocalization
parameter whenever this is practical.
Similarly, we don't assume that the localization is a field until we need it to
define ideal quotients. When this assumption is needed, we replace S
with R⁰
,
making the localization a field.
References #
Tags #
fractional ideal, fractional ideals, invertible ideal
A submodule I
is a fractional ideal if a I ⊆ R
for some a ≠ 0
.
Equations
- IsFractional S I = ∃ a ∈ S, ∀ b ∈ I, IsLocalization.IsInteger R (a • b)
Instances For
The fractional ideals of a domain R
are ideals of R
divided by some a ∈ R
.
More precisely, let P
be a localization of R
at some submonoid S
,
then a fractional ideal I ⊆ P
is an R
-submodule of P
,
such that there is a nonzero a : R
with a I ⊆ R
.
Equations
- FractionalIdeal S P = { I : Submodule R P // IsFractional S I }
Instances For
Map a fractional ideal I
to a submodule by forgetting that ∃ a, a I ⊆ R
.
This implements the coercion FractionalIdeal S P → Submodule R P
.
Equations
- ↑I = ↑I
Instances For
Map a fractional ideal I
to a submodule by forgetting that ∃ a, a I ⊆ R
.
This coercion is typically called coeToSubmodule
in lemma names
(or coe
when the coercion is clear from the context),
not to be confused with IsLocalization.coeSubmodule : Ideal R → Submodule R P
(which we use to define coe : Ideal R → FractionalIdeal S P
).
Equations
- FractionalIdeal.instCoeOutSubmodule = { coe := FractionalIdeal.coeToSubmodule }
An element of S
such that I.den • I = I.num
, see FractionalIdeal.num
and
FractionalIdeal.den_mul_self_eq_num
.
Equations
- I.den = ⟨Exists.choose ⋯, ⋯⟩
Instances For
An ideal of R
such that I.den • I = I.num
, see FractionalIdeal.den
and
FractionalIdeal.den_mul_self_eq_num
.
Equations
- I.num = Submodule.comap (Algebra.linearMap R P) (I.den • ↑I)
Instances For
The linear equivalence between the fractional ideal I
and the integral ideal I.num
defined by mapping x
to den I • x
.
Equations
- FractionalIdeal.equivNum h_nz = (LinearEquiv.ofBijective ((DistribMulAction.toLinearMap R P I.den).restrict ⋯) ⋯).trans (Submodule.equivMapOfInjective (Algebra.linearMap R P) ⋯ I.num).symm
Instances For
Equations
- FractionalIdeal.instSetLike = { coe := fun (I : FractionalIdeal S P) => ↑↑I, coe_injective' := ⋯ }
Copy of a FractionalIdeal
with a new underlying set equal to the old one.
Useful to fix definitional equalities.
Equations
- p.copy s hs = ⟨(↑p).copy s hs, ⋯⟩
Instances For
Transfer instances from Submodule R P
to FractionalIdeal S P
.
Map an ideal I
to a fractional ideal by forgetting I
is integral.
This is the function that implements the coercion Ideal R → FractionalIdeal S P
.
Equations
- ↑I = ⟨IsLocalization.coeSubmodule P I, ⋯⟩
Instances For
Map an ideal I
to a fractional ideal by forgetting I
is integral.
This is a bundled version of IsLocalization.coeSubmodule : Ideal R → Submodule R P
,
which is not to be confused with the coe : FractionalIdeal S P → Submodule R P
,
also called coeToSubmodule
in theorem names.
This map is available as a ring hom, called FractionalIdeal.coeIdealHom
.
Equations
- FractionalIdeal.instZero S = { zero := ↑0 }
(1 : FractionalIdeal S P)
is defined as the R-submodule f(R) ≤ P
.
However, this is not definitionally equal to 1 : Submodule R P
,
which is proved in the actual simp
lemma coe_one
.
Lattice
section #
Defines the order on fractional ideals as inclusion of their underlying sets, and ports the lattice structure on submodules to fractional ideals.
Equations
- FractionalIdeal.orderBot = OrderBot.mk ⋯
Equations
- FractionalIdeal.instInf = { inf := fun (I J : FractionalIdeal S P) => ⟨↑I ⊓ ↑J, ⋯⟩ }
Equations
- FractionalIdeal.instSup = { sup := fun (I J : FractionalIdeal S P) => ⟨↑I ⊔ ↑J, ⋯⟩ }
Equations
- FractionalIdeal.lattice = Function.Injective.lattice (fun (a : { I : Submodule R P // IsFractional S I }) => ↑a) ⋯ ⋯ ⋯
Equations
- FractionalIdeal.instSemilatticeSup = SemilatticeSup.mk ⋯ ⋯ ⋯
Equations
- FractionalIdeal.instAdd = { add := fun (x1 x2 : FractionalIdeal S P) => x1 ⊔ x2 }
Equations
- FractionalIdeal.instSMulNat = { smul := fun (n : ℕ) (I : FractionalIdeal S P) => ⟨n • ↑I, ⋯⟩ }
FractionalIdeal.mul
is the product of two fractional ideals,
used to define the Mul
instance.
This is only an auxiliary definition: the preferred way of writing I.mul J
is I * J
.
Elaborated terms involving FractionalIdeal
tend to grow quite large,
so by making definitions irreducible, we hope to avoid deep unfolds.
Equations
Instances For
Equations
- FractionalIdeal.instMul = { mul := fun (I J : FractionalIdeal S P) => I.mul J }
Equations
- FractionalIdeal.instPowNat = { pow := fun (I : FractionalIdeal S P) (n : ℕ) => ⟨↑I ^ n, ⋯⟩ }
Equations
- FractionalIdeal.commSemiring = Function.Injective.commSemiring (fun (a : { I : Submodule R P // IsFractional S I }) => ↑a) ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
FractionalIdeal.coeToSubmodule
as a bundled RingHom
.
Equations
- FractionalIdeal.coeSubmoduleHom S P = { toFun := FractionalIdeal.coeToSubmodule, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
coeIdealHom (S : Submonoid R) P
is (↑) : Ideal R → FractionalIdeal S P
as a ring hom
Equations
- FractionalIdeal.coeIdealHom S P = { toFun := FractionalIdeal.coeIdeal, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
The fractional ideals of a Noetherian ring are finitely generated.