Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat(RingTheory/Morita/Basic): Define Morita Equivalence #20640

Closed
wants to merge 42 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
42 commits
Select commit Hold shift + click to select a range
ca351a7
basic
Whysoserioushah Jan 10, 2025
b5b09d2
fix build
Whysoserioushah Jan 10, 2025
60c47ac
apply suggestions
jjaassoonn Jan 10, 2025
de69322
inline
Jan 10, 2025
6af581c
add todo
Jan 10, 2025
a66deb6
add 1
Jan 13, 2025
ab44429
change
Jan 13, 2025
2e80c78
Update Mathlib/RingTheory/Morita/Basic.lean
jjaassoonn Jan 13, 2025
90c9112
shorten line
Jan 15, 2025
f8bb3f4
Merge remote-tracking branch 'origin/master' into morita
Whysoserioushah Jan 16, 2025
273aaa7
generalize
Jan 18, 2025
13696f4
test
Jan 18, 2025
7d58c66
fix?
Jan 18, 2025
f79a4f2
fix
Jan 18, 2025
d7d9043
Update Mathlib/RingTheory/Morita/Basic.lean
jjaassoonn Jan 22, 2025
d1f7a09
Update Mathlib/RingTheory/Morita/Basic.lean
Whysoserioushah Jan 24, 2025
8a33222
Update Mathlib/RingTheory/Morita/Basic.lean
Whysoserioushah Jan 24, 2025
68ba2a3
Update Mathlib/RingTheory/Morita/Basic.lean
Whysoserioushah Jan 24, 2025
f0552fe
Update Mathlib/RingTheory/Morita/Basic.lean
Whysoserioushah Jan 25, 2025
2344c2a
Update Mathlib/RingTheory/Morita/Basic.lean
Whysoserioushah Jan 25, 2025
d2f7b46
Update Mathlib/Algebra/Category/ModuleCat/Basic.lean
jjaassoonn Jan 25, 2025
961d207
Merge remote-tracking branch 'origin/master' into morita
Whysoserioushah Feb 3, 2025
5ded14d
junyan suggestion
Whysoserioushah Feb 3, 2025
99a4bf2
Update Mathlib/RingTheory/Morita/Basic.lean
jjaassoonn Feb 3, 2025
6c34832
Update Mathlib/RingTheory/Morita/Basic.lean
jjaassoonn Feb 3, 2025
dffcf77
Update Mathlib/RingTheory/Morita/Basic.lean
jjaassoonn Feb 3, 2025
54cac5f
apply suggestions from @alreadydone
Feb 3, 2025
c7b5292
Remove Functor.Additive from def of Functor.Linear
alreadydone Feb 3, 2025
e556403
unused arguments
alreadydone Feb 4, 2025
47ce0e0
golf `restrictScalarsEquivalenceOfRingEquiv_linear`
alreadydone Feb 4, 2025
50cb794
Update Mathlib/RingTheory/Morita/Basic.lean
jjaassoonn Feb 4, 2025
bf8230e
remove ext
Feb 4, 2025
dde0e8b
Merge branch 'morita' of https://github.com/leanprover-community/math…
Feb 4, 2025
d046381
revert
Feb 4, 2025
e914cb1
Update Mathlib/RingTheory/Morita/Basic.lean
Whysoserioushah Feb 5, 2025
8717ed5
Update Mathlib/RingTheory/Morita/Basic.lean
Whysoserioushah Feb 5, 2025
e1f07fb
Merge remote-tracking branch 'origin/master' into morita
Whysoserioushah Feb 8, 2025
6244abb
Merge branch 'morita' of https://github.com/leanprover-community/math…
Whysoserioushah Feb 8, 2025
d3807df
apply kevin's suggestion
Whysoserioushah Feb 8, 2025
eb6ad79
apply andrew advice
Whysoserioushah Feb 12, 2025
fb51470
fix?
Whysoserioushah Feb 12, 2025
84846a1
fix large import
Whysoserioushah Feb 12, 2025
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
1 change: 1 addition & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4700,6 +4700,7 @@ import Mathlib.RingTheory.Localization.NumDen
import Mathlib.RingTheory.Localization.Pi
import Mathlib.RingTheory.Localization.Submodule
import Mathlib.RingTheory.MatrixAlgebra
import Mathlib.RingTheory.Morita.Basic
import Mathlib.RingTheory.Multiplicity
import Mathlib.RingTheory.MvPolynomial
import Mathlib.RingTheory.MvPolynomial.Basic
Expand Down
35 changes: 33 additions & 2 deletions Mathlib/Algebra/Category/ModuleCat/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -432,11 +432,40 @@ end Module

section

variable {S : Type u} [CommRing S]
universe u₀

namespace Algebra

variable {S₀ : Type u₀} [CommSemiring S₀] {S : Type u} [Ring S] [Algebra S₀ S]

