Skip to content

Commit 53f5571

Browse files
chore: backports for leanprover/lean4#4814 (part 17) (#15429)
see #15245 Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com>
1 parent 4bb053f commit 53f5571

File tree

17 files changed

+277
-215
lines changed

17 files changed

+277
-215
lines changed

Mathlib/Algebra/CharP/ExpChar.lean

Lines changed: 16 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -91,6 +91,22 @@ theorem char_eq_expChar_iff (p q : ℕ) [hp : CharP R p] [hq : ExpChar R q] : p
9191
decide
9292
· exact ⟨fun hpq => hpq.symm ▸ hq_prime, fun _ => CharP.eq R hp hq_hchar⟩
9393

94+
/-- The exponential characteristic is a prime number or one.
95+
See also `CharP.char_is_prime_or_zero`. -/
96+
theorem expChar_is_prime_or_one (q : ℕ) [hq : ExpChar R q] : Nat.Prime q ∨ q = 1 := by
97+
cases hq with
98+
| zero => exact .inr rfl
99+
| prime hp => exact .inl hp
100+
101+
/-- The exponential characteristic is positive. -/
102+
theorem expChar_pos (q : ℕ) [ExpChar R q] : 0 < q := by
103+
rcases expChar_is_prime_or_one R q with h | rfl
104+
exacts [Nat.Prime.pos h, Nat.one_pos]
105+
106+
/-- Any power of the exponential characteristic is positive. -/
107+
theorem expChar_pow_pos (q : ℕ) [ExpChar R q] (n : ℕ) : 0 < q ^ n :=
108+
Nat.pos_pow_of_pos n (expChar_pos R q)
109+
94110
section Nontrivial
95111

96112
variable [Nontrivial R]
@@ -126,22 +142,6 @@ theorem char_prime_of_ne_zero {p : ℕ} [hp : CharP R p] (p_ne_zero : p ≠ 0) :
126142
· exact h
127143
· contradiction
128144

129-
/-- The exponential characteristic is a prime number or one.
130-
See also `CharP.char_is_prime_or_zero`. -/
131-
theorem expChar_is_prime_or_one (q : ℕ) [hq : ExpChar R q] : Nat.Prime q ∨ q = 1 := by
132-
cases hq with
133-
| zero => exact .inr rfl
134-
| prime hp => exact .inl hp
135-
136-
/-- The exponential characteristic is positive. -/
137-
theorem expChar_pos (q : ℕ) [ExpChar R q] : 0 < q := by
138-
rcases expChar_is_prime_or_one R q with h | rfl
139-
exacts [Nat.Prime.pos h, Nat.one_pos]
140-
141-
/-- Any power of the exponential characteristic is positive. -/
142-
theorem expChar_pow_pos (q : ℕ) [ExpChar R q] (n : ℕ) : 0 < q ^ n :=
143-
Nat.pos_pow_of_pos n (expChar_pos R q)
144-
145145
end NoZeroDivisors
146146

147147
end Nontrivial

Mathlib/Algebra/Polynomial/AlgebraMap.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -217,7 +217,7 @@ end CommSemiring
217217
section aeval
218218

