Skip to content
This repository has been archived by the owner on Jul 24, 2024. It is now read-only.

old WIP trying to replace fractional_ideal by submodule in class_group: to be ported #19245

Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
27 changes: 21 additions & 6 deletions src/algebra/algebra/operations.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/algebra/algebra/operations.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.

Check warning on line 1 in src/algebra/algebra/operations.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2019 Kenny Lau. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Kenny Lau
Expand Down Expand Up @@ -325,7 +325,15 @@
..(by apply_instance : order_bot (submodule R A)),
..(by apply_instance : lattice (submodule R A)) }

variables (M)
variable R

def span_singleton : A →*₀ submodule R A :=
{ to_fun := λ x, span R {x},
map_zero' := span_zero_singleton R,
map_one' := one_eq_span.symm,
map_mul' := λ x y, by rw [span_mul_span, set.singleton_mul_singleton] }

variables {R} (M)

lemma span_pow (s : set A) : ∀ n : ℕ, span R s ^ n = span R (s ^ n)
| 0 := by rw [pow_zero, pow_zero, one_eq_span_one_set]
Expand Down Expand Up @@ -418,14 +426,21 @@
submodule.pow_induction_on_right' M
(by exact hr) (λ x y i hx hy, hadd x y) (λ i x hx, hmul _) hx

/-- `submonoid.map` as a `monoid_with_zero_hom`, when applied to `alg_hom`s. -/
/-- `submodule.map` as a `ring_hom`, when applied to an `alg_hom`. -/
@[simps]
def map_hom {A'} [semiring A'] [algebra R A'] (f : A →ₐ[R] A') :
submodule R A →*₀ submodule R A' :=
{ to_fun := map f.to_linear_map,
map_zero' := submodule.map_bot _,
submodule R A →+* submodule R A' :=
{ to_fun := map f,
map_zero' := map_bot _,
map_one' := submodule.map_one _,
map_mul' := λ _ _, submodule.map_mul _ _ _}
map_add' := (add_hom f).map_add,
map_mul' := λ _ _, submodule.map_mul _ _ _ }

/-- `submodule.map` as a `ring_equiv`, when applied to an `alg_equiv`. -/
@[simps]
def map_equiv {A'} [semiring A'] [algebra R A'] (f : A ≃ₐ[R] A') :
submodule R A ≃+* submodule R A' :=
{ map_mul' := λ _ _, submodule.map_mul _ _ f.to_alg_hom, .. add_equiv f.to_linear_equiv }

/-- The ring of submodules of the opposite algebra is isomorphic to the opposite ring of
submodules. -/
Expand Down
10 changes: 10 additions & 0 deletions src/algebra/module/equiv.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/algebra/module/equiv.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.

Check warning on line 1 in src/algebra/module/equiv.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2020 Anne Baanen. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Baanen,
Expand Down Expand Up @@ -102,6 +102,16 @@
coe_injective' := @fun_like.coe_injective F _ _ _,
..s }

variable {F}
def to_linear_equiv [ring_hom_inv_pair σ σ'] [ring_hom_inv_pair σ' σ]
[semilinear_equiv_class F σ M M₂] (f : F) : M ≃ₛₗ[σ] M₂ :=
{ to_fun := f,
map_add' := map_add f,
map_smul' := map_smulₛₗ f,
inv_fun := inv f,
left_inv := left_inv f,
right_inv := right_inv f }

end semilinear_equiv_class

namespace linear_equiv
Expand Down
23 changes: 23 additions & 0 deletions src/algebra/module/submodule/pointwise.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/algebra/module/submodule/pointwise.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.

Check warning on line 1 in src/algebra/module/submodule/pointwise.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2021 Eric Wieser. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Eric Wieser
Expand Down Expand Up @@ -149,6 +149,29 @@
..submodule.pointwise_add_comm_monoid,
..submodule.complete_lattice }

section hom

variables {R₂ M₂ F : Type*} [semiring R₂] [add_comm_monoid M₂] [module R₂ M₂]
{σ₁₂ : R →+* R₂} [ring_hom_surjective σ₁₂]

/-- `submodule.map` as an `add_hom`. -/
@[simps] def add_hom [semilinear_map_class F σ₁₂ M M₂] (f : F) : submodule R M →+ submodule R₂ M₂ :=
{ to_fun := map f,
map_zero' := map_bot f,
map_add' := by { intros, simp_rw add_eq_sup, apply map_sup } }

/-- `submodule.map` as an `add_equiv`, when applied to a `linear_equiv`.
TODO: allow `semilinear_equiv_class`. -/
@[simps] def add_equiv {σ₂₁ : R₂ →+* R} [ring_hom_inv_pair σ₁₂ σ₂₁] [ring_hom_inv_pair σ₂₁ σ₁₂]
(f : M ≃ₛₗ[σ₁₂] M₂) : submodule R M ≃+ submodule R₂ M₂ :=
{ to_fun := map f,
inv_fun := map f.symm,
left_inv := λ N, by rw map_symm_eq_iff,
right_inv := λ N, by rw ← map_symm_eq_iff,
map_add' := (add_hom f).map_add }

end hom

section
variables [monoid α] [distrib_mul_action α M] [smul_comm_class α R M]

Expand Down
29 changes: 18 additions & 11 deletions src/linear_algebra/basic.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/linear_algebra/basic.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.

Check warning on line 1 in src/linear_algebra/basic.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2017 Johannes Hölzl. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov, Frédéric Dupuis,
Expand Down Expand Up @@ -1996,7 +1996,9 @@

namespace submodule

variables [comm_semiring R] [comm_semiring R₂]
section

variables [semiring R] [semiring R₂]
variables [add_comm_monoid M] [add_comm_monoid M₂] [module R M] [module R₂ M₂]
variables [add_comm_monoid N] [add_comm_monoid N₂] [module R N] [module R N₂]
variables {τ₁₂ : R →+* R₂} {τ₂₁ : R₂ →+* R}
Expand Down Expand Up @@ -2043,16 +2045,6 @@
p.comap_equiv_eq_map_symm _
omit τ₂₁

lemma comap_le_comap_smul (fₗ : N →ₗ[R] N₂) (c : R) :
comap fₗ qₗ ≤ comap (c • fₗ) qₗ :=
begin
rw set_like.le_def,
intros m h,
change c • (fₗ m) ∈ qₗ,
change fₗ m ∈ qₗ at h,
apply qₗ.smul_mem _ h,
end

lemma inf_comap_le_comap_add (f₁ f₂ : M →ₛₗ[τ₁₂] M₂) :
comap f₁ q ⊓ comap f₂ q ≤ comap (f₁ + f₂) q :=
begin
Expand All @@ -2063,6 +2055,21 @@
apply q.add_mem h.1 h.2,
end

end

variables [comm_semiring R] [add_comm_monoid N] [add_comm_monoid N₂] [module R N] [module R N₂]
variables (pₗ : submodule R N) (qₗ : submodule R N₂)

lemma comap_le_comap_smul (fₗ : N →ₗ[R] N₂) (c : R) :
comap fₗ qₗ ≤ comap (c • fₗ) qₗ :=
begin
rw set_like.le_def,
intros m h,
change c • (fₗ m) ∈ qₗ,
change fₗ m ∈ qₗ at h,
apply qₗ.smul_mem _ h,
end

/-- Given modules `M`, `M₂` over a commutative ring, together with submodules `p ⊆ M`, `q ⊆ M₂`,
the set of maps $\{f ∈ Hom(M, M₂) | f(p) ⊆ q \}$ is a submodule of `Hom(M, M₂)`. -/
def compatible_maps : submodule R (N →ₗ[R] N₂) :=
Expand Down
47 changes: 26 additions & 21 deletions src/ring_theory/class_group.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/ring_theory/class_group.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.

Check warning on line 1 in src/ring_theory/class_group.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2021 Anne Baanen. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Anne Baanen
Expand Down Expand Up @@ -36,43 +36,40 @@

open_locale non_zero_divisors

open is_localization is_fraction_ring fractional_ideal units
open is_localization is_fraction_ring submodule units

section

variables (R K)

/-- The group of invertible fractional `R`-ideals in `K` is isomorphic to the group of
invertible `R`-submodules of `K`. We will use the latter henceforth for simplicity. -/
example : (fractional_ideal R⁰ K)ˣ ≃* (submodule R K)ˣ := fractional_ideal.unit_equiv

/-- `to_principal_ideal R K x` sends `x ≠ 0 : K` to the fractional `R`-ideal generated by `x` -/
@[irreducible]
def to_principal_ideal : Kˣ →* (fractional_ideal R⁰ K)ˣ :=
{ to_fun := λ x,
⟨span_singleton _ x,
span_singleton _ x⁻¹,
by simp only [span_singleton_one, units.mul_inv', span_singleton_mul_span_singleton],
by simp only [span_singleton_one, units.inv_mul', span_singleton_mul_span_singleton]⟩,
map_mul' := λ x y, ext
(by simp only [units.coe_mk, units.coe_mul, span_singleton_mul_span_singleton]),
map_one' := ext (by simp only [span_singleton_one, units.coe_mk, units.coe_one]) }
def to_principal_ideal : Kˣ →* (submodule R K)ˣ := units.map (span_singleton R).to_monoid_hom

variables {R K}

@[simp] lemma coe_to_principal_ideal (x : Kˣ) :
(to_principal_ideal R K x : fractional_ideal R⁰ K) = span_singleton _ x :=
(to_principal_ideal R K x : submodule R K) = span R {x} :=
by { simp only [to_principal_ideal], refl }

@[simp] lemma to_principal_ideal_eq_iff {I : (fractional_ideal R⁰ K)ˣ} {x : Kˣ} :
to_principal_ideal R K x = I ↔ span_singleton R⁰ (x : K) = I :=
@[simp] lemma to_principal_ideal_eq_iff {I : (submodule R K)ˣ} {x : Kˣ} :
to_principal_ideal R K x = I ↔ span R ({x} : set K) = I :=
by { simp only [to_principal_ideal], exact units.ext_iff }

lemma mem_principal_ideals_iff {I : (fractional_ideal R⁰ K)ˣ} :
I ∈ (to_principal_ideal R K).range ↔ ∃ x : K, span_singleton R⁰ x = I :=
lemma mem_principal_ideals_iff {I : (submodule R K)ˣ} :
I ∈ (to_principal_ideal R K).range ↔ ∃ x : K, span R ({x} : set K) = I :=
begin
simp only [monoid_hom.mem_range, to_principal_ideal_eq_iff],
split; rintros ⟨x, hx⟩,
{ exact ⟨x, hx⟩ },
{ refine ⟨units.mk0 x _, hx⟩,
rintro rfl,
simpa [I.ne_zero.symm] using hx }
apply I.ne_zero.symm,
simpa using hx }
end

instance principal_ideals.normal : (to_principal_ideal R K).range.normal :=
Expand All @@ -86,18 +83,26 @@
modulo the principal ideals. -/
@[derive comm_group]
def class_group :=
(fractional_ideal R⁰ (fraction_ring R))ˣ ⧸ (to_principal_ideal R (fraction_ring R)).range
(submodule R (fraction_ring R))ˣ ⧸ (to_principal_ideal R (fraction_ring R)).range

noncomputable instance : inhabited (class_group R) := ⟨1⟩

variables {R K}

/-- Send a nonzero fractional ideal to the corresponding class in the class group. -/
noncomputable def class_group.mk : (fractional_ideal R⁰ K)ˣ →* class_group R :=
(quotient_group.mk' (to_principal_ideal R (fraction_ring R)).range).comp
(units.map (fractional_ideal.canonical_equiv R⁰ K (fraction_ring R)))
noncomputable def class_group.mk : (submodule R K)ˣ →* class_group R :=
(quotient_group.mk' _).comp $ mul_equiv.to_monoid_hom $ units.map_equiv $
ring_equiv.to_mul_equiv $ submodule.map_equiv $ is_localization.alg_equiv R⁰ _ _

lemma class_group.mk_apply (I : (submodule R $ fraction_ring R)ˣ) :
class_group.mk I = quotient_group.mk' _ I :=
begin
rw [class_group.mk, monoid_hom.comp_apply],
congr, ext1, squeeze_simp,
convert submodule.map_id I,
end

lemma class_group.mk_eq_mk {I J : (fractional_ideal R⁰ $ fraction_ring R)ˣ} :
lemma class_group.mk_eq_mk {I J : (submodule R $ fraction_ring R)ˣ} :
class_group.mk I = class_group.mk J
↔ ∃ x : (fraction_ring R)ˣ, I * to_principal_ideal R (fraction_ring R) x = J :=
by { erw [quotient_group.mk'_eq_mk', canonical_equiv_self, units.map_id, set.exists_range_iff],
Expand Down
42 changes: 38 additions & 4 deletions src/ring_theory/fractional_ideal.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/ring_theory/fractional_ideal.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.

Check warning on line 1 in src/ring_theory/fractional_ideal.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2020 Anne Baanen. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Anne Baanen, Filippo A. E. Nuccio
Expand Down Expand Up @@ -78,17 +78,19 @@

variables (S)

/-- A submodule `I` is a fractional ideal if `a I ⊆ R` for some `a ≠ 0`. -/
/-- A submodule `I` is a fractional ideal with respect to a submonoid `S`
if `a I ⊆ R` for some `a ∈ S`. -/
def is_fractional (I : submodule R P) :=
∃ a ∈ S, ∀ b ∈ I, is_integer R (a • b)

variables (S P)

/-- The fractional ideals of a domain `R` are ideals of `R` divided by some `a ∈ R`.
/-- The fractional ideals of a domain `R` with respect to a submonoid `S`
are ideals of `R` divided by some `a ∈ S`.

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`.
such that there is an `a ∈ S` with `a I ⊆ R`.
-/
def fractional_ideal :=
{I : submodule R P // is_fractional S I}
Expand Down Expand Up @@ -691,7 +693,7 @@

include loc

lemma is_fractional_of_fg {I : submodule R P} (hI : I.fg) :
lemma _root_.is_fractional_of_fg {I : submodule R P} (hI : I.fg) :
is_fractional S I :=
begin
rcases hI with ⟨I, rfl⟩,
Expand All @@ -700,6 +702,29 @@
exact ⟨s, hs1, hs⟩,
end

lemma _root_.is_fractional_unit (I : (submodule R P)ˣ) : is_fractional S (I : submodule R P) :=
is_fractional_of_fg $ fg_unit I

lemma _root_.is_fractional_of_is_unit {I : submodule R P} (hI : is_unit I) : is_fractional S I :=
is_fractional_of_fg $ fg_of_is_unit hI

/-- The group of invertible fractional `R`-ideals in `P` is isomorphic to
the group of invertible `R`-submodules of `P`. -/
def unit_equiv : (fractional_ideal S P)ˣ ≃* (submodule R P)ˣ :=
{ to_fun := λ I, by use [I, ↑I⁻¹];
simp only [coe_coe, ← fractional_ideal.coe_mul, I.mul_inv, I.inv_mul, fractional_ideal.coe_one],
inv_fun := λ I, begin
use [⟨I, is_fractional_unit I⟩, ⟨_, is_fractional_unit I⁻¹⟩];
apply subtype.ext;
simp only [fractional_ideal.coe_mul, fractional_ideal.coe_mk,
I.mul_inv, I.inv_mul, fractional_ideal.coe_one],
end,
left_inv := λ I, by { ext1, ext1, refl },
right_inv := λ I, by { ext1, refl },
map_mul' := λ I J, begin
ext1, simp only [coe_coe, fractional_ideal.coe_mul, units.coe_mul, units.coe_mk],
end }

omit loc

lemma mem_span_mul_finite_of_mem_mul {I J : fractional_ideal S P} {x : P} (hx : x ∈ I * J) :
Expand Down Expand Up @@ -731,6 +756,15 @@

include loc loc'

@[irreducible]
noncomputable def canonical_equiv :
fractional_ideal S P ≃+* fractional_ideal S P' :=
submodule.map_equiv
{ commutes' := λ r, ring_equiv_of_ring_equiv_eq _ _,
..ring_equiv_of_ring_equiv P P' (ring_equiv.refl R)
(show S.map _ = S, by rw [ring_equiv.to_monoid_hom_refl, submonoid.map_id]) }


/-- `canonical_equiv f f'` is the canonical equivalence between the fractional
ideals in `P` and in `P'` -/
@[irreducible]
Expand Down
4 changes: 2 additions & 2 deletions src/ring_theory/localization/submodule.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
/-

Check warning on line 1 in src/ring_theory/localization/submodule.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.

Check warning on line 1 in src/ring_theory/localization/submodule.lean

View workflow job for this annotation

GitHub Actions / Check for modifications to ported files

Changes to this file will need to be ported to mathlib 4! Please consider retracting the changes to this file unless you are willing to immediately forward-port them.
Copyright (c) 2018 Kenny Lau. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baanen
Expand All @@ -18,8 +18,8 @@
localization, ring localization, commutative ring localization, characteristic predicate,
commutative ring, field of fractions
-/
variables {R : Type*} [comm_ring R] (M : submonoid R) (S : Type*) [comm_ring S]
variables [algebra R S] {P : Type*} [comm_ring P]
variables {R : Type*} [comm_semiring R] (M : submonoid R) (S : Type*) [comm_semiring S]
variables [algebra R S] {P : Type*} [comm_semiring P]

namespace is_localization

Expand Down
Loading