|
2 | 2 | Copyright (c) 2018 Robert Y. Lewis. All rights reserved.
|
3 | 3 | Released under Apache 2.0 license as described in the file LICENSE.
|
4 | 4 | Authors: Robert Y. Lewis
|
5 |
| -
|
6 |
| -A proof of Hensel's lemma on ℤ_p, roughly following Keith Conrad's writeup: |
7 |
| -http://www.math.uconn.edu/~kconrad/blurbs/gradnumthy/hensel.pdf |
8 | 5 | -/
|
9 | 6 |
|
10 | 7 | import data.padics.padic_integers data.polynomial topology.metric_space.cau_seq_filter
|
11 | 8 | import analysis.specific_limits topology.instances.polynomial
|
12 |
| -import tactic.basic |
| 9 | + |
| 10 | +/-! |
| 11 | +# Hensel's lemma on ℤ_p |
| 12 | +
|
| 13 | +This file proves Hensel's lemma on ℤ_p, roughly following Keith Conrad's writeup: |
| 14 | +http://www.math.uconn.edu/~kconrad/blurbs/gradnumthy/hensel.pdf |
| 15 | +
|
| 16 | +Hensel's lemma gives a simple condition for the existence of a root of a polynomial. |
| 17 | +
|
| 18 | +The proof and motivation are described in the paper |
| 19 | +[R. Y. Lewis, *A formal proof of Hensel's lemma over the p-adic integers*][lewis2019]. |
| 20 | +
|
| 21 | +## References |
| 22 | +
|
| 23 | +* http://www.math.uconn.edu/~kconrad/blurbs/gradnumthy/hensel.pdf |
| 24 | +* [R. Y. Lewis, *A formal proof of Hensel's lemma over the p-adic integers*][lewis2019] |
| 25 | +* https://en.wikipedia.org/wiki/Hensel%27s_lemma |
| 26 | +
|
| 27 | +## Tags |
| 28 | +
|
| 29 | +p-adic, p adic, padic, p-adic integer |
| 30 | +-/ |
13 | 31 |
|
14 | 32 | noncomputable theory
|
15 | 33 |
|
16 | 34 | local attribute [instance] classical.prop_decidable
|
17 | 35 |
|
| 36 | +-- We begin with some general lemmas that are used below in the computation. |
| 37 | + |
18 | 38 | lemma padic_polynomial_dist {p : ℕ} [p.prime] (F : polynomial ℤ_[p]) (x y : ℤ_[p]) :
|
19 | 39 | ∥F.eval x - F.eval y∥ ≤ ∥x - y∥ :=
|
20 | 40 | let ⟨z, hz⟩ := F.eval_sub_factor x y in calc
|
@@ -69,6 +89,7 @@ parameters {p : ℕ} [nat.prime p] {F : polynomial ℤ_[p]} {a : ℤ_[p]}
|
69 | 89 | (hnorm : ∥F.eval a∥ < ∥F.derivative.eval a∥^2) (hnsol : F.eval a ≠ 0)
|
70 | 90 | include hnorm
|
71 | 91 |
|
| 92 | +/-- `T` is an auxiliary value that is used to control the behavior of the polynomial `F`. -/ |
72 | 93 | private def T : ℝ := ∥(F.eval a).val / ((F.derivative.eval a).val)^2∥
|
73 | 94 |
|
74 | 95 | private lemma deriv_sq_norm_pos : 0 < ∥F.derivative.eval a∥ ^ 2 :=
|
@@ -101,6 +122,7 @@ private lemma T_pow' (n : ℕ) : T ^ (2 ^ n) < 1 := (T_pow (nat.pow_pos (by norm
|
101 | 122 |
|
102 | 123 | private lemma T_pow_nonneg (n : ℕ) : T ^ n ≥ 0 := pow_nonneg (norm_nonneg _) _
|
103 | 124 |
|
| 125 | +/-- We will construct a sequence of elements of ℤ_p satisfying successive values of `ih`. -/ |
104 | 126 | private def ih (n : ℕ) (z : ℤ_[p]) : Prop :=
|
105 | 127 | ∥F.derivative.eval z∥ = ∥F.derivative.eval a∥ ∧ ∥F.eval z∥ ≤ ∥F.derivative.eval a∥^2 * T ^ (2^n)
|
106 | 128 |
|
@@ -164,7 +186,8 @@ calc ∥F.eval z'∥
|
164 | 186 |
|
165 | 187 | set_option eqn_compiler.zeta true
|
166 | 188 |
|
167 |
| --- we need (ih k) in order to construct the value for k+1, otherwise it might not be an integer. |
| 189 | +/-- Given `z : ℤ_[p]` satisfying `ih n z`, construct `z' : ℤ_[p]` satisfying `ih (n+1) z'`. We need |
| 190 | +the hypothesis `ih n z`, since otherwise `z'` is not necessarily an integer. -/ |
168 | 191 | private def ih_n {n : ℕ} {z : ℤ_[p]} (hz : ih n z) : {z' : ℤ_[p] // ih (n+1) z'} :=
|
169 | 192 | have h1 : ∥(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)∥ ≤ 1, from calc_norm_le_one hz,
|
170 | 193 | let z1 : ℤ_[p] := ⟨_, h1⟩,
|
@@ -261,7 +284,6 @@ private lemma newton_seq_dist_aux (n : ℕ) :
|
261 | 284 | private lemma newton_seq_dist {n k : ℕ} (hnk : n ≤ k) :
|
262 | 285 | ∥newton_seq k - newton_seq n∥ ≤ ∥F.derivative.eval a∥ * T^(2^n) :=
|
263 | 286 | have hex : ∃ m, k = n + m, from exists_eq_add_of_le hnk,
|
264 |
| --- ⟨k - n, by rw [←nat.add_sub_assoc hnk, add_comm, nat.add_sub_assoc (le_refl n), nat.sub_self, nat.add_zero]⟩, |
265 | 287 | let ⟨_, hex'⟩ := hex in
|
266 | 288 | by rw hex'; apply newton_seq_dist_aux; assumption
|
267 | 289 |
|
|
0 commit comments