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

Commit 671d57d

Browse files
committed
chore(number_theory/padics/padic_integers): golf the comm_ring instance (#15590)
This results in nicer definitional equalities that don't involve the application of a recursor. This also renames `padic_int.coe_coe` to `padic_int.coe_nat_cast` and `padic_int.coe_coe_int` to `padic_int.coe_int_cast` to match other similar lemmas in mathlib. Finally, this fixes the TODO comment ```lean -- TODO: define nat_cast/int_cast so that coe_coe and coe_coe_int are rfl ```
1 parent e35aa92 commit 671d57d

File tree

2 files changed

+47
-105
lines changed

2 files changed

+47
-105
lines changed

src/number_theory/padics/padic_integers.lean

Lines changed: 43 additions & 100 deletions
Original file line numberDiff line numberDiff line change
@@ -61,31 +61,37 @@ variables {p : ℕ} [fact p.prime]
6161

6262
instance : has_coe ℤ_[p] ℚ_[p] := ⟨subtype.val⟩
6363

64-
lemma ext {x y : ℤ_[p]} : (x : ℚ_[p]) = y → x = y := subtype.ext_iff_val.2
64+
lemma ext {x y : ℤ_[p]} : (x : ℚ_[p]) = y → x = y := subtype.ext
65+
66+
variables (p)
67+
68+
/-- The padic integers as a subring of the padics -/
69+
def subring : subring (ℚ_[p]) :=
70+
{ carrier := {x : ℚ_[p] | ∥x∥ ≤ 1},
71+
zero_mem' := by norm_num,
72+
one_mem' := by norm_num,
73+
add_mem' := λ x y hx hy, (padic_norm_e.nonarchimedean _ _).trans $ max_le_iff.2 ⟨hx, hy⟩,
74+
mul_mem' := λ x y hx hy, (padic_norm_e.mul _ _).trans_le $ mul_le_one hx (norm_nonneg _) hy,
75+
neg_mem' := λ x hx, (norm_neg _).trans_le hx }
76+
77+
@[simp] lemma mem_subring_iff {x : ℚ_[p]} : x ∈ subring p ↔ ∥x∥ ≤ 1 := iff.rfl
78+
79+
variables {p}
6580

6681
/-- Addition on ℤ_p is inherited from ℚ_p. -/
67-
instance : has_add ℤ_[p] :=
68-
⟨λ ⟨x, hx⟩ ⟨y, hy⟩, ⟨x+y,
69-
le_trans (padic_norm_e.nonarchimedean _ _) (max_le_iff.2 ⟨hx,hy⟩)⟩⟩
82+
instance : has_add ℤ_[p] := (by apply_instance : has_add (subring p))
7083

7184
/-- Multiplication on ℤ_p is inherited from ℚ_p. -/
72-
instance : has_mul ℤ_[p] :=
73-
⟨λ ⟨x, hx⟩ ⟨y, hy⟩, ⟨x*y,
74-
begin rw padic_norm_e.mul, apply mul_le_one; {assumption <|> apply norm_nonneg} end⟩⟩
85+
instance : has_mul ℤ_[p] := (by apply_instance : has_mul (subring p))
7586

7687
/-- Negation on ℤ_p is inherited from ℚ_p. -/
77-
instance : has_neg ℤ_[p] :=
78-
⟨λ ⟨x, hx⟩, ⟨-x, by simpa⟩⟩
88+
instance : has_neg ℤ_[p] := (by apply_instance : has_neg (subring p))
7989

8090
/-- Subtraction on ℤ_p is inherited from ℚ_p. -/
81-
instance : has_sub ℤ_[p] :=
82-
⟨λ ⟨x, hx⟩ ⟨y, hy⟩, ⟨x - y,
83-
by { rw sub_eq_add_neg, rw ← norm_neg at hy,
84-
exact le_trans (padic_norm_e.nonarchimedean _ _) (max_le_iff.2 ⟨hx, hy⟩) }⟩⟩
91+
instance : has_sub ℤ_[p] := (by apply_instance : has_sub (subring p))
8592

8693
/-- Zero on ℤ_p is inherited from ℚ_p. -/
87-
instance : has_zero ℤ_[p] :=
88-
⟨⟨0, by norm_num⟩⟩
94+
instance : has_zero ℤ_[p] := (by apply_instance : has_zero (subring p))
8995

9096
instance : inhabited ℤ_[p] := ⟨0
9197

@@ -97,69 +103,28 @@ instance : has_one ℤ_[p] :=
97103

98104
@[simp] lemma val_eq_coe (z : ℤ_[p]) : z.val = z := rfl
99105

100-
@[simp, norm_cast] lemma coe_add : ∀ (z1 z2 : ℤ_[p]), ((z1 + z2 : ℤ_[p]) : ℚ_[p]) = z1 + z2
101-
| ⟨_, _⟩ ⟨_, _⟩ := rfl
102-
103-
@[simp, norm_cast] lemma coe_mul : ∀ (z1 z2 : ℤ_[p]), ((z1 * z2 : ℤ_[p]) : ℚ_[p]) = z1 * z2
104-
| ⟨_, _⟩ ⟨_, _⟩ := rfl
105-
106-
@[simp, norm_cast] lemma coe_neg : ∀ (z1 : ℤ_[p]), ((-z1 : ℤ_[p]) : ℚ_[p]) = -z1
107-
| ⟨_, _⟩ := rfl
108-
109-
@[simp, norm_cast] lemma coe_sub : ∀ (z1 z2 : ℤ_[p]), ((z1 - z2 : ℤ_[p]) : ℚ_[p]) = z1 - z2
110-
| ⟨_, _⟩ ⟨_, _⟩ := rfl
111-
106+
@[simp, norm_cast] lemma coe_add (z1 z2 : ℤ_[p]) : ((z1 + z2 : ℤ_[p]) : ℚ_[p]) = z1 + z2 := rfl
107+
@[simp, norm_cast] lemma coe_mul (z1 z2 : ℤ_[p]) : ((z1 * z2 : ℤ_[p]) : ℚ_[p]) = z1 * z2 := rfl
108+
@[simp, norm_cast] lemma coe_neg (z1 : ℤ_[p]) : ((-z1 : ℤ_[p]) : ℚ_[p]) = -z1 := rfl
109+
@[simp, norm_cast] lemma coe_sub (z1 z2 : ℤ_[p]) : ((z1 - z2 : ℤ_[p]) : ℚ_[p]) = z1 - z2 := rfl
112110
@[simp, norm_cast] lemma coe_one : ((1 : ℤ_[p]) : ℚ_[p]) = 1 := rfl
113-
114111
@[simp, norm_cast] lemma coe_zero : ((0 : ℤ_[p]) : ℚ_[p]) = 0 := rfl
115112

116113
instance : add_comm_group ℤ_[p] :=
117-
by refine_struct
118-
{ add := (+),
119-
neg := has_neg.neg,
120-
zero := (0 : ℤ_[p]),
121-
sub := has_sub.sub,
122-
nsmul := @nsmul_rec _ ⟨(0 : ℤ_[p])⟩ ⟨(+)⟩,
123-
zsmul := @zsmul_rec _ ⟨(0 : ℤ_[p])⟩ ⟨(+)⟩ ⟨has_neg.neg⟩ };
124-
intros; try { refl }; ext; simp; ring
125-
126-
instance : add_group_with_one ℤ_[p] :=
127-
-- TODO: define nat_cast/int_cast so that coe_coe and coe_coe_int are rfl
128-
{ one := 1, .. padic_int.add_comm_group }
129-
130-
instance : ring ℤ_[p] :=
131-
by refine_struct
132-
{ add := (+),
133-
mul := (*),
134-
neg := has_neg.neg,
135-
zero := (0 : ℤ_[p]),
136-
one := 1,
137-
sub := has_sub.sub,
138-
npow := @npow_rec _ ⟨(1 : ℤ_[p])⟩ ⟨(*)⟩,
139-
.. padic_int.add_group_with_one };
140-
intros; try { refl }; ext; simp; ring
141-
142-
@[simp, norm_cast] lemma coe_coe : ∀ n : ℕ, ((n : ℤ_[p]) : ℚ_[p]) = n
143-
| 0 := rfl
144-
| (k+1) := by simp [coe_coe]
145-
146-
@[simp, norm_cast] lemma coe_coe_int : ∀ (z : ℤ), ((z : ℤ_[p]) : ℚ_[p]) = z
147-
| (int.of_nat n) := by simp
148-
| -[1+n] := by simp
114+
(by apply_instance : add_comm_group (subring p))
115+
116+
instance : comm_ring ℤ_[p] :=
117+
(by apply_instance : comm_ring (subring p))
118+
119+
@[simp, norm_cast] lemma coe_nat_cast (n : ℕ) : ((n : ℤ_[p]) : ℚ_[p]) = n := rfl
120+
@[simp, norm_cast] lemma coe_int_cast (z : ℤ) : ((z : ℤ_[p]) : ℚ_[p]) = z := rfl
149121

150122
/-- The coercion from ℤ[p] to ℚ[p] as a ring homomorphism. -/
151-
def coe.ring_hom : ℤ_[p] →+* ℚ_[p] :=
152-
{ to_fun := (coe : ℤ_[p] → ℚ_[p]),
153-
map_zero' := rfl,
154-
map_one' := rfl,
155-
map_mul' := coe_mul,
156-
map_add' := coe_add }
123+
def coe.ring_hom : ℤ_[p] →+* ℚ_[p] := (subring p).subtype
157124

158-
@[simp, norm_cast] lemma coe_pow (x : ℤ_[p]) (n : ℕ) : (↑(x^n) : ℚ_[p]) = (↑x : ℚ_[p])^n :=
159-
(coe.ring_hom : ℤ_[p] →+* ℚ_[p]).map_pow x n
125+
@[simp, norm_cast] lemma coe_pow (x : ℤ_[p]) (n : ℕ) : (↑(x^n) : ℚ_[p]) = (↑x : ℚ_[p])^n := rfl
160126

161-
@[simp] lemma mk_coe : ∀ (k : ℤ_[p]), (⟨k, k.2⟩ : ℤ_[p]) = k
162-
| ⟨_, _⟩ := rfl
127+
@[simp] lemma mk_coe (k : ℤ_[p]) : (⟨k, k.2⟩ : ℤ_[p]) = k := subtype.coe_eta _ _
163128

164129
/-- The inverse of a p-adic integer with norm equal to 1 is also a p-adic integer. Otherwise, the
165130
inverse is defined to be 0. -/
@@ -212,29 +177,14 @@ instance : has_norm ℤ_[p] := ⟨λ z, ∥(z : ℚ_[p])∥⟩
212177

213178
variables {p}
214179

215-
protected lemma mul_comm : ∀ z1 z2 : ℤ_[p], z1*z2 = z2*z1
216-
| ⟨q1, h1⟩ ⟨q2, h2⟩ := show (⟨q1*q2, _⟩ : ℤ_[p]) = ⟨q2*q1, _⟩, by simp [_root_.mul_comm]
217-
218-
protected lemma zero_ne_one : (0 : ℤ_[p]) ≠ 1 :=
219-
show (⟨(0 : ℚ_[p]), _⟩ : ℤ_[p]) ≠ ⟨(1 : ℚ_[p]), _⟩, from mt subtype.ext_iff_val.1 zero_ne_one
220-
221-
protected lemma eq_zero_or_eq_zero_of_mul_eq_zero :
222-
∀ (a b : ℤ_[p]), a * b = 0 → a = 0 ∨ b = 0
223-
| ⟨a, ha⟩ ⟨b, hb⟩ := λ h : (⟨a * b, _⟩ : ℤ_[p]) = ⟨0, _⟩,
224-
have a * b = 0, from subtype.ext_iff_val.1 h,
225-
(mul_eq_zero.1 this).elim
226-
(λ h1, or.inl (by simp [h1]; refl))
227-
(λ h2, or.inr (by simp [h2]; refl))
228-
229180
lemma norm_def {z : ℤ_[p]} : ∥z∥ = ∥(z : ℚ_[p])∥ := rfl
230181

231182
variables (p)
232183

233184
instance : normed_comm_ring ℤ_[p] :=
234185
{ dist_eq := λ ⟨_, _⟩ ⟨_, _⟩, rfl,
235186
norm_mul := by simp [norm_def],
236-
mul_comm := padic_int.mul_comm,
237-
norm := norm, .. padic_int.ring, .. padic_int.metric_space p }
187+
norm := norm, .. padic_int.comm_ring, .. padic_int.metric_space p }
238188

239189
instance : norm_one_class ℤ_[p] := ⟨norm_def.trans norm_one⟩
240190

@@ -247,18 +197,15 @@ instance is_absolute_value : is_absolute_value (λ z : ℤ_[p], ∥z∥) :=
247197
variables {p}
248198

249199
instance : is_domain ℤ_[p] :=
250-
{ eq_zero_or_eq_zero_of_mul_eq_zero := λ x y, padic_int.eq_zero_or_eq_zero_of_mul_eq_zero x y,
251-
exists_pair_ne := ⟨0, 1, padic_int.zero_ne_one⟩,
252-
.. padic_int.normed_comm_ring p }
200+
function.injective.is_domain (subring p).subtype subtype.coe_injective
253201

254202
end padic_int
255203

256204
namespace padic_int
257205
/-! ### Norm -/
258206
variables {p : ℕ} [fact p.prime]
259207

260-
lemma norm_le_one : ∀ z : ℤ_[p], ∥z∥ ≤ 1
261-
| ⟨_, h⟩ := h
208+
lemma norm_le_one (z : ℤ_[p]) : ∥z∥ ≤ 1 := z.2
262209

263210
@[simp] lemma norm_mul (z1 z2 : ℤ_[p]) : ∥z1 * z2∥ = ∥z1∥ * ∥z2∥ :=
264211
by simp [norm_def]
@@ -267,11 +214,10 @@ by simp [norm_def]
267214
| 0 := by simp
268215
| (k+1) := by { rw [pow_succ, pow_succ, norm_mul], congr, apply norm_pow }
269216

270-
theorem nonarchimedean : ∀ (q r : ℤ_[p]), ∥q + r∥ ≤ max (∥q∥) (∥r∥)
271-
| ⟨_, _⟩ ⟨_, _⟩ := padic_norm_e.nonarchimedean _ _
217+
theorem nonarchimedean (q r : ℤ_[p]) : ∥q + r∥ ≤ max (∥q∥) (∥r∥) := padic_norm_e.nonarchimedean _ _
272218

273-
theorem norm_add_eq_max_of_ne : ∀ {q r : ℤ_[p]}, ∥q∥ ≠ ∥r∥ → ∥q+r∥ = max (∥q∥) (∥r∥)
274-
| ⟨_, _⟩ ⟨_, _⟩ := padic_norm_e.add_eq_max_of_ne
219+
theorem norm_add_eq_max_of_ne {q r : ℤ_[p]} : ∥q∥ ≠ ∥r∥ → ∥q+r∥ = max (∥q∥) (∥r∥) :=
220+
padic_norm_e.add_eq_max_of_ne
275221

276222
lemma norm_eq_of_norm_add_lt_right {z1 z2 : ℤ_[p]}
277223
(h : ∥z1 + z2∥ < ∥z2∥) : ∥z1∥ = ∥z2∥ :=
@@ -292,12 +238,9 @@ by simp [norm_def]
292238
@[simp] lemma norm_eq_padic_norm {q : ℚ_[p]} (hq : ∥q∥ ≤ 1) :
293239
@norm ℤ_[p] _ ⟨q, hq⟩ = ∥q∥ := rfl
294240

295-
@[simp] lemma norm_p : ∥(p : ℤ_[p])∥ = p⁻¹ :=
296-
show ∥((p : ℤ_[p]) : ℚ_[p])∥ = p⁻¹, by exact_mod_cast padic_norm_e.norm_p
241+
@[simp] lemma norm_p : ∥(p : ℤ_[p])∥ = p⁻¹ := padic_norm_e.norm_p
297242

298-
@[simp] lemma norm_p_pow (n : ℕ) : ∥(p : ℤ_[p])^n∥ = p^(-n:ℤ) :=
299-
show ∥((p^n : ℤ_[p]) : ℚ_[p])∥ = p^(-n:ℤ),
300-
by { convert padic_norm_e.norm_p_pow n, simp, }
243+
@[simp] lemma norm_p_pow (n : ℕ) : ∥(p : ℤ_[p])^n∥ = p^(-n:ℤ) := padic_norm_e.norm_p_pow n
301244

302245
private def cau_seq_to_rat_cau_seq (f : cau_seq ℤ_[p] norm) :
303246
cau_seq ℚ_[p] (λ a, ∥a∥) :=

src/number_theory/padics/ring_homs.lean

Lines changed: 4 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -81,7 +81,7 @@ lemma is_unit_denom (r : ℚ) (h : ∥(r : ℚ_[p])∥ ≤ 1) : is_unit (r.denom
8181
begin
8282
rw is_unit_iff,
8383
apply le_antisymm (r.denom : ℤ_[p]).2,
84-
rw [← not_lt, val_eq_coe, coe_coe],
84+
rw [← not_lt, val_eq_coe, coe_nat_cast],
8585
intro norm_denom_lt,
8686
have hr : ∥(r * r.denom : ℚ_[p])∥ = ∥(r.num : ℚ_[p])∥,
8787
{ rw_mod_cast @rat.mul_denom_eq_num r, refl, },
@@ -126,7 +126,7 @@ begin
126126
simp only [sub_mul, int.cast_coe_nat, ring_hom.eq_int_cast, int.cast_mul,
127127
sub_left_inj, int.cast_sub],
128128
apply subtype.coe_injective,
129-
simp only [coe_mul, subtype.coe_mk, coe_coe],
129+
simp only [coe_mul, subtype.coe_mk, coe_nat_cast],
130130
rw_mod_cast @rat.mul_denom_eq_num r, refl },
131131
exact norm_sub_mod_part_aux r h
132132
end
@@ -182,7 +182,7 @@ begin
182182
lift n to ℕ using hzn,
183183
use n,
184184
split, {exact_mod_cast hnp},
185-
simp only [norm_def, coe_sub, subtype.coe_mk, coe_coe] at hn ⊢,
185+
simp only [norm_def, coe_sub, subtype.coe_mk, coe_nat_cast] at hn ⊢,
186186
rw show (x - n : ℚ_[p]) = (x - r) + (r - n), by ring,
187187
apply lt_of_le_of_lt (padic_norm_e.nonarchimedean _ _),
188188
apply max_lt hr,
@@ -587,8 +587,7 @@ begin
587587
apply lt_trans _ hε',
588588
change ↑(padic_norm_e _) < _,
589589
norm_cast,
590-
convert hN _ hn,
591-
simp [nth_hom, lim_nth_hom, nth_hom_seq, of_int_seq],
590+
exact hN _ hn,
592591
end
593592

594593
lemma lim_nth_hom_zero : lim_nth_hom f_compat 0 = 0 :=

0 commit comments

Comments
 (0)