mathlib documentation

ring_theory.dedekind_domain

Dedekind domains

This file defines the notion of a Dedekind domain (or Dedekind ring), giving three equivalent definitions (TODO: and shows that they are equivalent).

Main definitions

Implementation notes

The definitions that involve a field of fractions choose a canonical field of fractions, but are independent of that choice. The ..._iff lemmas express this independence.

References

Tags

dedekind domain, dedekind ring

def ring.dimension_le_one (R : Type u_1) [comm_ring R] :
Prop

A ring R has Krull dimension at most one if all nonzero prime ideals are maximal.

Equations
@[class]
structure is_dedekind_domain (A : Type u_2) [integral_domain A] :
Prop

A Dedekind domain is an integral domain that is Noetherian, integrally closed, and has Krull dimension exactly one (not_is_field and dimension_le_one).

The integral closure condition is independent of the choice of field of fractions: use is_dedekind_domain_iff to prove is_dedekind_domain for a given fraction_map.

This is the default implementation, but there are equivalent definitions, is_dedekind_domain_dvr and is_dedekind_domain_inv. TODO: Prove that these are actually equivalent definitions.

An integral domain is a Dedekind domain iff and only if it is not a field, is Noetherian, has dimension ≤ 1, and is integrally closed in a given fraction field. In particular, this definition does not depend on the choice of this fraction field.

structure is_dedekind_domain_dvr (A : Type u_2) [integral_domain A] :
Prop

A Dedekind domain is an integral domain that is not a field, is Noetherian, and the localization at every nonzero prime is a discrete valuation ring.

This is equivalent to is_dedekind_domain. TODO: prove the equivalence.

structure is_dedekind_domain_inv (A : Type u_2) [integral_domain A] :
Prop

A Dedekind domain is an integral domain that is not a field such that every fractional ideal has an inverse.

This is equivalent to is_dedekind_domain. TODO: prove the equivalence.

theorem is_dedekind_domain_inv_iff (A : Type u_2) (K : Type u_3) [integral_domain A] [field K] (f : fraction_map A K) :