Skip to content

Commit bf26137

Browse files
committed
Feat (NumberTheory/Ostrowski): Proof of the non-archimedean case of Ostrowski's Theorem (#14026)
In this PR we prove that a non-archimedean non-trivial absolute value on the rational is equivalent to a padic one. Main result ```lean theorem mulRingNorm_equiv_padic_of_bounded : ∃! p, ∃ (hp : Fact (Nat.Prime p)), MulRingNorm.equiv f (mulRingNorm_padic p) ``` Co-authored-by: David Kurniadi Angdinata [[email protected]](mailto:[email protected]) Laura Capuano [[email protected]](mailto:[email protected]) Nirvana Coppola [[email protected]](mailto:[email protected]) María Inés de Frutos Fernández [[email protected]](mailto:[email protected]) Sam van Gool [[email protected]](mailto:[email protected]) Silvain Rideau-Kikuchi [[email protected]](mailto:[email protected]) Amos Turchet [[email protected]](mailto:[email protected]) Francesco Veneziano [[email protected]](mailto:[email protected])
1 parent 5bdfd1a commit bf26137

File tree

1 file changed

+173
-5
lines changed

1 file changed

+173
-5
lines changed

Mathlib/NumberTheory/Ostrowski.lean

+173-5
Original file line numberDiff line numberDiff line change
@@ -7,8 +7,10 @@ Francesco Veneziano
77
-/
88

99
import Mathlib.Analysis.Normed.Field.Basic
10+
import Mathlib.Analysis.SpecialFunctions.Log.Base
1011
import Mathlib.Analysis.SpecialFunctions.Pow.Real
1112
import Mathlib.Analysis.Normed.Ring.Seminorm
13+
import Mathlib.NumberTheory.Padics.PadicNorm
1214

1315
/-!
1416
# Ostrowski’s Theorem
@@ -32,27 +34,193 @@ open Int
3234
variable {f g : MulRingNorm ℚ}
3335

3436
/-- Values of a multiplicative norm of the rationals coincide on ℕ if and only if they coincide
35-
on . -/
36-
lemma eq_on_Nat_iff_eq_on_Int : (∀ n : ℕ , f n = g n) ↔ (∀ n : ℤ , f n = g n) := by
37+
on `ℤ`. -/
38+
lemma eq_on_nat_iff_eq_on_Int : (∀ n : ℕ , f n = g n) ↔ (∀ n : ℤ , f n = g n) := by
3739
refine ⟨fun h z ↦ ?_, fun a n ↦ a n⟩
3840
obtain ⟨n , rfl | rfl⟩ := eq_nat_or_neg z <;>
3941
simp only [Int.cast_neg, Int.cast_natCast, map_neg_eq_map, h n]
4042

4143
/-- Values of a multiplicative norm of the rationals are determined by the values on the natural
4244
numbers. -/
43-
lemma eq_on_Nat_iff_eq : (∀ n : ℕ , f n = g n) ↔ f = g := by
45+
lemma eq_on_nat_iff_eq : (∀ n : ℕ , f n = g n) ↔ f = g := by
4446
refine ⟨fun h ↦ ?_, fun h n ↦ congrFun (congrArg DFunLike.coe h) ↑n⟩
4547
ext z
46-
rw [← Rat.num_div_den z, map_div₀, map_div₀, h, eq_on_Nat_iff_eq_on_Int.mp h]
48+
rw [← Rat.num_div_den z, map_div₀, map_div₀, h, eq_on_nat_iff_eq_on_Int.mp h]
4749

4850
/-- The equivalence class of a multiplicative norm on the rationals is determined by its values on
4951
the natural numbers. -/
50-
lemma equiv_on_Nat_iff_equiv : (∃ c : ℝ, 0 < c ∧ (∀ n : ℕ , (f n)^c = g n)) ↔
52+
lemma equiv_on_nat_iff_equiv : (∃ c : ℝ, 0 < c ∧ (∀ n : ℕ , (f n) ^ c = g n)) ↔
5153
f.equiv g := by
5254
refine ⟨fun ⟨c, hc, h⟩ ↦ ⟨c, ⟨hc, ?_⟩⟩, fun ⟨c, hc, h⟩ ↦ ⟨c, ⟨hc, fun n ↦ by rw [← h]⟩⟩⟩
5355
ext x
5456
rw [← Rat.num_div_den x, map_div₀, map_div₀, Real.div_rpow (apply_nonneg f _)
5557
(apply_nonneg f _), h x.den, ← MulRingNorm.apply_natAbs_eq,← MulRingNorm.apply_natAbs_eq,
5658
h (natAbs x.num)]
5759

