Skip to content

Commit 1c080cc

Browse files
committed
refactor: make PadicInt.valuation -valued (#19858)
It is currently `ℤ`-valued, but always nonnegative. Also rename the misnamed `norm_eq_pow_val` to `norm_eq_zpow_neg_valuation` From FLT Closes ImperialCollegeLondon/FLT#277
1 parent 9ab8279 commit 1c080cc

File tree

4 files changed

+113
-131
lines changed

4 files changed

+113
-131
lines changed

Mathlib/NumberTheory/Padics/MahlerBasis.lean

Lines changed: 4 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -181,9 +181,10 @@ private lemma bojanic_mahler_step2 {f : C(ℤ_[p], E)} {s t : ℕ}
181181
refine mul_le_mul_of_nonneg_right ?_ (by simp only [zero_le])
182182
-- remains to show norm of binomial coeff is `≤ p⁻¹`
183183
have : 0 < (p ^ t).choose (i + 1) := Nat.choose_pos (by omega)
184-
rw [← zpow_neg_one, ← coe_le_coe, coe_nnnorm, Padic.norm_eq_pow_val (mod_cast this.ne'),
185-
coe_zpow, NNReal.coe_natCast, (zpow_right_strictMono₀ (mod_cast hp.out.one_lt)).le_iff_le,
186-
neg_le_neg_iff, Padic.valuation_natCast, Nat.one_le_cast]
184+
rw [← zpow_neg_one, ← coe_le_coe, coe_nnnorm, Padic.norm_eq_zpow_neg_valuation
185+
(mod_cast this.ne'), coe_zpow, NNReal.coe_natCast,
186+
zpow_le_zpow_iff_right₀ (mod_cast hp.out.one_lt), neg_le_neg_iff, Padic.valuation_natCast,
187+
Nat.one_le_cast]
187188
exact one_le_padicValNat_of_dvd this <| hp.out.dvd_choose_pow (by omega) (by omega)
188189
· -- Bounding the sum over `range (n + 1)`: every term is small by the choice of `t`
189190
refine norm_sum_le_of_forall_le_of_nonempty nonempty_range_succ (fun i _ ↦ ?_)

Mathlib/NumberTheory/Padics/PadicIntegers.lean

Lines changed: 54 additions & 96 deletions
Original file line numberDiff line numberDiff line change
@@ -50,20 +50,19 @@ open Padic Metric IsLocalRing
5050

5151
noncomputable section
5252

53+
variable (p : ℕ) [hp : Fact p.Prime]
54+
5355
/-- The `p`-adic integers `ℤ_[p]` are the `p`-adic numbers with norm `≤ 1`. -/
54-
def PadicInt (p : ℕ) [Fact p.Prime] :=
55-
{ x : ℚ_[p] // ‖x‖ ≤ 1 }
56+
def PadicInt : Type := {x : ℚ_[p] // ‖x‖ ≤ 1}
5657

5758
/-- The ring of `p`-adic integers. -/
5859
notation "ℤ_[" p "]" => PadicInt p
5960

6061
namespace PadicInt
62+
variable {p} {x y : ℤ_[p]}
6163

6264
/-! ### Ring structure and coercion to `ℚ_[p]` -/
6365

64-
65-
variable {p : ℕ} [Fact p.Prime]
66-
6766
instance : Coe ℤ_[p] ℚ_[p] :=
6867
⟨Subtype.val⟩
6968

@@ -127,9 +126,9 @@ theorem coe_one : ((1 : ℤ_[p]) : ℚ_[p]) = 1 := rfl
127126
@[simp, norm_cast]
128127
theorem coe_zero : ((0 : ℤ_[p]) : ℚ_[p]) = 0 := rfl
129128

130-
theorem coe_eq_zero (z : ℤ_[p]) : (z : ℚ_[p]) = 0z = 0 := by rw [← coe_zero, Subtype.coe_inj]
129+
@[simp] lemma coe_eq_zero : (x : ℚ_[p]) = 0x = 0 := by rw [← coe_zero, Subtype.coe_inj]
131130

132-
theorem coe_ne_zero (z : ℤ_[p]) : (z : ℚ_[p]) ≠ 0z0 := z.coe_eq_zero.not
131+
lemma coe_ne_zero : (x : ℚ_[p]) ≠ 0x0 := coe_eq_zero.not
133132

134133
instance : AddCommGroup ℤ_[p] := (by infer_instance : AddCommGroup (subring p))
135134

@@ -178,10 +177,6 @@ def ofIntSeq (seq : ℕ → ℤ) (h : IsCauSeq (padicNorm p) fun n => seq n) :
178177
split_ifs with hne <;> norm_cast
179178
apply padicNorm.of_int⟩
180179

181-
end PadicInt
182-
183-
namespace PadicInt
184-
185180
/-! ### Instances
186181
187182
We now show that `ℤ_[p]` is a
@@ -190,8 +185,7 @@ We now show that `ℤ_[p]` is a
190185
* integral domain
191186
-/
192187

193-
194-
variable (p : ℕ) [Fact p.Prime]
188+
variable (p)
195189

196190
instance : MetricSpace ℤ_[p] := Subtype.metricSpace
197191

@@ -228,15 +222,8 @@ variable {p}
228222

229223
instance : IsDomain ℤ_[p] := Function.Injective.isDomain (subring p).subtype Subtype.coe_injective
230224

231-
end PadicInt
232-
233-
namespace PadicInt
234-
235225
/-! ### Norm -/
236226

237-
238-
variable {p : ℕ} [Fact p.Prime]
239-
240227
theorem norm_le_one (z : ℤ_[p]) : ‖z‖ ≤ 1 := z.2
241228

242229
@[simp]
@@ -291,12 +278,6 @@ instance complete : CauSeq.IsComplete ℤ_[p] norm :=
291278
⟨⟨_, hqn⟩, fun ε => by
292279
simpa [norm, norm_def] using CauSeq.equiv_lim (cauSeq_to_rat_cauSeq f) ε⟩⟩
293280

294-
end PadicInt
295-
296-
namespace PadicInt
297-
298-
variable (p : ℕ) [hp : Fact p.Prime]
299-
300281
theorem exists_pow_neg_lt {ε : ℝ} (hε : 0 < ε) : ∃ k : ℕ, (p : ℝ) ^ (-(k : ℤ)) < ε := by
301282
obtain ⟨k, hk⟩ := exists_nat_gt ε⁻¹
302283
use k
@@ -329,50 +310,45 @@ theorem norm_int_le_pow_iff_dvd {k : ℤ} {n : ℕ} :
329310

330311
/-! ### Valuation on `ℤ_[p]` -/
331312

313+
lemma valuation_coe_nonneg : 0 ≤ (x : ℚ_[p]).valuation := by
314+
obtain rfl | hx := eq_or_ne x 0
315+
· simp
316+
have := x.2
317+
rwa [Padic.norm_eq_zpow_neg_valuation <| coe_ne_zero.2 hx, zpow_le_one_iff_right₀, neg_nonpos]
318+
at this
319+
exact mod_cast hp.out.one_lt
332320

333321
/-- `PadicInt.valuation` lifts the `p`-adic valuation on `ℚ` to `ℤ_[p]`. -/
334-
def valuation (x : ℤ_[p]) :=
335-
Padic.valuation (x : ℚ_[p])
322+
def valuation (x : ℤ_[p]) : ℕ := (x : ℚ_[p]).valuation.toNat
336323

337-
theorem norm_eq_pow_val {x : ℤ_[p]} (hx : x ≠ 0) : ‖x‖ = (p : ℝ) ^ (-x.valuation) := by
338-
refine @Padic.norm_eq_pow_val p hp x ?_
339-
contrapose! hx
340-
exact Subtype.val_injective hx
324+
@[simp, norm_cast] lemma valuation_coe (x : ℤ_[p]) : (x : ℚ_[p]).valuation = x.valuation := by
325+
simp [valuation, valuation_coe_nonneg]
341326

342-
@[simp]
343-
theorem valuation_zero : valuation (0 : ℤ_[p]) = 0 := Padic.valuation_zero
327+
@[simp] lemma valuation_zero : valuation (0 : ℤ_[p]) = 0 := by simp [valuation]
328+
@[simp] lemma valuation_one : valuation (1 : ℤ_[p]) = 0 := by simp [valuation]
329+
@[simp] lemma valuation_p : valuation (p : ℤ_[p]) = 1 := by simp [valuation]
344330

345-
@[simp]
346-
theorem valuation_one : valuation (1 : ℤ_[p]) = 0 := Padic.valuation_one
331+
lemma le_valuation_add (hxy : x + y ≠ 0) : min x.valuation y.valuation ≤ (x + y).valuation := by
332+
zify; simpa [← valuation_coe] using Padic.le_valuation_add <| coe_ne_zero.2 hxy
333+
334+
@[simp] lemma valuation_mul (hx : x ≠ 0) (hy : y ≠ 0) :
335+
(x * y).valuation = x.valuation + y.valuation := by
336+
zify; simp [← valuation_coe, Padic.valuation_mul (coe_ne_zero.2 hx) (coe_ne_zero.2 hy)]
347337

348338
@[simp]
349-
theorem valuation_p : valuation (p : ℤ_[p]) = 1 := by simp [valuation]
339+
lemma valuation_pow (x : ℤ_[p]) (n : ℕ) : (x ^ n).valuation = n * x.valuation := by
340+
zify; simp [← valuation_coe]
350341

351-
theorem valuation_nonneg (x : ℤ_[p]) : 0 ≤ x.valuation := by
352-
by_cases hx : x = 0
353-
· simp [hx]
354-
have h : (1 : ℝ) < p := mod_cast hp.1.one_lt
355-
rw [← neg_nonpos, ← (zpow_right_strictMono₀ h).le_iff_le]
356-
show (p : ℝ) ^ (-valuation x) ≤ (p : ℝ) ^ (0 : ℤ)
357-
rw [← norm_eq_pow_val hx]
358-
simpa using x.property
342+
lemma norm_eq_zpow_neg_valuation {x : ℤ_[p]} (hx : x ≠ 0) : ‖x‖ = p ^ (-x.valuation : ℤ) := by
343+
simp [norm_def, Padic.norm_eq_zpow_neg_valuation <| coe_ne_zero.2 hx]
344+
345+
@[deprecated (since := "2024-12-10")] alias norm_eq_pow_val := norm_eq_zpow_neg_valuation
359346

347+
-- TODO: Do we really need this lemma?
360348
@[simp]
361349
theorem valuation_p_pow_mul (n : ℕ) (c : ℤ_[p]) (hc : c ≠ 0) :
362350
((p : ℤ_[p]) ^ n * c).valuation = n + c.valuation := by
363-
have : ‖(p : ℤ_[p]) ^ n * c‖ = ‖(p : ℤ_[p]) ^ n‖ * ‖c‖ := norm_mul _ _
364-
have aux : (p : ℤ_[p]) ^ n * c ≠ 0 := by
365-
contrapose! hc
366-
rw [mul_eq_zero] at hc
367-
cases' hc with hc hc
368-
· refine (hp.1.ne_zero ?_).elim
369-
exact_mod_cast pow_eq_zero hc
370-
· exact hc
371-
rwa [norm_eq_pow_val aux, norm_p_pow, norm_eq_pow_val hc, ← zpow_add₀, ← neg_add,
372-
zpow_right_inj₀, neg_inj] at this
373-
· exact mod_cast hp.1.pos
374-
· exact mod_cast hp.1.ne_one
375-
· exact mod_cast hp.1.ne_zero
351+
rw [valuation_mul (NeZero.ne _) hc, valuation_pow, valuation_p, mul_one]
376352

377353
section Units
378354

@@ -429,25 +405,22 @@ theorem norm_units (u : ℤ_[p]ˣ) : ‖(u : ℤ_[p])‖ = 1 := isUnit_iff.mp <|
429405
/-- `unitCoeff hx` is the unit `u` in the unique representation `x = u * p ^ n`.
430406
See `unitCoeff_spec`. -/
431407
def unitCoeff {x : ℤ_[p]} (hx : x ≠ 0) : ℤ_[p]ˣ :=
432-
let u : ℚ_[p] := x * (p : ℚ_[p]) ^ (-x.valuation)
408+
let u : ℚ_[p] := x * (p : ℚ_[p]) ^ (-x.valuation : ℤ)
433409
have hu : ‖u‖ = 1 := by
434-
simp [u, hx, zpow_ne_zero (G₀ := ℝ) _ (Nat.cast_ne_zero.2 hp.1.pos.ne'), norm_eq_pow_val]
410+
simp [u, hx, pow_ne_zero _ (NeZero.ne _), norm_eq_zpow_neg_valuation]
435411
mkUnits hu
436412

437413
@[simp]
438414
theorem unitCoeff_coe {x : ℤ_[p]} (hx : x ≠ 0) :
439-
(unitCoeff hx : ℚ_[p]) = x * (p : ℚ_[p]) ^ (-x.valuation) := rfl
415+
(unitCoeff hx : ℚ_[p]) = x * (p : ℚ_[p]) ^ (-x.valuation : ℤ) := rfl
440416

441417
theorem unitCoeff_spec {x : ℤ_[p]} (hx : x ≠ 0) :
442-
x = (unitCoeff hx : ℤ_[p]) * (p : ℤ_[p]) ^ Int.natAbs (valuation x) := by
418+
x = (unitCoeff hx : ℤ_[p]) * (p : ℤ_[p]) ^ x.valuation := by
443419
apply Subtype.coe_injective
444420
push_cast
445-
have repr : (x : ℚ_[p]) = unitCoeff hx * (p : ℚ_[p]) ^ x.valuation := by
446-
rw [unitCoeff_coe, mul_assoc, ← zpow_add₀]
447-
· simp
448-
· exact mod_cast hp.1.ne_zero
449-
convert repr using 2
450-
rw [← zpow_natCast, Int.natAbs_of_nonneg (valuation_nonneg x)]
421+
rw [unitCoeff_coe, mul_assoc, ← zpow_natCast, ← zpow_add₀]
422+
· simp
423+
· exact NeZero.ne _
451424

452425
end Units
453426

@@ -457,35 +430,22 @@ section NormLeIff
457430

458431

459432
theorem norm_le_pow_iff_le_valuation (x : ℤ_[p]) (hx : x ≠ 0) (n : ℕ) :
460-
‖x‖ ≤ (p : ℝ) ^ (-n : ℤ) ↔ ↑n ≤ x.valuation := by
461-
rw [norm_eq_pow_val hx]
462-
lift x.valuation to ℕ using x.valuation_nonneg with k
463-
simp only [Int.ofNat_le, zpow_neg, zpow_natCast]
464-
have aux : ∀ m : ℕ, 0 < (p : ℝ) ^ m := by
465-
intro m
466-
refine pow_pos ?_ m
467-
exact mod_cast hp.1.pos
468-
rw [inv_le_inv₀ (aux _) (aux _)]
469-
have : p ^ n ≤ p ^ k ↔ n ≤ k := (pow_right_strictMono₀ hp.1.one_lt).le_iff_le
470-
rw [← this]
471-
norm_cast
433+
‖x‖ ≤ (p : ℝ) ^ (-n : ℤ) ↔ n ≤ x.valuation := by
434+
rw [norm_eq_zpow_neg_valuation hx, zpow_le_zpow_iff_right₀, neg_le_neg_iff, Nat.cast_le]
435+
exact mod_cast hp.out.one_lt
472436

473437
theorem mem_span_pow_iff_le_valuation (x : ℤ_[p]) (hx : x ≠ 0) (n : ℕ) :
474-
x ∈ (Ideal.span {(p : ℤ_[p]) ^ n} : Ideal ℤ_[p]) ↔ n ≤ x.valuation := by
438+
x ∈ (Ideal.span {(p : ℤ_[p]) ^ n} : Ideal ℤ_[p]) ↔ n ≤ x.valuation := by
475439
rw [Ideal.mem_span_singleton]
476440
constructor
477441
· rintro ⟨c, rfl⟩
478442
suffices c ≠ 0 by
479-
rw [valuation_p_pow_mul _ _ this, le_add_iff_nonneg_right]
480-
apply valuation_nonneg
443+
rw [valuation_p_pow_mul _ _ this]
444+
exact le_self_add
481445
contrapose! hx
482446
rw [hx, mul_zero]
483447
· nth_rewrite 2 [unitCoeff_spec hx]
484-
lift x.valuation to ℕ using x.valuation_nonneg with k
485-
simp only [Int.natAbs_ofNat, Units.isUnit, IsUnit.dvd_mul_left, Int.ofNat_le]
486-
intro H
487-
obtain ⟨k, rfl⟩ := Nat.exists_eq_add_of_le H
488-
simp only [pow_add, dvd_mul_right]
448+
simpa [Units.isUnit, IsUnit.dvd_mul_left] using pow_dvd_pow _
489449

490450
theorem norm_le_pow_iff_mem_span_pow (x : ℤ_[p]) (n : ℕ) :
491451
‖x‖ ≤ (p : ℝ) ^ (-n : ℤ) ↔ x ∈ (Ideal.span {(p : ℤ_[p]) ^ n} : Ideal ℤ_[p]) := by
@@ -525,7 +485,7 @@ instance : IsLocalRing ℤ_[p] :=
525485
IsLocalRing.of_nonunits_add <| by simp only [mem_nonunits]; exact fun x y => norm_lt_one_add
526486

527487
theorem p_nonnunit : (p : ℤ_[p]) ∈ nonunits ℤ_[p] := by
528-
have : (p : ℝ)⁻¹ < 1 := inv_lt_one_of_one_lt₀ <| mod_cast hp.1.one_lt
488+
have : (p : ℝ)⁻¹ < 1 := inv_lt_one_of_one_lt₀ <| mod_cast hp.out.one_lt
529489
rwa [← norm_p, ← mem_nonunits] at this
530490

531491
theorem maximalIdeal_eq_span_p : maximalIdeal ℤ_[p] = Ideal.span {(p : ℤ_[p])} := by
@@ -539,14 +499,14 @@ theorem maximalIdeal_eq_span_p : maximalIdeal ℤ_[p] = Ideal.span {(p : ℤ_[p]
539499
theorem prime_p : Prime (p : ℤ_[p]) := by
540500
rw [← Ideal.span_singleton_prime, ← maximalIdeal_eq_span_p]
541501
· infer_instance
542-
· exact mod_cast hp.1.ne_zero
502+
· exact NeZero.ne _
543503

544504
theorem irreducible_p : Irreducible (p : ℤ_[p]) := Prime.irreducible prime_p
545505

546506
instance : IsDiscreteValuationRing ℤ_[p] :=
547507
IsDiscreteValuationRing.ofHasUnitMulPowIrreducibleFactorization
548508
⟨p, irreducible_p, fun {x hx} =>
549-
⟨x.valuation.natAbs, unitCoeff hx, by rw [mul_comm, ← unitCoeff_spec hx]⟩⟩
509+
⟨x.valuation, unitCoeff hx, by rw [mul_comm, ← unitCoeff_spec hx]⟩⟩
550510

551511
theorem ideal_eq_span_pow_p {s : Ideal ℤ_[p]} (hs : s ≠ ⊥) :
552512
∃ n : ℕ, s = Ideal.span {(p : ℤ_[p]) ^ n} :=
@@ -565,9 +525,7 @@ instance : IsAdicComplete (maximalIdeal ℤ_[p]) ℤ_[p] where
565525
rw [← neg_sub, norm_neg]
566526
exact hx hn
567527
· refine ⟨x'.lim, fun n => ?_⟩
568-
have : (0 : ℝ) < (p : ℝ) ^ (-n : ℤ) := by
569-
apply zpow_pos
570-
exact mod_cast hp.1.pos
528+
have : (0 : ℝ) < (p : ℝ) ^ (-n : ℤ) := zpow_pos (mod_cast hp.out.pos) _
571529
obtain ⟨i, hi⟩ := equiv_def₃ (equiv_lim x') this
572530
by_cases hin : i ≤ n
573531
· exact (hi i le_rfl n hin).le
@@ -609,8 +567,8 @@ instance isFractionRing : IsFractionRing ℤ_[p] ℚ_[p] where
609567
intro h0
610568
rw [h0, norm_zero] at hx
611569
exact hx zero_le_one
612-
rw [ha, padicNormE.mul, padicNormE.norm_p_pow, Padic.norm_eq_pow_val hx, ← zpow_add',
613-
hn_coe, neg_neg, neg_add_cancel, zpow_zero]
570+
rw [ha, padicNormE.mul, padicNormE.norm_p_pow, Padic.norm_eq_zpow_neg_valuation hx,
571+
← zpow_add', hn_coe, neg_neg, neg_add_cancel, zpow_zero]
614572
exact Or.inl (Nat.cast_ne_zero.mpr (NeZero.ne p))
615573
use
616574
(⟨a, le_of_eq ha_norm⟩,

0 commit comments

Comments
 (0)