variable {M N : ModuleCat.{v} S}

instance : Linear S (ModuleCat.{v} S) where
/--
Let `S` be an `S₀`-algebra. Then `S`-modules are modules over `S₀`.
-/
scoped instance : Module S₀ M := Module.compHom _ (algebraMap S₀ S)

scoped instance : IsScalarTower S₀ S M where
smul_assoc _ _ _ := by rw [Algebra.smul_def, mul_smul]; rfl

scoped instance : SMulCommClass S S₀ M where
smul_comm s s₀ n :=
show s • algebraMap S₀ S s₀ • n = algebraMap S₀ S s₀ • s • n by
rw [← smul_assoc, smul_eq_mul, ← Algebra.commutes, mul_smul]
Comment on lines +451 to +454
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

You should give IsScalarTower instead. And use IsScalarTower.of_algebraMap_smul in the proof.

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

changed!

Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

It still gives a large-import label. Perhaps add both IsScalarTower and SMulCommClass by hand?


/--
Let `S` be an `S₀`-algebra. Then the category of `S`-modules is `S₀`-linear.
-/
scoped instance instLinear : Linear S₀ (ModuleCat.{v} S) where
smul_comp _ M N s₀ f g := by ext; simp

end Algebra

section

variable {S : Type u} [CommRing S]

instance : Linear S (ModuleCat.{v} S) := ModuleCat.Algebra.instLinear

variable {X Y X' Y' : ModuleCat.{v} S}