60+
open Rat.MulRingNorm
61+
62+
section Non_archimedean
63+
64+
-- ## Non-archimedean case
65+
66+
/-- The mulRingNorm corresponding to the p-adic norm on `ℚ`. -/
67+
def mulRingNorm_padic (p : ℕ) [Fact p.Prime] : MulRingNorm ℚ :=
68+
{ toFun := fun x : ℚ ↦ (padicNorm p x : ℝ),
69+
map_zero' := by simp only [padicNorm.zero, Rat.cast_zero]
70+
add_le' := by simp only; norm_cast; exact fun r s ↦ padicNorm.triangle_ineq r s
71+
neg' := by simp only [forall_const, padicNorm.neg];
72+
eq_zero_of_map_eq_zero' := by
73+
simp only [Rat.cast_eq_zero]
74+
apply padicNorm.zero_of_padicNorm_eq_zero
75+
map_one' := by simp only [ne_eq, one_ne_zero, not_false_eq_true, padicNorm.eq_zpow_of_nonzero,
76+
padicValRat.one, neg_zero, zpow_zero, Rat.cast_one]
77+
map_mul' := by simp only [padicNorm.mul, Rat.cast_mul, forall_const]
78+
}
79+
80+
@[simp] lemma mulRingNorm_eq_padic_norm (p : ℕ) [Fact p.Prime] (r : ℚ) :
81+
mulRingNorm_padic p r = padicNorm p r := rfl
82+
83+
-- ## Step 1: define `p = minimal n s. t. 0 < f n < 1`
84+
85+
variable (hf_nontriv : f ≠ 1) (bdd : ∀ n : ℕ, f n ≤ 1)
86+
87+
/-- There exists a minimal positive integer with absolute value smaller than 1. -/
88+
lemma exists_minimal_nat_zero_lt_mulRingNorm_lt_one : ∃ p : ℕ, (0 < f p ∧ f p < 1) ∧
89+
∀ m : ℕ, 0 < f m ∧ f m < 1 → p ≤ m := by
90+
-- There is a positive integer with absolute value different from one.
91+
obtain ⟨n, hn1, hn2⟩ : ∃ n : ℕ, n ≠ 0 ∧ f n ≠ 1 := by
92+
contrapose! hf_nontriv
93+
rw [← eq_on_nat_iff_eq]
94+
intro n
95+
rcases eq_or_ne n 0 with rfl | hn0
96+
· simp only [Nat.cast_zero, map_zero]
97+
· simp only [MulRingNorm.apply_one, Nat.cast_eq_zero, hn0, ↓reduceIte, hf_nontriv n hn0]
98+
set P := {m : ℕ | 0 < f ↑m ∧ f ↑m < 1} -- p is going to be the minimum of this set.
99+
have hP : P.Nonempty :=
100+
⟨n, map_pos_of_ne_zero f (Nat.cast_ne_zero.mpr hn1), lt_of_le_of_ne (bdd n) hn2⟩
101+
exact ⟨sInf P, Nat.sInf_mem hP, fun m hm ↦ Nat.sInf_le hm⟩
102+
103+
-- ## Step 2: p is prime
104+
105+
variable {p : ℕ} (hp0 : 0 < f p) (hp1 : f p < 1) (hmin : ∀ m : ℕ, 0 < f m ∧ f m < 1 → p ≤ m)
106+
107+
/-- The minimal positive integer with absolute value smaller than 1 is a prime number.-/
108+
lemma is_prime_of_minimal_nat_zero_lt_mulRingNorm_lt_one : p.Prime := by
109+
rw [← Nat.irreducible_iff_nat_prime]
110+
constructor -- Two goals: p is not a unit and any product giving p must contain a unit.
111+
· rw [Nat.isUnit_iff]
112+
rintro rfl
113+
simp only [Nat.cast_one, map_one, lt_self_iff_false] at hp1
114+
· rintro a b rfl
115+
rw [Nat.isUnit_iff, Nat.isUnit_iff]
116+
by_contra! con
117+
obtain ⟨ha₁, hb₁⟩ := con
118+
obtain ⟨ha₀, hb₀⟩ : a ≠ 0 ∧ b ≠ 0 := by
119+
refine mul_ne_zero_iff.1 fun h ↦ ?_
120+
rwa [h, Nat.cast_zero, map_zero, lt_self_iff_false] at hp0
121+
have hap : a < a * b := lt_mul_of_one_lt_right (by omega) (by omega)
122+
have hbp : b < a * b := lt_mul_of_one_lt_left (by omega) (by omega)
123+
have ha :=
124+
le_of_not_lt <| not_and.1 ((hmin a).mt hap.not_le) (map_pos_of_ne_zero f (mod_cast ha₀))
125+
have hb :=
126+
le_of_not_lt <| not_and.1 ((hmin b).mt hbp.not_le) (map_pos_of_ne_zero f (mod_cast hb₀))
127+
rw [Nat.cast_mul, map_mul] at hp1
128+
exact ((one_le_mul_of_one_le_of_one_le ha hb).trans_lt hp1).false
129+
130+
-- ## Step 3: if p does not divide m, then f m = 1
131+
132+
open Real
133+
134+
/-- A natural number not divible by `p` has absolute value 1. -/
135+
lemma mulRingNorm_eq_one_of_not_dvd {m : ℕ} (hpm : ¬ p ∣ m) : f m = 1 := by
136+
apply le_antisymm (bdd m)
137+
by_contra! hm
138+
set M := f p ⊔ f m with hM
139+
set k := Nat.ceil (M.logb (1 / 2)) + 1 with hk
140+
obtain ⟨a, b, bezout⟩ : IsCoprime (p ^ k : ℤ) (m ^ k) := by
141+
apply IsCoprime.pow (Nat.Coprime.isCoprime _)
142+
exact (Nat.Prime.coprime_iff_not_dvd
143+
(is_prime_of_minimal_nat_zero_lt_mulRingNorm_lt_one hp0 hp1 hmin)).2 hpm
144+
have le_half {x} (hx0 : 0 < x) (hx1 : x < 1) (hxM : x ≤ M) : x ^ k < 1 / 2 := by
145+
calc
146+
x ^ k = x ^ (k : ℝ) := (rpow_natCast x k).symm
147+
_ < x ^ M.logb (1 / 2) := by
148+
apply rpow_lt_rpow_of_exponent_gt hx0 hx1
149+
rw [hk]
150+
push_cast
151+
exact lt_add_of_le_of_pos (Nat.le_ceil (M.logb (1 / 2))) zero_lt_one
152+
_ ≤ x ^ x.logb (1 / 2) := by
153+
apply rpow_le_rpow_of_exponent_ge hx0 (le_of_lt hx1)
154+
simp only [one_div, ← log_div_log, Real.log_inv, neg_div, ← div_neg, hM]
155+
gcongr
156+
simp only [Left.neg_pos_iff]
157+
exact log_neg (lt_sup_iff.2 <| Or.inl hp0) (sup_lt_iff.2 ⟨hp1, hm⟩)
158+
_ = 1 / 2 := rpow_logb hx0 (ne_of_lt hx1) one_half_pos
159+
apply lt_irrefl (1 : ℝ)
160+
calc
161+
1 = f 1 := (map_one f).symm
162+
_ = f (a * p ^ k + b * m ^ k) := by rw_mod_cast [bezout]; norm_cast
163+
_ ≤ f (a * p ^ k) + f (b * m ^ k) := f.add_le' (a * p ^ k) (b * m ^ k)
164+
_ ≤ 1 * (f p) ^ k + 1 * (f m) ^ k := by
165+
simp only [map_mul, map_pow, le_refl]
166+
gcongr
167+
all_goals rw [← MulRingNorm.apply_natAbs_eq]; apply bdd
168+
_ = (f p) ^ k + (f m) ^ k := by simp only [one_mul]
169+
_ < 1 := by
170+
have hm₀ : 0 < f m :=
171+
map_pos_of_ne_zero _ <| Nat.cast_ne_zero.2 fun H ↦ hpm <| H ▸ dvd_zero p
172+
linarith only [le_half hp0 hp1 le_sup_left, le_half hm₀ hm le_sup_right]
173+
174+
-- ## Step 4: f p = p ^ (- t) for some positive real t
175+
176+
/-- The absolute value of `p` is `p ^ (-t)` for some positive real number `t`. -/
177+
lemma exists_pos_mulRingNorm_eq_pow_neg : ∃ t : ℝ, 0 < t ∧ f p = p ^ (-t) := by
178+
have pprime := is_prime_of_minimal_nat_zero_lt_mulRingNorm_lt_one hp0 hp1 hmin
179+
refine ⟨- logb p (f p), Left.neg_pos_iff.2 <| logb_neg (mod_cast pprime.one_lt) hp0 hp1, ?_⟩
180+
rw [neg_neg]
181+
refine (rpow_logb (mod_cast pprime.pos) ?_ hp0).symm
182+
simp only [ne_eq, Nat.cast_eq_one,Nat.Prime.ne_one pprime, not_false_eq_true]
183+
184+
-- ## Non-archimedean case: end goal
185+
186+
/-- If `f` is bounded and not trivial, then it is equivalent to a p-adic absolute value. -/
187+
theorem mulRingNorm_equiv_padic_of_bounded :
188+
∃! p, ∃ (hp : Fact (p.Prime)), MulRingNorm.equiv f (mulRingNorm_padic p) := by
189+
obtain ⟨p, hfp, hmin⟩ := exists_minimal_nat_zero_lt_mulRingNorm_lt_one hf_nontriv bdd
190+
have hprime := is_prime_of_minimal_nat_zero_lt_mulRingNorm_lt_one hfp.1 hfp.2 hmin
191+
have hprime_fact : Fact (p.Prime) := ⟨hprime⟩
192+
obtain ⟨t, h⟩ := exists_pos_mulRingNorm_eq_pow_neg hfp.1 hfp.2 hmin
193+
simp_rw [← equiv_on_nat_iff_equiv]
194+
use p
195+
constructor -- 2 goals: MulRingNorm.equiv f (mulRingNorm_padic p) and p is unique.
196+
· use hprime_fact
197+
refine ⟨t⁻¹, by simp only [inv_pos, h.1], fun n ↦ ?_⟩
198+
have ht : t⁻¹ ≠ 0 := inv_ne_zero h.1.ne'
199+
rcases eq_or_ne n 0 with rfl | hn -- Separate cases n=0 and n ≠ 0
200+
· simp only [Nat.cast_zero, map_zero, ne_eq, ht, not_false_eq_true, zero_rpow]
201+
· /- Any natural number can be written as a power of p times a natural number not divisible
202+
by p -/
203+
rcases Nat.exists_eq_pow_mul_and_not_dvd hn p hprime.ne_one with ⟨e, m, hpm, rfl⟩
204+
simp only [Nat.cast_mul, Nat.cast_pow, map_mul, map_pow, mulRingNorm_eq_padic_norm,
205+
padicNorm.padicNorm_p_of_prime, Rat.cast_inv, Rat.cast_natCast, inv_pow,
206+
mulRingNorm_eq_one_of_not_dvd bdd hfp.1 hfp.2 hmin hpm, h.2]
207+
rw [← padicNorm.nat_eq_one_iff] at hpm
208+
simp only [← rpow_natCast, p.cast_nonneg, ← rpow_mul, mul_one, ← rpow_neg, hpm, cast_one]
209+
congr
210+
field_simp [h.1.ne']
211+
ring
212+
· intro q ⟨hq_prime, h_equiv⟩
213+
by_contra! hne
214+
apply Prime.ne_one (Nat.Prime.prime (Fact.elim hq_prime))
215+
rw [ne_comm, ← Nat.coprime_primes hprime (Fact.elim hq_prime),
216+
Nat.Prime.coprime_iff_not_dvd hprime] at hne
217+
rcases h_equiv with ⟨c, _, h_eq⟩
218+
have h_eq' := h_eq q
219+
simp only [mulRingNorm_eq_one_of_not_dvd bdd hfp.1 hfp.2 hmin hne, one_rpow,
220+
mulRingNorm_eq_padic_norm, padicNorm.padicNorm_p_of_prime, cast_inv, cast_natCast, eq_comm,
221+
inv_eq_one] at h_eq'
222+
norm_cast at h_eq'
223+
224+
end Non_archimedean
225+
58226
end Rat.MulRingNorm

0 commit comments

Comments
 (0)