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

Commit 49be50f

Browse files
robertylewismergify[bot]
authored andcommitted
doc(data/padics, data/real/cau_seq, algebra): add doc strings, remove unnecessary assumptions (#1283)
* doc(data/padics): add doc strings, remove unnecessary prime assumptions * fix(data/real/cau_seq): remove unnecessary hypotheses * fix(algebra/{field, ordered_field}): remove unused assumptions * doc(data/real/cau_seq): document Cauchy sequences * fix(algebra/field): remove obsolete lemma * fix build * fix build * more unnecessary arguments * Update src/data/padics/padic_numbers.lean * Update src/data/padics/padic_numbers.lean * remove another unnecessary argument (suggested by @sgouezel)
1 parent 88d60dc commit 49be50f

File tree

10 files changed

+242
-76
lines changed

10 files changed

+242
-76
lines changed

docs/references.bib

Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -48,6 +48,24 @@ @book {james1999
4848
URL = {https://doi.org/10.1007/978-1-4471-3994-2},
4949
}
5050

51+
@inproceedings{lewis2019,
52+
author = {Lewis, Robert Y.},
53+
title = {A Formal Proof of Hensel\&\#39;s Lemma over the P-adic Integers},
54+
booktitle = {Proceedings of the 8th ACM SIGPLAN International Conference on Certified Programs and Proofs},
55+
series = {CPP 2019},
56+
year = {2019},
57+
isbn = {978-1-4503-6222-1},
58+
location = {Cascais, Portugal},
59+
pages = {15--26},
60+
numpages = {12},
61+
url = {http://doi.acm.org/10.1145/3293880.3294089},
62+
doi = {10.1145/3293880.3294089},
63+
acmid = {3294089},
64+
publisher = {ACM},
65+
address = {New York, NY, USA},
66+
keywords = {Hensel\&\#39;s lemma, Lean, formal proof, p-adic},
67+
}
68+
5169
@book {riehl2017,
5270
AUTHOR = {Riehl, Emily},
5371
TITLE = {Category theory in context},

src/algebra/field.lean

Lines changed: 2 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -39,7 +39,7 @@ def mk0 (a : α) (ha : a ≠ 0) : units α :=
3939
@[simp] lemma mk0_coe (u : units α) (h : (u : α) ≠ 0) : mk0 (u : α) h = u :=
4040
units.ext rfl
4141

42-
@[simp] lemma units.mk0_inj [field α] {a b : α} (ha : a ≠ 0) (hb : b ≠ 0) :
42+
@[simp] lemma units.mk0_inj {a b : α} (ha : a ≠ 0) (hb : b ≠ 0) :
4343
units.mk0 a ha = units.mk0 b hb ↔ a = b :=
4444
⟨λ h, by injection h, λ h, units.ext h⟩
4545

@@ -144,7 +144,7 @@ by rw [field.div_div_eq_div_mul _ hb hc, field.div_div_eq_div_mul _ hc hb, mul_c
144144
lemma field.div_div_div_cancel_right (a : α) (hb : b ≠ 0) (hc : c ≠ 0) : (a / c) / (b / c) = a / b :=
145145
by rw [field.div_div_eq_mul_div _ hb hc, div_mul_cancel _ hc]
146146

147-
lemma field.div_mul_div_cancel (a : α) (hb : b ≠ 0) (hc : c ≠ 0) : (a / c) * (c / b) = a / b :=
147+
lemma div_mul_div_cancel (a : α) (hc : c ≠ 0) : (a / c) * (c / b) = a / b :=
148148
by rw [← mul_div_assoc, div_mul_cancel _ hc]
149149

150150
lemma div_eq_div_iff (hb : b ≠ 0) (hd : d ≠ 0) : a / b = c / d ↔ a * d = c * b :=
@@ -171,10 +171,6 @@ lemma div_div_div_cancel_right (a b : α) (hc : c ≠ 0) : (a / c) / (b / c) = a
171171
if b0 : b = 0 then by simp only [b0, div_zero, zero_div] else
172172
field.div_div_div_cancel_right _ b0 hc
173173

174-
lemma div_mul_div_cancel (a : α) (hb : b ≠ 0) (hc : c ≠ 0) : (a / c) * (c / b) = a / b :=
175-
if b0 : b = 0 then by simp only [b0, div_zero, mul_zero] else
176-
field.div_mul_div_cancel _ b0 hc
177-
178174
lemma div_div_cancel (ha : a ≠ 0) : a / (a / b) = b :=
179175
if b0 : b = 0 then by simp only [b0, div_zero] else
180176
field.div_div_cancel ha b0

src/algebra/ordered_field.lean

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -139,17 +139,17 @@ calc (λx, x * c) '' {r | a ≤ r ∧ r ≤ b } = (λx, x / c) ⁻¹' {r | a ≤
139139
... = {r | a * c ≤ r ∧ r ≤ b * c} :
140140
set.ext $ by simp [le_div_iff, div_le_iff, hc]
141141

142-
instance linear_ordered_field.to_densely_ordered [linear_ordered_field α] : densely_ordered α :=
142+
instance linear_ordered_field.to_densely_ordered : densely_ordered α :=
143143
{ dense := assume a₁ a₂ h, ⟨(a₁ + a₂) / 2,
144144
calc a₁ = (a₁ + a₁) / 2 : (add_self_div_two a₁).symm
145145
... < (a₁ + a₂) / 2 : div_lt_div_of_lt_of_pos (add_lt_add_left h _) two_pos,
146146
calc (a₁ + a₂) / 2 < (a₂ + a₂) / 2 : div_lt_div_of_lt_of_pos (add_lt_add_right h _) two_pos
147147
... = a₂ : add_self_div_two a₂⟩ }
148148

149-
instance linear_ordered_field.to_no_top_order [linear_ordered_field α] : no_top_order α :=
149+
instance linear_ordered_field.to_no_top_order : no_top_order α :=
150150
{ no_top := assume a, ⟨a + 1, lt_add_of_le_of_pos (le_refl a) zero_lt_one ⟩ }
151151

152-
instance linear_ordered_field.to_no_bot_order [linear_ordered_field α] : no_bot_order α :=
152+
instance linear_ordered_field.to_no_bot_order : no_bot_order α :=
153153
{ no_bot := assume a, ⟨a + -1,
154154
add_lt_of_le_of_neg (le_refl _) (neg_lt_of_neg_lt $ by simp [zero_lt_one]) ⟩ }
155155

@@ -181,10 +181,10 @@ namespace nat
181181

182182
variables {α : Type*} [linear_ordered_field α]
183183

184-
lemma inv_pos_of_nat [linear_ordered_field α] {n : ℕ} : 0 < ((n : α) + 1)⁻¹ :=
184+
lemma inv_pos_of_nat {n : ℕ} : 0 < ((n : α) + 1)⁻¹ :=
185185
inv_pos $ add_pos_of_nonneg_of_pos n.cast_nonneg zero_lt_one
186186

187-
lemma one_div_pos_of_nat [linear_ordered_field α] {n : ℕ} : 0 < 1 / ((n : α) + 1) :=
187+
lemma one_div_pos_of_nat {n : ℕ} : 0 < 1 / ((n : α) + 1) :=
188188
by { rw one_div_eq_inv, exact inv_pos_of_nat }
189189

190190
end nat

src/data/padics/hensel.lean

Lines changed: 28 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -2,19 +2,39 @@
22
Copyright (c) 2018 Robert Y. Lewis. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
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
85
-/
96

107
import data.padics.padic_integers data.polynomial topology.metric_space.cau_seq_filter
118
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+
-/
1331

1432
noncomputable theory
1533

1634
local attribute [instance] classical.prop_decidable
1735

36+
-- We begin with some general lemmas that are used below in the computation.
37+
1838
lemma padic_polynomial_dist {p : ℕ} [p.prime] (F : polynomial ℤ_[p]) (x y : ℤ_[p]) :
1939
∥F.eval x - F.eval y∥ ≤ ∥x - y∥ :=
2040
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]}
6989
(hnorm : ∥F.eval a∥ < ∥F.derivative.eval a∥^2) (hnsol : F.eval a ≠ 0)
7090
include hnorm
7191

92+
/-- `T` is an auxiliary value that is used to control the behavior of the polynomial `F`. -/
7293
private def T : ℝ := ∥(F.eval a).val / ((F.derivative.eval a).val)^2
7394

7495
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
101122

102123
private lemma T_pow_nonneg (n : ℕ) : T ^ n ≥ 0 := pow_nonneg (norm_nonneg _) _
103124

125+
/-- We will construct a sequence of elements of ℤ_p satisfying successive values of `ih`. -/
104126
private def ih (n : ℕ) (z : ℤ_[p]) : Prop :=
105127
∥F.derivative.eval z∥ = ∥F.derivative.eval a∥ ∧ ∥F.eval z∥ ≤ ∥F.derivative.eval a∥^2 * T ^ (2^n)
106128

@@ -164,7 +186,8 @@ calc ∥F.eval z'∥
164186

165187
set_option eqn_compiler.zeta true
166188

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. -/
168191
private def ih_n {n : ℕ} {z : ℤ_[p]} (hz : ih n z) : {z' : ℤ_[p] // ih (n+1) z'} :=
169192
have h1 : ∥(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)∥ ≤ 1, from calc_norm_le_one hz,
170193
let z1 : ℤ_[p] := ⟨_, h1⟩,
@@ -261,7 +284,6 @@ private lemma newton_seq_dist_aux (n : ℕ) :
261284
private lemma newton_seq_dist {n k : ℕ} (hnk : n ≤ k) :
262285
∥newton_seq k - newton_seq n∥ ≤ ∥F.derivative.eval a∥ * T^(2^n) :=
263286
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]⟩,
265287
let ⟨_, hex'⟩ := hex in
266288
by rw hex'; apply newton_seq_dist_aux; assumption
267289

src/data/padics/padic_integers.lean

Lines changed: 38 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -2,31 +2,64 @@
22
Copyright (c) 2018 Robert Y. Lewis. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Robert Y. Lewis, Mario Carneiro
5-
6-
Define the p-adic integers ℤ_p as a subtype of ℚ_p. Construct algebraic structures on ℤ_p.
75
-/
86

97
import data.padics.padic_numbers ring_theory.ideals data.int.modeq
10-
import tactic.linarith
8+
9+
/-!
10+
# p-adic integers
11+
12+
This file defines the p-adic integers ℤ_p as the subtype of ℚ_p with norm ≤ 1. We show that ℤ_p is a
13+
complete nonarchimedean normed local ring.
14+
15+
## Important definitions
16+
17+
* `padic_int` : the type of p-adic numbers
18+
19+
## Notation
20+
21+
We introduce the notation ℤ_[p] for the p-adic integers.
22+
23+
## Implementation notes
24+
25+
Much, but not all, of this file assumes that `p` is prime. This assumption is inferred automatically
26+
by taking (prime p) as a type class argument.
27+
28+
Coercions into ℤ_p are set up to work with the `norm_cast` tactic.
29+
30+
## References
31+
32+
* [F. Q. Gouêva, *p-adic numbers*][gouvea1997]
33+
* [R. Y. Lewis, *A formal proof of Hensel's lemma over the p-adic integers*][lewis2019]
34+
* https://en.wikipedia.org/wiki/P-adic_number
35+
36+
## Tags
37+
38+
p-adic, p adic, padic, p-adic integer
39+
-/
1140

1241
open nat padic metric
1342
noncomputable theory
1443
local attribute [instance] classical.prop_decidable
1544

45+
/-- The p-adic integers ℤ_p are the p-adic numbers with norm ≤ 1. -/
1646
def padic_int (p : ℕ) [p.prime] := {x : ℚ_[p] // ∥x∥ ≤ 1}
1747
notation `ℤ_[`p`]` := padic_int p
1848

1949
namespace padic_int
2050
variables {p : ℕ} [nat.prime p]
2151

52+
/-- Addition on ℤ_p is inherited from ℚ_p. -/
2253
def add : ℤ_[p] → ℤ_[p] → ℤ_[p]
2354
| ⟨x, hx⟩ ⟨y, hy⟩ := ⟨x+y,
2455
le_trans (padic_norm_e.nonarchimedean _ _) (max_le_iff.2 ⟨hx,hy⟩)⟩
2556

57+
/-- Multiplication on ℤ_p is inherited from ℚ_p. -/
2658
def mul : ℤ_[p] → ℤ_[p] → ℤ_[p]
2759
| ⟨x, hx⟩ ⟨y, hy⟩ := ⟨x*y,
2860
begin rw padic_norm_e.mul, apply mul_le_one; {assumption <|> apply norm_nonneg} end
2961

62+
/-- Negation on ℤ_p is inherited from ℚ_p. -/
3063
def neg : ℤ_[p] → ℤ_[p]
3164
| ⟨x, hx⟩ := ⟨-x, by simpa⟩
3265

@@ -83,6 +116,8 @@ instance : has_coe ℤ_[p] ℚ_[p] := ⟨subtype.val⟩
83116
lemma mk_coe : ∀ (k : ℤ_[p]), (⟨↑k, k.2⟩ : ℤ_[p]) = k
84117
| ⟨_, _⟩ := rfl
85118

119+
/-- The inverse of a p-adic integer with norm equal to 1 is also a p-adic integer. Otherwise, the
120+
inverse is defined to be 0. -/
86121
def inv : ℤ_[p] → ℤ_[p]
87122
| ⟨k, _⟩ := if h : ∥k∥ = 1 then1/k, by simp [h]⟩ else 0
88123

src/data/padics/padic_norm.lean

Lines changed: 27 additions & 24 deletions
Original file line numberDiff line numberDiff line change
@@ -33,6 +33,7 @@ by taking (prime p) as a type class argument.
3333
## References
3434
3535
* [F. Q. Gouêva, *p-adic numbers*][gouvea1997]
36+
* [R. Y. Lewis, *A formal proof of Hensel's lemma over the p-adic integers*][lewis2019]
3637
* https://en.wikipedia.org/wiki/P-adic_number
3738
3839
## Tags
@@ -264,8 +265,26 @@ namespace padic_norm
264265

265266
section padic_norm
266267
open padic_val_rat
267-
variables (p : ℕ) [hp : p.prime]
268-
include hp
268+
variables (p : ℕ)
269+
270+
/--
271+
Unfolds the definition of the p-adic norm of `q` when `q ≠ 0`.
272+
-/
273+
@[simp] protected lemma eq_fpow_of_nonzero {q : ℚ} (hq : q ≠ 0) :
274+
padic_norm p q = p ^ (-(padic_val_rat p q)) :=
275+
by simp [hq, padic_norm]
276+
277+
/--
278+
The p-adic norm is nonnegative.
279+
-/
280+
protected lemma nonneg (q : ℚ) : padic_norm p q ≥ 0 :=
281+
if hq : q = 0 then by simp [hq]
282+
else
283+
begin
284+
unfold padic_norm; split_ifs,
285+
apply fpow_nonneg_of_nonneg,
286+
exact_mod_cast nat.zero_le _
287+
end
269288

270289
/--
271290
The p-adic norm of 0 is 0.
@@ -278,11 +297,13 @@ The p-adic norm of 1 is 1.
278297
@[simp] protected lemma one : padic_norm p 1 = 1 := by simp [padic_norm]
279298

280299
/--
281-
Unfolds the definition of the p-adic norm of `q` when `q ≠ 0`.
300+
The image of `padic_norm p` is {0} ∪ {p^(-n) | n ∈ ℤ}.
282301
-/
283-
@[simp] protected lemma eq_fpow_of_nonzero {q : ℚ} (hq : q ≠ 0) :
284-
padic_norm p q = p ^ (-(padic_val_rat p q)) :=
285-
by simp [hq, padic_norm]
302+
protected theorem image {q : ℚ} (hq : q ≠ 0) : ∃ n : ℤ, padic_norm p q = p ^ (-n) :=
303+
⟨ (padic_val_rat p q), by simp [padic_norm, hq] ⟩
304+
305+
variable [hp : p.prime]
306+
include hp
286307

287308
/--
288309
If `q ≠ 0`, then `padic_norm p q ≠ 0`.
@@ -313,18 +334,6 @@ begin
313334
exact_mod_cast hp.ne_zero
314335
end
315336

316-
/--
317-
The p-adic norm is nonnegative.
318-
-/
319-
protected lemma nonneg (q : ℚ) : padic_norm p q ≥ 0 :=
320-
if hq : q = 0 then by simp [hq]
321-
else
322-
begin
323-
unfold padic_norm; split_ifs,
324-
apply fpow_nonneg_of_nonneg,
325-
exact_mod_cast nat.zero_le _
326-
end
327-
328337
/--
329338
The p-adic norm is multiplicative.
330339
-/
@@ -438,12 +447,6 @@ begin
438447
assumption }
439448
end
440449

441-
/--
442-
The image of `padic_norm p` is {0} ∪ {p^(-n) | n ∈ ℤ}.
443-
-/
444-
protected theorem image {q : ℚ} (hq : q ≠ 0) : ∃ n : ℤ, padic_norm p q = p ^ (-n) :=
445-
⟨ (padic_val_rat p q), by simp [padic_norm, hq] ⟩
446-
447450
/--
448451
The p-adic norm is an absolute value: positive-definite and multiplicative, satisfying the triangle
449452
inequality.

0 commit comments

Comments
 (0)