Expand All @@ -450,6 +479,8 @@ theorem Iso.conj_eq_conj (i : X ≅ X') (f : End X) :

end

end

variable (M N : ModuleCat.{v} R)

/-- `ModuleCat.Hom.hom` as an isomorphism of rings. -/
Expand Down
19 changes: 19 additions & 0 deletions Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,6 +8,7 @@ import Mathlib.Algebra.Category.ModuleCat.Colimits
import Mathlib.Algebra.Category.ModuleCat.Limits
import Mathlib.RingTheory.TensorProduct.Basic
import Mathlib.CategoryTheory.Adjunction.Mates
import Mathlib.CategoryTheory.Linear.LinearFunctor

/-!
# Change Of Rings
Expand Down Expand Up @@ -264,6 +265,24 @@ instance restrictScalars_isEquivalence_of_ringEquiv {R S} [Ring R] [Ring S] (e :
(ModuleCat.restrictScalars e.toRingHom).IsEquivalence :=
(restrictScalarsEquivalenceOfRingEquiv e).isEquivalence_functor

instance {R S} [Ring R] [Ring S] (f : R →+* S) : (restrictScalars f).Additive where

instance restrictScalarsEquivalenceOfRingEquiv_additive {R S} [Ring R] [Ring S] (e : R ≃+* S) :
(restrictScalarsEquivalenceOfRingEquiv e).functor.Additive where

namespace Algebra

instance {R₀ R S} [CommSemiring R₀] [Ring R] [Ring S] [Algebra R₀ R] [Algebra R₀ S]
(f : R →ₐ[R₀] S) : (restrictScalars f.toRingHom).Linear R₀ where
map_smul {M N} g r₀ := by ext m; exact congr_arg (· • g.hom m) (f.commutes r₀).symm

instance restrictScalarsEquivalenceOfRingEquiv_linear
{R₀ R S} [CommSemiring R₀] [Ring R] [Ring S] [Algebra R₀ R] [Algebra R₀ S] (e : R ≃ₐ[R₀] S) :
(restrictScalarsEquivalenceOfRingEquiv e.toRingEquiv).functor.Linear R₀ :=
inferInstanceAs ((restrictScalars e.toAlgHom.toRingHom).Linear R₀)

end Algebra

open TensorProduct

variable {R : Type u₁} {S : Type u₂} [CommRing R] [CommRing S] (f : R →+* S)
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/CategoryTheory/Action/Limits.lean
Original file line number Diff line number Diff line change
Expand Up @@ -328,6 +328,6 @@ instance mapAction_preadditive [F.Additive] : (F.mapAction G).Additive where

variable {R : Type*} [Semiring R] [CategoryTheory.Linear R V] [CategoryTheory.Linear R W]

instance mapAction_linear [F.Additive] [F.Linear R] : (F.mapAction G).Linear R where
instance mapAction_linear [F.Linear R] : (F.mapAction G).Linear R where

end CategoryTheory.Functor
10 changes: 5 additions & 5 deletions Mathlib/CategoryTheory/Linear/LinearFunctor.lean
Original file line number Diff line number Diff line change
Expand Up @@ -28,7 +28,7 @@ variable (R : Type*) [Semiring R]

/-- An additive functor `F` is `R`-linear provided `F.map` is an `R`-module morphism. -/
class Functor.Linear {C D : Type*} [Category C] [Category D] [Preadditive C] [Preadditive D]
[Linear R C] [Linear R D] (F : C ⥤ D) [F.Additive] : Prop where
[Linear R C] [Linear R D] (F : C ⥤ D) : Prop where
/-- the functor induces a linear map on morphisms -/
map_smul : ∀ {X Y : C} (f : X ⟶ Y) (r : R), F.map (r • f) = r • F.map f := by aesop_cat

Expand All @@ -40,7 +40,7 @@ section

variable {R}
variable {C D : Type*} [Category C] [Category D] [Preadditive C] [Preadditive D]
[CategoryTheory.Linear R C] [CategoryTheory.Linear R D] (F : C ⥤ D) [Additive F] [Linear R F]
[CategoryTheory.Linear R C] [CategoryTheory.Linear R D] (F : C ⥤ D) [Linear R F]

@[simp]
theorem map_smul {X Y : C} (r : R) (f : X ⟶ Y) : F.map (r • f) = r • F.map f :=
Expand All @@ -53,9 +53,9 @@ theorem map_units_smul {X Y : C} (r : Rˣ) (f : X ⟶ Y) : F.map (r • f) = r
instance : Linear R (𝟭 C) where

instance {E : Type*} [Category E] [Preadditive E] [CategoryTheory.Linear R E] (G : D ⥤ E)
[Additive G] [Linear R G] : Linear R (F ⋙ G) where
[Linear R G] : Linear R (F ⋙ G) where

variable (R)
variable (R) [F.Additive]

/-- `F.mapLinearMap` is an `R`-linear map whose underlying function is `F.map`. -/
@[simps]
Expand Down Expand Up @@ -103,7 +103,7 @@ namespace Equivalence
variable {C D : Type*} [Category C] [Category D] [Preadditive C] [Linear R C] [Preadditive D]
[Linear R D]

instance inverseLinear (e : C ≌ D) [e.functor.Additive] [e.functor.Linear R] :
instance inverseLinear (e : C ≌ D) [e.functor.Linear R] :
e.inverse.Linear R where
map_smul r f := by
apply e.functor.map_injective
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/CategoryTheory/Monoidal/Linear.lean
Original file line number Diff line number Diff line change
Expand Up @@ -50,7 +50,7 @@ instance tensoringRight_linear (X : C) : ((tensoringRight C).obj X).Linear R whe
ensures that the domain is linear monoidal. -/
theorem monoidalLinearOfFaithful {D : Type*} [Category D] [Preadditive D] [Linear R D]
[MonoidalCategory D] [MonoidalPreadditive D] (F : D ⥤ C) [F.Monoidal] [F.Faithful]
[F.Additive] [F.Linear R] : MonoidalLinear R D :=
[F.Linear R] : MonoidalLinear R D :=
{ whiskerLeft_smul := by
intros X Y Z r f
apply F.map_injective
Expand Down
143 changes: 143 additions & 0 deletions Mathlib/RingTheory/Morita/Basic.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,143 @@
/-
Copyright (c) 2025 Jujian Zhang. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jujian Zhang, Yunzhou Xie
-/
import Mathlib.Algebra.Category.ModuleCat.ChangeOfRings
import Mathlib.CategoryTheory.Linear.LinearFunctor
import Mathlib.Algebra.Category.ModuleCat.Basic
import Mathlib.CategoryTheory.Adjunction.Limits

/-!
# Morita equivalence
Two `R`-algebras `A` and `B` are Morita equivalent if the categories of modules over `A` and
`B` are `R`-linearly equivalent. In this file, we prove that Morita equivalence is an equivalence
relation and that isomorphic algebras are Morita equivalent.
# Main definitions
- `MoritaEquivalence R A B`: a structure containing an `R`-linear equivalence of categories between
the module categories of `A` and `B`.
- `IsMoritaEquivalent R A B`: a predicate asserting that `R`-algebras `A` and `B` are Morita
equivalent.
## TODO
- For any ring `R`, `R` and `Matₙ(R)` are Morita equivalent.
- Morita equivalence in terms of projective generators.
- Morita equivalence in terms of full idempotents.
Whysoserioushah marked this conversation as resolved.
Show resolved Hide resolved
- Morita equivalence in terms of existence of an invertible bimodule.
- If `R ≈ S`, then `R` is simple iff `S` is simple.
jjaassoonn marked this conversation as resolved.
Show resolved Hide resolved
## References
* [Nathan Jacobson, *Basic Algebra II*][jacobson1989]
## Tags
Morita Equivalence, Category Theory, Noncommutative Ring, Module Theory
-/

universe u₀ u₁ u₂ u₃

open CategoryTheory

variable (R : Type u₀) [CommSemiring R]

open scoped ModuleCat.Algebra

/--
Let `A` and `B` be `R`-algebras. A Morita equivalence between `A` and `B` is an `R`-linear
equivalence between the categories of `A`-modules and `B`-modules.
-/
structure MoritaEquivalence
(A : Type u₁) [Ring A] [Algebra R A]
(B : Type u₂) [Ring B] [Algebra R B] where
/-- The underlying equivalence of categories -/
eqv : ModuleCat.{max u₁ u₂} A ≌ ModuleCat.{max u₁ u₂} B
linear : eqv.functor.Linear R := by infer_instance

namespace MoritaEquivalence

attribute [instance] MoritaEquivalence.linear

instance {A : Type u₁} [Ring A] [Algebra R A] {B : Type u₂} [Ring B] [Algebra R B]
(e : MoritaEquivalence R A B) : e.eqv.functor.Additive :=
e.eqv.functor.additive_of_preserves_binary_products

/--
For any `R`-algebra `A`, `A` is Morita equivalent to itself.
-/
def refl (A : Type u₁) [Ring A] [Algebra R A] : MoritaEquivalence R A A where
eqv := CategoryTheory.Equivalence.refl
linear := Functor.instLinearId

/--
For any `R`-algebras `A` and `B`, if `A` is Morita equivalent to `B`, then `B` is Morita equivalent
to `A`.
-/
def symm {A : Type u₁} [Ring A] [Algebra R A] {B : Type u₂} [Ring B] [Algebra R B]
(e : MoritaEquivalence R A B) : MoritaEquivalence R B A where
eqv := e.eqv.symm
linear := e.eqv.inverseLinear R

-- TODO: We have restricted all the rings to the same universe here because of the complication
-- `max u₁ u₂`, `max u₂ u₃` vs `max u₁ u₃`. But once we proved the definition of Morita
-- equivalence is equivalent to the existence of a full idempotent element, we can remove this
-- restriction in the universe.
-- Or alternatively, @alreadydone has sketched an argument on how the universe restriction can be
-- removed via a categorical argument,
-- see [here](https://github.com/leanprover-community/mathlib4/pull/20640#discussion_r1912189931)
/--
For any `R`-algebras `A`, `B`, and `C`, if `A` is Morita equivalent to `B` and `B` is Morita
equivalent to `C`, then `A` is Morita equivalent to `C`.
-/
def trans {A B C : Type u₁}
[Ring A] [Algebra R A] [Ring B] [Algebra R B] [Ring C] [Algebra R C]
(e : MoritaEquivalence R A B) (e' : MoritaEquivalence R B C) :
MoritaEquivalence R A C where
eqv := e.eqv.trans e'.eqv
linear := e.eqv.functor.instLinearComp e'.eqv.functor

variable {R} in
/--
Isomorphic `R`-algebras are Morita equivalent.
-/
noncomputable def ofAlgEquiv {A : Type u₁} {B : Type u₂}
[Ring A] [Algebra R A] [Ring B] [Algebra R B] (f : A ≃ₐ[R] B) :
MoritaEquivalence R A B where
eqv := ModuleCat.restrictScalarsEquivalenceOfRingEquiv f.symm.toRingEquiv
linear := ModuleCat.Algebra.restrictScalarsEquivalenceOfRingEquiv_linear f.symm

end MoritaEquivalence

/--
Let `A` and `B` be `R`-algebras. We say that `A` and `B` are Morita equivalent if the categories of
`A`-modules and `B`-modules are equivalent as `R`-linear categories.
-/
structure IsMoritaEquivalent
(A : Type u₁) [Ring A] [Algebra R A]
(B : Type u₂) [Ring B] [Algebra R B] : Prop where
cond : Nonempty <| MoritaEquivalence R A B

namespace IsMoritaEquivalent

lemma refl (A : Type u₁) [Ring A] [Algebra R A] : IsMoritaEquivalent R A A where
cond := ⟨.refl R A⟩

lemma symm {A : Type u₁} [Ring A] [Algebra R A] {B : Type u₂} [Ring B] [Algebra R B]
(h : IsMoritaEquivalent R A B) : IsMoritaEquivalent R B A where
cond := h.cond.map <| .symm R

lemma trans {A B C : Type u₁} [Ring A] [Ring B] [Ring C] [Algebra R A] [Algebra R B] [Algebra R C]
(h : IsMoritaEquivalent R A B) (h' : IsMoritaEquivalent R B C) :
IsMoritaEquivalent R A C where
cond := Nonempty.map2 (.trans R) h.cond h'.cond

lemma of_algEquiv {A : Type u₁} [Ring A] [Algebra R A] {B : Type u₂} [Ring B] [Algebra R B]
(f : A ≃ₐ[R] B) : IsMoritaEquivalent R A B where
cond := ⟨.ofAlgEquiv f⟩

end IsMoritaEquivalent
Loading