Skip to content

Commit 921bb50

Browse files
fix: align Int.norm_eq_abs with its mathlib3 meaning (#11841)
1 parent a3021ae commit 921bb50

File tree

4 files changed

+8
-9
lines changed

4 files changed

+8
-9
lines changed

Mathlib/Analysis/Normed/Group/Basic.lean

Lines changed: 3 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -1917,9 +1917,8 @@ theorem norm_cast_real (m : ℤ) : ‖(m : ℝ)‖ = ‖m‖ :=
19171917
rfl
19181918
#align int.norm_cast_real Int.norm_cast_real
19191919

1920-
theorem norm_eq_abs (n : ℤ) : ‖n‖ = |n| :=
1921-
show ‖(n : ℝ)‖ = |n| by rw [Real.norm_eq_abs, cast_abs]
1922-
-- Porting note: I'm not sure why this isn't `rfl` anymore, but I suspect it's about coercions
1920+
theorem norm_eq_abs (n : ℤ) : ‖n‖ = |(n : ℝ)| :=
1921+
rfl
19231922
#align int.norm_eq_abs Int.norm_eq_abs
19241923

19251924
@[simp]
@@ -1930,7 +1929,7 @@ theorem _root_.NNReal.coe_natAbs (n : ℤ) : (n.natAbs : ℝ≥0) = ‖n‖₊ :
19301929
NNReal.eq <|
19311930
calc
19321931
((n.natAbs : ℝ≥0) : ℝ) = (n.natAbs : ℤ) := by simp only [Int.cast_ofNat, NNReal.coe_nat_cast]
1933-
_ = |n| := by simp only [Int.coe_natAbs, Int.cast_abs]
1932+
_ = |(n : ℝ)| := by simp only [Int.coe_natAbs, Int.cast_abs]
19341933
_ = ‖n‖ := (norm_eq_abs n).symm
19351934
#align nnreal.coe_nat_abs NNReal.coe_natAbs
19361935

Mathlib/Analysis/NormedSpace/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -101,8 +101,8 @@ instance NormedSpace.discreteTopology_zmultiples
101101
obtain ⟨k, rfl⟩ := AddSubgroup.mem_zmultiples_iff.mp hx
102102
rw [mem_preimage, mem_ball_zero_iff, AddSubgroup.coe_mk, mem_singleton_iff, Subtype.ext_iff,
103103
AddSubgroup.coe_mk, AddSubgroup.coe_zero, norm_zsmul ℚ k e, Int.norm_cast_rat,
104-
Int.norm_eq_abs, mul_lt_iff_lt_one_left (norm_pos_iff.mpr he), ←
105-
@Int.cast_one ℝ _, Int.cast_lt, Int.abs_lt_one_iff, smul_eq_zero, or_iff_left he]
104+
Int.norm_eq_abs, mul_lt_iff_lt_one_left (norm_pos_iff.mpr he), ← @Int.cast_one ℝ _,
105+
Int.cast_abs, Int.cast_lt, Int.abs_lt_one_iff, smul_eq_zero, or_iff_left he]
106106

107107
open NormedField
108108

Mathlib/NumberTheory/NumberField/Discriminant.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -309,8 +309,8 @@ theorem minkowskiBound_lt_boundOfDiscBdd : minkowskiBound K ↑1 < boundOfDiscBd
309309
coe_mul, ENNReal.coe_pow, coe_ofNat, show sqrt N = (1:ℝ≥0∞) * sqrt N by rw [one_mul]]
310310
gcongr
311311
· exact pow_le_one _ (by positivity) (by norm_num)
312-
· rw [sqrt_le_sqrt, ← NNReal.coe_le_coe, coe_nnnorm, Int.norm_eq_abs]
313-
exact Int.cast_le.mpr hK
312+
· rwa [sqrt_le_sqrt, ← NNReal.coe_le_coe, coe_nnnorm, Int.norm_eq_abs, ← Int.cast_abs,
313+
NNReal.coe_nat_cast, ← Int.cast_ofNat, Int.cast_le]
314314
· exact one_le_two
315315
· exact rank_le_rankOfDiscrBdd hK
316316

Mathlib/NumberTheory/NumberField/Embeddings.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -112,7 +112,7 @@ theorem finite_of_norm_le (B : ℝ) : {x : K | IsIntegral ℤ x ∧ ∀ φ : K
112112
exact minpoly.natDegree_le x
113113
rw [mem_Icc, ← abs_le, ← @Int.cast_le ℝ]
114114
refine (Eq.trans_le ?_ <| coeff_bdd_of_norm_le hx.2 i).trans (Nat.le_ceil _)
115-
rw [h_map_ℚ_minpoly, coeff_map, eq_intCast, Int.norm_cast_rat, Int.norm_eq_abs]
115+
rw [h_map_ℚ_minpoly, coeff_map, eq_intCast, Int.norm_cast_rat, Int.norm_eq_abs, Int.cast_abs]
116116
#align number_field.embeddings.finite_of_norm_le NumberField.Embeddings.finite_of_norm_le
117117

118118
/-- An algebraic integer whose conjugates are all of norm one is a root of unity. -/

0 commit comments

Comments
 (0)