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

Commit add100a

Browse files
committed
feat(ring_theory/perfection): perfection.map (#5503)
1 parent 768497c commit add100a

File tree

1 file changed

+80
-7
lines changed

1 file changed

+80
-7
lines changed

src/ring_theory/perfection.lean

Lines changed: 80 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -24,7 +24,7 @@ Define the valuation on the tilt, and define a characteristic predicate for the
2424
2525
-/
2626

27-
universes u₁ u₂ u₃
27+
universes u₁ u₂ u₃ u₄
2828

2929
open_locale nnreal
3030

@@ -86,6 +86,10 @@ lemma coeff_pow_p (f : ring.perfection R p) (n : ℕ) :
8686
coeff R p (n + 1) (f ^ p) = coeff R p n f :=
8787
by { rw ring_hom.map_pow, exact f.2 n }
8888

89+
lemma coeff_pow_p' (f : ring.perfection R p) (n : ℕ) :
90+
coeff R p (n + 1) f ^ p = coeff R p n f :=
91+
f.2 n
92+
8993
lemma coeff_frobenius (f : ring.perfection R p) (n : ℕ) :
9094
coeff R p (n + 1) (frobenius _ p f) = coeff R p n f :=
9195
by convert coeff_pow_p f n
@@ -151,6 +155,25 @@ any homomorphism `R →+* S` can be lifted to a homomorphism `R →+* perfection
151155
by rw [← coeff_iterate_frobenius _ 0 n, zero_add, ← ring_hom.map_iterate_frobenius,
152156
right_inverse_pth_root_frobenius.iterate] }
153157

158+
lemma hom_ext {R : Type u₁} [comm_semiring R] [char_p R p] [perfect_ring R p]
159+
{S : Type u₂} [comm_semiring S] [char_p S p] {f g : R →+* ring.perfection S p}
160+
(hfg : ∀ x, coeff S p 0 (f x) = coeff S p 0 (g x)) : f = g :=
161+
(lift p R S).symm.injective $ ring_hom.ext hfg
162+
163+
variables {R} {S : Type u₂} [comm_semiring S] [char_p S p]
164+
165+
/-- A ring homomorphism `R →+* S` induces `perfection R p →+* perfection S p` -/
166+
@[simps] def map (φ : R →+* S) : ring.perfection R p →+* ring.perfection S p :=
167+
{ to_fun := λ f, ⟨λ n, φ (coeff R p n f), λ n, by rw [← φ.map_pow, coeff_pow_p']⟩,
168+
map_one' := subtype.eq $ funext $ λ n, φ.map_one,
169+
map_mul' := λ f g, subtype.eq $ funext $ λ n, φ.map_mul _ _,
170+
map_zero' := subtype.eq $ funext $ λ n, φ.map_zero,
171+
map_add' := λ f g, subtype.eq $ funext $ λ n, φ.map_add _ _ }
172+
173+
lemma coeff_map (φ : R →+* S) (f : ring.perfection R p) (n : ℕ) :
174+
coeff S p n (map p φ f) = φ (coeff R p n f) :=
175+
rfl
176+
154177
end perfection
155178

156179
/-- A perfection map to a ring of characteristic `p` is a map that is isomorphic
@@ -166,7 +189,7 @@ namespace perfection_map
166189

167190
variables {p : ℕ} [fact p.prime]
168191
variables {R : Type u₁} [comm_semiring R] [char_p R p]
169-
variables {P : Type u} [comm_semiring P] [char_p P p] [perfect_ring P p]
192+
variables {P : Type u} [comm_semiring P] [char_p P p] [perfect_ring P p]
170193

171194
/-- Create a `perfection_map` from an isomorphism to the perfection. -/
172195
@[simps] lemma mk' {f : P →+* R} (g : P ≃+* ring.perfection R p)
@@ -199,6 +222,26 @@ ring_equiv.of_bijective (perfection.lift p P R π)
199222
⟨λ x y hxy, m.injective $ λ n, (congr_arg (perfection.coeff R p n) hxy : _),
200223
λ f, let ⟨x, hx⟩ := m.surjective f.1 f.2 in ⟨x, perfection.ext $ hx⟩⟩
201224

225+
lemma equiv_apply {π : P →+* R} (m : perfection_map p π) (x : P) :
226+
m.equiv x = perfection.lift p P R π x :=
227+
rfl
228+
229+
lemma comp_equiv {π : P →+* R} (m : perfection_map p π) (x : P) :
230+
perfection.coeff R p 0 (m.equiv x) = π x :=
231+
rfl
232+
233+
lemma comp_equiv' {π : P →+* R} (m : perfection_map p π) :
234+
(perfection.coeff R p 0).comp ↑m.equiv = π :=
235+
ring_hom.ext $ λ x, rfl
236+
237+
lemma comp_symm_equiv {π : P →+* R} (m : perfection_map p π) (f : ring.perfection R p) :
238+
π (m.equiv.symm f) = perfection.coeff R p 0 f :=
239+
(m.comp_equiv _).symm.trans $ congr_arg _ $ m.equiv.apply_symm_apply f
240+
241+
lemma comp_symm_equiv' {π : P →+* R} (m : perfection_map p π) :
242+
π.comp ↑m.equiv.symm = perfection.coeff R p 0 :=
243+
ring_hom.ext m.comp_symm_equiv
244+
202245
variables (p R P)
203246
/-- Given rings `R` and `S` of characteristic `p`, with `R` being perfect,
204247
any homomorphism `R →+* S` can be lifted to a homomorphism `R →+* P`,
@@ -208,11 +251,41 @@ where `P` is any perfection of `S`. -/
208251
(π : P →+* S) (m : perfection_map p π) :
209252
(R →+* S) ≃ (R →+* P) :=
210253
{ to_fun := λ f, ring_hom.comp ↑m.equiv.symm $ perfection.lift p R S f,
211-
inv_fun := λ f, (perfection.lift p R S).symm (ring_hom.comp ↑m.equiv f),
212-
left_inv := λ f, (equiv.symm_apply_eq _).2 $ ring_hom.ext $ λ x,
213-
by simp_rw [ring_hom.comp_apply, ring_equiv.coe_ring_hom, ring_equiv.apply_symm_apply],
214-
right_inv := λ f, ring_hom.ext $ λ x, by simp_rw [equiv.apply_symm_apply,
215-
ring_hom.comp_apply, ring_equiv.coe_ring_hom, ring_equiv.symm_apply_apply] }
254+
inv_fun := λ f, π.comp f,
255+
left_inv := λ f, by { simp_rw [← ring_hom.comp_assoc, comp_symm_equiv'],
256+
exact (perfection.lift p R S).symm_apply_apply f },
257+
right_inv := λ f, ring_hom.ext $ λ x, m.equiv.injective $ (m.equiv.apply_symm_apply _).trans $
258+
show perfection.lift p R S (π.comp f) x = ring_hom.comp ↑m.equiv f x,
259+
from ring_hom.ext_iff.1 ((perfection.lift p R S).apply_eq_iff_eq_symm_apply.2 rfl) _ }
260+
261+
variables {R p}
262+
263+
lemma hom_ext [perfect_ring R p] {S : Type u₂} [comm_semiring S] [char_p S p]
264+
{P : Type u₃} [comm_semiring P] [char_p P p] [perfect_ring P p]
265+
(π : P →+* S) (m : perfection_map p π) {f g : R →+* P}
266+
(hfg : ∀ x, π (f x) = π (g x)) : f = g :=
267+
(lift p R S P π m).symm.injective $ ring_hom.ext hfg
268+
269+
variables {R P} (p) {S : Type u₂} [comm_semiring S] [char_p S p]
270+
variables {Q : Type u₄} [comm_semiring Q] [char_p Q p] [perfect_ring Q p]
271+
272+
/-- A ring homomorphism `R →+* S` induces `P →+* Q`, a map of the respective perfections -/
273+
@[nolint unused_arguments]
274+
noncomputable def map {π : P →+* R} (m : perfection_map p π) {σ : Q →+* S} (n : perfection_map p σ)
275+
(φ : R →+* S) : P →+* Q :=
276+
lift p P S Q σ n $ φ.comp π
277+
278+
lemma comp_map {π : P →+* R} (m : perfection_map p π) {σ : Q →+* S} (n : perfection_map p σ)
279+
(φ : R →+* S) : σ.comp (map p m n φ) = φ.comp π :=
280+
(lift p P S Q σ n).symm_apply_apply _
281+
282+
lemma map_map {π : P →+* R} (m : perfection_map p π) {σ : Q →+* S} (n : perfection_map p σ)
283+
(φ : R →+* S) (x : P) : σ (map p m n φ x) = φ (π x) :=
284+
ring_hom.ext_iff.1 (comp_map p m n φ) x
285+
286+
-- Why is this slow?
287+
lemma map_eq_map (φ : R →+* S) : map p (of p R) (of p S) φ = perfection.map p φ :=
288+
hom_ext _ (of p S) $ λ f, by rw [map_map, perfection.coeff_map]
216289

217290
end perfection_map
218291

0 commit comments

Comments
 (0)