Skip to content

Commit 38c98ef

Browse files
authored
Merge pull request #127 from ImperialCollegeLondon/Ore-Localization-for-algebras
Add algebra instances for localizations
2 parents 683d007 + bc76c01 commit 38c98ef

File tree

1 file changed

+192
-0
lines changed

1 file changed

+192
-0
lines changed

FLT/ForMathlib/Algebra.lean

+192
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,192 @@
1+
import Mathlib.RingTheory.OreLocalization.Ring
2+
3+
-- extend localization theory to algebras
4+
5+
/-!
6+
7+
# Algebra instances of Ore Localizations
8+
9+
If `R` is a commutative ring with submonoid `S`, and if `A` is a commutative `R`-algebra,
10+
then my guess is that `A[S⁻¹]` is an `R[S⁻¹]`-algebra. Let's see if I'm right and if so,
11+
in what generality.
12+
13+
-/
14+
15+
namespace OreLocalization
16+
17+
section CommMagma
18+
19+
variable {R A : Type*} [CommMonoid R] [CommMagma A] [MulAction R A] [IsScalarTower R A A]
20+
{S : Submonoid R}
21+
22+
@[to_additive]
23+
private def mul' (a₁ : A) (s₁ : S) (a₂ : A) (s₂ : S) : A[S⁻¹] :=
24+
a₁ * a₂ /ₒ (s₂ * s₁)
25+
26+
@[to_additive]
27+
private def mul''
28+
(a : A) (s : S) : A[S⁻¹] → A[S⁻¹] :=
29+
liftExpand (mul' a s) fun a₁ r₁ ⟨s₁, hs₁⟩ hr₁s₁ => by
30+
unfold OreLocalization.mul'
31+
simp only at hr₁s₁ ⊢
32+
rw [oreDiv_eq_iff]
33+
use ⟨s₁, hs₁⟩, r₁ * s₁
34+
simp only [Submonoid.mk_smul, Submonoid.coe_mul]
35+
constructor
36+
· rw [← smul_mul_assoc]
37+
rw [← smul_mul_assoc]
38+
rw [mul_comm]
39+
rw [smul_mul_assoc, smul_mul_assoc]
40+
rw [mul_comm]
41+
-- it's like a bloody Rubik's cube
42+
rw [smul_mul_assoc]
43+
rw [← mul_smul]
44+
· obtain ⟨s₂, hs₂⟩ := s
45+
simpa only [Submonoid.mk_smul, Submonoid.coe_mul] using mul_left_comm s₁ (r₁ * s₁) s₂
46+
47+
@[to_additive]
48+
protected def mul : A[S⁻¹] → A[S⁻¹] → A[S⁻¹] :=
49+
liftExpand mul'' fun a₁ r₁ s hs => by
50+
obtain ⟨s₁, hs₁⟩ := s
51+
simp only at hs
52+
unfold OreLocalization.mul''
53+
simp
54+
unfold OreLocalization.mul'
55+
dsimp
56+
ext sa
57+
induction sa
58+
rename_i a s_temp
59+
obtain ⟨s, hs⟩ := s_temp
60+
rw [liftExpand_of]
61+
rw [liftExpand_of]
62+
rw [oreDiv_eq_iff]
63+
simp only [Submonoid.mk_smul, Submonoid.coe_mul]
64+
use ⟨s₁, hs₁⟩, r₁ * s₁
65+
simp only [Submonoid.mk_smul, Submonoid.coe_mul]
66+
constructor
67+
· rw [smul_mul_assoc]
68+
rw [← mul_smul]
69+
rw [mul_comm]
70+
· repeat rw [mul_assoc]
71+
repeat rw [mul_left_comm s₁]
72+
rw [mul_left_comm s]
73+
74+
instance : Mul (A[S⁻¹]) where
75+
mul := OreLocalization.mul
76+
77+
protected def mul_def (a : A) (s : { x // x ∈ S }) (b : A) (t : { x // x ∈ S }) :
78+
a /ₒ s * (b /ₒ t) = a * b /ₒ (t * s) := rfl
79+
80+
unseal OreLocalization.smul in
81+
example (as bt : R[S⁻¹]) : as * bt = as • bt := rfl
82+
83+
end CommMagma
84+
85+
section One
86+
87+
variable {R A : Type*} [CommMonoid R] [MulAction R A] [One A] --[MulAction R A] [IsScalarTower R A A]
88+
{S : Submonoid R}
89+
90+
instance : One (A[S⁻¹]) where
91+
one := 1 /ₒ 1
92+
93+
protected lemma one_def' : (1 : A[S⁻¹]) = 1 /ₒ 1 := rfl
94+
95+
end One
96+
97+
section CommMonoid
98+
99+
variable {R A : Type*} [CommMonoid R] [CommMonoid A] [MulAction R A] [IsScalarTower R A A]
100+
{S : Submonoid R}
101+
102+
instance commMonoid : CommMonoid (A[S⁻¹]) where
103+
mul_assoc a b c := by
104+
induction' a with a
105+
induction' b with b
106+
induction' c with c
107+
change (a * b * c) /ₒ _ = (a * (b * c)) /ₒ _
108+
simp [mul_assoc]
109+
one_mul a := by
110+
induction' a with a
111+
change (1 * a) /ₒ _ = a /ₒ _
112+
simp
113+
mul_one a := by
114+
induction' a with a s
115+
simp [OreLocalization.mul_def, OreLocalization.one_def']
116+
mul_comm a b := by
117+
induction' a with a
118+
induction' b with b
119+
simp only [OreLocalization.mul_def, mul_comm]
120+
121+
end CommMonoid
122+
123+
section CommSemiring
124+
125+
variable {R A : Type*} [CommMonoid R] [CommRing A] [DistribMulAction R A] [IsScalarTower R A A]
126+
{S : Submonoid R}
127+
128+
lemma left_distrib' (a b c : A[S⁻¹]) :
129+
a * (b + c) = a * b + a * c := by
130+
induction' a with a s
131+
induction' b with b t
132+
induction' c with c u
133+
rcases oreDivAddChar' b c t u with ⟨r₁, s₁, h₁, q⟩; rw [q]; clear q
134+
simp only [OreLocalization.mul_def]
135+
rcases oreDivAddChar' (a * b) (a * c) (t * s) (u * s) with ⟨r₂, s₂, h₂, q⟩; rw [q]; clear q
136+
rw [oreDiv_eq_iff]
137+
sorry
138+
139+
instance : CommSemiring A[S⁻¹] where
140+
__ := commMonoid
141+
left_distrib := left_distrib'
142+
right_distrib a b c := by
143+
rw [mul_comm, mul_comm a, mul_comm b, left_distrib']
144+
zero_mul a := by
145+
induction' a with a s
146+
rw [OreLocalization.zero_def, OreLocalization.mul_def]
147+
simp
148+
mul_zero a := by
149+
induction' a with a s
150+
rw [OreLocalization.zero_def, OreLocalization.mul_def]
151+
simp
152+
153+
end CommSemiring
154+
155+
156+
-- fails on mathlib master;
157+
-- works if irreducibiliy of OreLocalization.smul is removed
158+
159+
-- Next job: make API so I can prove `Algebra (R[S⁻¹]) A[S⁻¹]`
160+
-- Might also want `numeratorAlgHom : A →ₐ[R] A[S⁻¹]`
161+
-- Might also want some universal property a la
162+
/-
163+
variable {T : Type*} [Semiring T]
164+
variable (f : R →+* T) (fS : S →* Units T)
165+
variable (hf : ∀ s : S, f s = fS s)
166+
167+
/-- The universal lift from a ring homomorphism `f : R →+* T`, which maps elements in `S` to
168+
units of `T`, to a ring homomorphism `R[S⁻¹] →+* T`. This extends the construction on
169+
monoids. -/
170+
def universalHom : R[S⁻¹] →+* T :=
171+
-/
172+
173+
174+
section first_goal
175+
176+
variable {R A : Type*} [CommRing R] [CommRing A] [Algebra R A] {S : Submonoid R}
177+
178+
abbrev SR := R[S⁻¹]
179+
abbrev SA := A[S⁻¹]
180+
181+
182+
--unseal OreLocalization.smul in
183+
--instance : Algebra (R[S⁻¹]) (A[S⁻¹]) where
184+
/-
185+
error:
186+
failed to synthesize
187+
Semiring (OreLocalization S A)
188+
-/
189+
190+
end first_goal
191+
192+
end OreLocalization

0 commit comments

Comments
 (0)