219219
variable [CommSemiring R] [Semiring A] [CommSemiring A'] [Semiring B]
220-
variable [Algebra R A] [Algebra R A'] [Algebra R B]
220+
variable [Algebra R A] [Algebra R B]
221221
variable {p q : R[X]} (x : A)
222222

223223
/-- Given a valuation `x` of the variable in an `R`-algebra `A`, `aeval R A x` is

Mathlib/Algebra/Polynomial/Degree/Lemmas.lean

Lines changed: 15 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -305,19 +305,7 @@ end Ring
305305

306306
section NoZeroDivisors
307307

308-
variable [Semiring R] [NoZeroDivisors R] {p q : R[X]} {a : R}
309-
310-
theorem degree_mul_C (a0 : a ≠ 0) : (p * C a).degree = p.degree := by
311-
rw [degree_mul, degree_C a0, add_zero]
312-
313-
theorem degree_C_mul (a0 : a ≠ 0) : (C a * p).degree = p.degree := by
314-
rw [degree_mul, degree_C a0, zero_add]
315-
316-
theorem natDegree_mul_C (a0 : a ≠ 0) : (p * C a).natDegree = p.natDegree := by
317-
simp only [natDegree, degree_mul_C a0]
318-
319-
theorem natDegree_C_mul (a0 : a ≠ 0) : (C a * p).natDegree = p.natDegree := by
320-
simp only [natDegree, degree_C_mul a0]
308+
variable [Semiring R] {p q : R[X]} {a : R}
321309

322310
@[simp]
323311
lemma nextCoeff_C_mul_X_add_C (ha : a ≠ 0) (c : R) : nextCoeff (C a * X + C c) = c := by
@@ -337,6 +325,20 @@ lemma natDegree_eq_one : p.natDegree = 1 ↔ ∃ a ≠ 0, ∃ b, C a * X + C b =
337325
· rintro ⟨a, ha, b, rfl⟩
338326
simp [ha]
339327

328+
variable [NoZeroDivisors R]
329+
330+
theorem degree_mul_C (a0 : a ≠ 0) : (p * C a).degree = p.degree := by
331+
rw [degree_mul, degree_C a0, add_zero]
332+
333+
theorem degree_C_mul (a0 : a ≠ 0) : (C a * p).degree = p.degree := by
334+
rw [degree_mul, degree_C a0, zero_add]
335+
336+
theorem natDegree_mul_C (a0 : a ≠ 0) : (p * C a).natDegree = p.natDegree := by
337+
simp only [natDegree, degree_mul_C a0]
338+
339+
theorem natDegree_C_mul (a0 : a ≠ 0) : (C a * p).natDegree = p.natDegree := by
340+
simp only [natDegree, degree_C_mul a0]
341+
340342
theorem natDegree_comp : natDegree (p.comp q) = natDegree p * natDegree q := by
341343
by_cases q0 : q.natDegree = 0
342344
· rw [degree_le_zero_iff.mp (natDegree_eq_zero_iff_degree_le_zero.mp q0), comp_C, natDegree_C,

Mathlib/Algebra/Polynomial/Module/Basic.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -221,7 +221,7 @@ lemma equivPolynomial_single {S : Type*} [CommRing S] [Algebra R S] (n : ℕ) (x
221221
equivPolynomial (single R n x) = monomial n x := rfl
222222

223223
variable (R' : Type*) {M' : Type*} [CommRing R'] [AddCommGroup M'] [Module R' M']
224-
variable [Algebra R R'] [Module R M'] [IsScalarTower R R' M']
224+
variable [Module R M']
225225

226226
/-- The image of a polynomial under a linear map. -/
227227
noncomputable def map (f : M →ₗ[R] M') : PolynomialModule R M →ₗ[R] PolynomialModule R' M' :=
@@ -231,6 +231,8 @@ noncomputable def map (f : M →ₗ[R] M') : PolynomialModule R M →ₗ[R] Poly
231231
theorem map_single (f : M →ₗ[R] M') (i : ℕ) (m : M) : map R' f (single R i m) = single R' i (f m) :=
232232
Finsupp.mapRange_single (hf := f.map_zero)
233233

234+
variable [Algebra R R'] [IsScalarTower R R' M']
235+
234236
theorem map_smul (f : M →ₗ[R] M') (p : R[X]) (q : PolynomialModule R M) :
235237
map R' f (p • q) = p.map (algebraMap R R') • map R' f q := by
236238
apply induction_linear q

Mathlib/Algebra/Polynomial/Smeval.lean

Lines changed: 41 additions & 31 deletions
Original file line numberDiff line numberDiff line change
@@ -181,8 +181,16 @@ octonions), we replace the `[Semiring S]` with `[NonAssocSemiring S] [Pow S ℕ]
181181
-/
182182

183183
variable (R : Type*) [Semiring R] {p : R[X]} (r : R) (p q : R[X]) {S : Type*}
184-
[NonAssocSemiring S] [Module R S] [IsScalarTower R S S] [SMulCommClass R S S] [Pow S ℕ]
185-
[NatPowAssoc S] (x : S)
184+
[NonAssocSemiring S] [Module R S] [Pow S ℕ] (x : S)
185+
186+
theorem smeval_C_mul : (C r * p).smeval x = r • p.smeval x := by
187+
induction p using Polynomial.induction_on' with
188+
| h_add p q ph qh =>
189+
simp only [mul_add, smeval_add, ph, qh, smul_add]
190+
| h_monomial n b =>
191+
simp only [C_mul_monomial, smeval_monomial, mul_smul]
192+
193+
variable [NatPowAssoc S]
186194

187195
theorem smeval_at_natCast (q : ℕ[X]) : ∀(n : ℕ), q.smeval (n : S) = q.smeval n := by
188196
induction q using Polynomial.induction_on' with
@@ -206,13 +214,8 @@ theorem smeval_at_zero : p.smeval (0 : S) = (p.coeff 0) • (1 : S) := by
206214
| succ n => rw [coeff_monomial_succ, smeval_monomial, npow_add, npow_one, mul_zero, zero_smul,
207215
smul_zero]
208216

209-
theorem smeval_mul_X : (p * X).smeval x = p.smeval x * x := by
210-
induction p using Polynomial.induction_on' with
211-
| h_add p q ph qh =>
212-
simp only [add_mul, smeval_add, ph, qh]
213-
| h_monomial n a =>
214-
simp only [← monomial_one_one_eq_X, monomial_mul_monomial, smeval_monomial, mul_one, pow_succ',
215-
mul_assoc, npow_add, smul_mul_assoc, npow_one]
217+
section
218+
variable [SMulCommClass R S S]
216219

217220
theorem smeval_X_mul : (X * p).smeval x = x * p.smeval x := by
218221
induction p using Polynomial.induction_on' with
@@ -222,21 +225,6 @@ theorem smeval_X_mul : (X * p).smeval x = x * p.smeval x := by
222225
rw [← monomial_one_one_eq_X, monomial_mul_monomial, smeval_monomial, one_mul, npow_add,
223226
npow_one, ← mul_smul_comm, smeval_monomial]
224227

225-
theorem smeval_assoc_X_pow (m n : ℕ) :
226-
p.smeval x * x ^ m * x ^ n = p.smeval x * (x ^ m * x ^ n) := by
227-
induction p using Polynomial.induction_on' with
228-
| h_add p q ph qh =>
229-
simp only [smeval_add, ph, qh, add_mul]
230-
| h_monomial n a =>
231-
rw [smeval_monomial, smul_mul_assoc, smul_mul_assoc, npow_mul_assoc, ← smul_mul_assoc]
232-
233-
theorem smeval_mul_X_pow : ∀ (n : ℕ), (p * X^n).smeval x = p.smeval x * x^n
234-
| 0 => by
235-
simp only [npow_zero, mul_one]
236-
| n + 1 => by
237-
rw [npow_add, ← mul_assoc, npow_one, smeval_mul_X, smeval_mul_X_pow n, npow_add,
238-
← smeval_assoc_X_pow, npow_one]
239-
240228
theorem smeval_X_pow_assoc (m n : ℕ) :
241229
x ^ m * x ^ n * p.smeval x = x ^ m * (x ^ n * p.smeval x) := by
242230
induction p using Polynomial.induction_on' with
@@ -252,13 +240,6 @@ theorem smeval_X_pow_mul : ∀ (n : ℕ), (X^n * p).smeval x = x^n * p.smeval x
252240
rw [add_comm, npow_add, mul_assoc, npow_one, smeval_X_mul, smeval_X_pow_mul n, npow_add,
253241
smeval_X_pow_assoc, npow_one]
254242

255-
theorem smeval_C_mul : (C r * p).smeval x = r • p.smeval x := by
256-
induction p using Polynomial.induction_on' with
257-
| h_add p q ph qh =>
258-
simp only [mul_add, smeval_add, ph, qh, smul_add]
259-
| h_monomial n b =>
260-
simp only [C_mul_monomial, smeval_monomial, mul_smul]
261-
262243
theorem smeval_monomial_mul (n : ℕ) :
263244
(monomial n r * p).smeval x = r • (x ^ n * p.smeval x) := by
264245
induction p using Polynomial.induction_on' with
@@ -268,6 +249,35 @@ theorem smeval_monomial_mul (n : ℕ) :
268249
| h_monomial n a =>
269250
rw [smeval_monomial, monomial_mul_monomial, smeval_monomial, npow_add, mul_smul, mul_smul_comm]
270251

252+
end
253+
254+
variable [IsScalarTower R S S]
255+
256+
theorem smeval_mul_X : (p * X).smeval x = p.smeval x * x := by
257+
induction p using Polynomial.induction_on' with
258+
| h_add p q ph qh =>
259+
simp only [add_mul, smeval_add, ph, qh]
260+
| h_monomial n a =>
261+
simp only [← monomial_one_one_eq_X, monomial_mul_monomial, smeval_monomial, mul_one, pow_succ',
262+
mul_assoc, npow_add, smul_mul_assoc, npow_one]
263+
264+
theorem smeval_assoc_X_pow (m n : ℕ) :
265+
p.smeval x * x ^ m * x ^ n = p.smeval x * (x ^ m * x ^ n) := by
266+
induction p using Polynomial.induction_on' with
267+
| h_add p q ph qh =>
268+
simp only [smeval_add, ph, qh, add_mul]
269+
| h_monomial n a =>
270+
rw [smeval_monomial, smul_mul_assoc, smul_mul_assoc, npow_mul_assoc, ← smul_mul_assoc]
271+
272+
theorem smeval_mul_X_pow : ∀ (n : ℕ), (p * X^n).smeval x = p.smeval x * x^n
273+
| 0 => by
274+
simp only [npow_zero, mul_one]
275+
| n + 1 => by
276+
rw [npow_add, ← mul_assoc, npow_one, smeval_mul_X, smeval_mul_X_pow n, npow_add,
277+
← smeval_assoc_X_pow, npow_one]
278+
279+
variable [SMulCommClass R S S]
280+
271281
theorem smeval_mul : (p * q).smeval x = p.smeval x * q.smeval x := by
272282
induction p using Polynomial.induction_on' with
273283
| h_add r s hr hs =>

Mathlib/Algebra/Quaternion.lean

Lines changed: 52 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -89,11 +89,14 @@ theorem equivTuple_apply {R : Type*} (c₁ c₂ : R) (x : ℍ[R,c₁,c₂]) :
8989
@[simp]
9090
theorem mk.eta {R : Type*} {c₁ c₂} (a : ℍ[R,c₁,c₂]) : mk a.1 a.2 a.3 a.4 = a := rfl
9191

92-
variable {S T R : Type*} [CommRing R] {c₁ c₂ : R} (r x y z : R) (a b c : ℍ[R,c₁,c₂])
92+
variable {S T R : Type*} {c₁ c₂ : R} (r x y z : R) (a b c : ℍ[R,c₁,c₂])
9393

9494
instance [Subsingleton R] : Subsingleton ℍ[R, c₁, c₂] := (equivTuple c₁ c₂).subsingleton
9595
instance [Nontrivial R] : Nontrivial ℍ[R, c₁, c₂] := (equivTuple c₁ c₂).surjective.nontrivial
9696

97+
section Zero
98+
variable [Zero R]
99+
97100
/-- The imaginary part of a quaternion. -/
98101
def im (x : ℍ[R,c₁,c₂]) : ℍ[R,c₁,c₂] :=
99102
0, x.imI, x.imJ, x.imK⟩
@@ -159,6 +162,9 @@ theorem coe_zero : ((0 : R) : ℍ[R,c₁,c₂]) = 0 := rfl
159162

160163
instance : Inhabited ℍ[R,c₁,c₂] := ⟨0
161164

165+
section One
166+
variable [One R]
167+
162168
-- Porting note: removed `simps`, added simp lemmas manually
163169
instance : One ℍ[R,c₁,c₂] := ⟨⟨1, 0, 0, 0⟩⟩
164170

@@ -175,6 +181,11 @@ instance : One ℍ[R,c₁,c₂] := ⟨⟨1, 0, 0, 0⟩⟩
175181
@[simp, norm_cast]
176182
theorem coe_one : ((1 : R) : ℍ[R,c₁,c₂]) = 1 := rfl
177183

184+
end One
185+
end Zero
186+
section Add
187+
variable [Add R]
188+
178189
-- Porting note: removed `simps`, added simp lemmas manually
179190
instance : Add ℍ[R,c₁,c₂] :=
180191
fun a b => ⟨a.1 + b.1, a.2 + b.2, a.3 + b.3, a.4 + b.4⟩⟩
@@ -187,17 +198,27 @@ instance : Add ℍ[R,c₁,c₂] :=
187198

188199
@[simp] theorem add_imK : (a + b).imK = a.imK + b.imK := rfl
189200

190-
@[simp] theorem add_im : (a + b).im = a.im + b.im :=
191-
QuaternionAlgebra.ext (zero_add _).symm rfl rfl rfl
192-
193201
@[simp]
194202
theorem mk_add_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) :
195203
(mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) + mk b₁ b₂ b₃ b₄ = mk (a₁ + b₁) (a₂ + b₂) (a₃ + b₃) (a₄ + b₄) :=
196204
rfl
197205

206+
end Add
207+
208+
section AddZeroClass
209+
variable [AddZeroClass R]
210+
211+
@[simp] theorem add_im : (a + b).im = a.im + b.im :=
212+
QuaternionAlgebra.ext (zero_add _).symm rfl rfl rfl
213+
198214
@[simp, norm_cast]
199215
theorem coe_add : ((x + y : R) : ℍ[R,c₁,c₂]) = x + y := by ext <;> simp
200216

217+
end AddZeroClass
218+
219+
section Neg
220+
variable [Neg R]
221+
201222
-- Porting note: removed `simps`, added simp lemmas manually
202223
instance : Neg ℍ[R,c₁,c₂] := ⟨fun a => ⟨-a.1, -a.2, -a.3, -a.4⟩⟩
203224

@@ -209,13 +230,18 @@ instance : Neg ℍ[R,c₁,c₂] := ⟨fun a => ⟨-a.1, -a.2, -a.3, -a.4⟩⟩
209230

210231
@[simp] theorem neg_imK : (-a).imK = -a.imK := rfl
211232

212-
@[simp] theorem neg_im : (-a).im = -a.im :=
213-
QuaternionAlgebra.ext neg_zero.symm rfl rfl rfl
214-
215233
@[simp]
216234
theorem neg_mk (a₁ a₂ a₃ a₄ : R) : -(mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) = ⟨-a₁, -a₂, -a₃, -a₄⟩ :=
217235
rfl
218236

237+
end Neg
238+
239+
section AddGroup
240+
variable [AddGroup R]
241+
242+
@[simp] theorem neg_im : (-a).im = -a.im :=
243+
QuaternionAlgebra.ext neg_zero.symm rfl rfl rfl
244+
219245
@[simp, norm_cast]
220246
theorem coe_neg : ((-x : R) : ℍ[R,c₁,c₂]) = -x := by ext <;> simp
221247

@@ -254,6 +280,11 @@ theorem sub_self_im : a - a.im = a.re :=
254280
theorem sub_self_re : a - a.re = a.im :=
255281
QuaternionAlgebra.ext (sub_self _) (sub_zero _) (sub_zero _) (sub_zero _)
256282

283+
end AddGroup
284+
285+
section Ring
286+
variable [Ring R]
287+
257288
/-- Multiplication is given by
258289
259290
* `1 * x = x * 1 = x`;
@@ -290,11 +321,11 @@ theorem mk_mul_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) :
290321
a₁ * b₃ + c₁ * a₂ * b₄ + a₃ * b₁ - c₁ * a₄ * b₂, a₁ * b₄ + a₂ * b₃ - a₃ * b₂ + a₄ * b₁⟩ :=
291322
rfl
292323

293-
section
324+
end Ring
325+
section SMul
294326

295327
variable [SMul S R] [SMul T R] (s : S)
296328

297-
-- Porting note: Lean 4 auto drops the unused `[Ring R]` argument
298329
instance : SMul S ℍ[R,c₁,c₂] where smul s a := ⟨s • a.1, s • a.2, s • a.3, s • a.4
299330

300331
instance [SMul S T] [IsScalarTower S T R] : IsScalarTower S T ℍ[R,c₁,c₂] where
@@ -311,25 +342,28 @@ instance [SMulCommClass S T R] : SMulCommClass S T ℍ[R,c₁,c₂] where
311342

312343
@[simp] theorem smul_imK : (s • a).imK = s • a.imK := rfl
313344

314-
@[simp] theorem smul_im {S} [SMulZeroClass S R] (s : S) : (s • a).im = s • a.im :=
345+
@[simp] theorem smul_im {S} [CommRing R] [SMulZeroClass S R] (s : S) : (s • a).im = s • a.im :=
315346
QuaternionAlgebra.ext (smul_zero s).symm rfl rfl rfl
316347

317348
@[simp]
318349
theorem smul_mk (re im_i im_j im_k : R) :
319350
s • (⟨re, im_i, im_j, im_k⟩ : ℍ[R,c₁,c₂]) = ⟨s • re, s • im_i, s • im_j, s • im_k⟩ :=
320351
rfl
321352

322-
end
353+
end SMul
323354

324355
@[simp, norm_cast]
325-
theorem coe_smul [SMulZeroClass S R] (s : S) (r : R) :
356+
theorem coe_smul [Zero R] [SMulZeroClass S R] (s : S) (r : R) :
326357
(↑(s • r) : ℍ[R,c₁,c₂]) = s • (r : ℍ[R,c₁,c₂]) :=
327358
QuaternionAlgebra.ext rfl (smul_zero s).symm (smul_zero s).symm (smul_zero s).symm
328359

329-
instance : AddCommGroup ℍ[R,c₁,c₂] :=
360+
instance [AddCommGroup R] : AddCommGroup ℍ[R,c₁,c₂] :=
330361
(equivProd c₁ c₂).injective.addCommGroup _ rfl (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl)
331362
(fun _ _ ↦ rfl) (fun _ _ ↦ rfl)
332363

364+
section AddCommGroupWithOne
365+
variable [AddCommGroupWithOne R]
366+
333367
instance : AddCommGroupWithOne ℍ[R,c₁,c₂] where
334368
natCast n := ((n : R) : ℍ[R,c₁,c₂])
335369
natCast_zero := by simp
@@ -424,6 +458,11 @@ theorem coe_intCast (z : ℤ) : ↑(z : R) = (z : ℍ[R,c₁,c₂]) :=
424458
@[deprecated (since := "2024-04-17")]
425459
alias coe_int_cast := coe_intCast
426460

461+
end AddCommGroupWithOne
462+
463+
-- For the remainder of the file we assume `CommRing R`.
464+
variable [CommRing R]
465+
427466
instance instRing : Ring ℍ[R,c₁,c₂] where
428467
__ := inferInstanceAs (AddCommGroupWithOne ℍ[R,c₁,c₂])
429468
left_distrib _ _ _ := by ext <;> simp <;> ring

Mathlib/GroupTheory/HNNExtension.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -345,7 +345,7 @@ theorem unitsSMulGroup_snd (u : ℤˣ) (g : G) :
345345
(unitsSMulGroup φ d u g).2 = ((d.compl u).equiv g).2 := by
346346
rcases Int.units_eq_one_or u with rfl | rfl <;> rfl
347347

348-
variable {d} [DecidableEq G]
348+
variable {d}
349349

350350
/-- `Cancels u w` is a predicate expressing whether `t^u` cancels with some occurence
351351
of `t^-u` when when we multiply `t^u` by `w`. -/

0 commit comments

Comments
 (0)