Skip to content

Commit 3646044

Browse files
committed
chore(LinearAlgebra,BigOperators): drop some decidability assumptions (#10237)
Found by a linter from #10235.
1 parent a619d25 commit 3646044

File tree

5 files changed

+45
-33
lines changed

5 files changed

+45
-33
lines changed

Mathlib/Algebra/BigOperators/Ring.lean

Lines changed: 3 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -73,12 +73,13 @@ section CommSemiring
7373
variable [CommSemiring α]
7474

7575
section DecidableEq
76-
variable [DecidableEq ι] [∀ i, DecidableEq (κ i)]
76+
variable [DecidableEq ι]
7777

7878
/-- The product over a sum can be written as a sum over the product of sets, `Finset.Pi`.
7979
`Finset.prod_univ_sum` is an alternative statement when the product is over `univ`. -/
8080
lemma prod_sum (s : Finset ι) (t : ∀ i, Finset (κ i)) (f : ∀ i, κ i → α) :
8181
∏ a in s, ∑ b in t a, f a b = ∑ p in s.pi t, ∏ x in s.attach, f x.1 (p x.1 x.2) := by
82+
classical
8283
induction' s using Finset.induction with a s ha ih
8384
· rw [pi_empty, sum_singleton]
8485
rfl
@@ -121,7 +122,7 @@ lemma sum_prod_piFinset {κ : Type*} [Fintype ι] (s : Finset κ) (g : ι → κ
121122

122123
lemma sum_pow' (s : Finset ι) (f : ι → α) (n : ℕ) :
123124
(∑ a in s, f a) ^ n = ∑ p in piFinset fun _i : Fin n ↦ s, ∏ i, f (p i) := by
124-
classical convert @prod_univ_sum (Fin n) _ _ _ _ _ _ (fun _i ↦ s) fun _i d ↦ f d; simp
125+
classical convert @prod_univ_sum (Fin n) _ _ _ _ _ (fun _i ↦ s) fun _i d ↦ f d; simp
125126

126127
/-- The product of `f a + g a` over all of `s` is the sum over the powerset of `s` of the product of
127128
`f` over a subset `t` times the product of `g` over the complement of `t` -/

Mathlib/LinearAlgebra/DFinsupp.lean

Lines changed: 27 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -37,14 +37,16 @@ function with finite support, module, linear algebra
3737

3838
variable {ι : Type*} {R : Type*} {S : Type*} {M : ι → Type*} {N : Type*}
3939

40-
variable [dec_ι : DecidableEq ι]
41-
4240
namespace DFinsupp
4341

4442
variable [Semiring R] [∀ i, AddCommMonoid (M i)] [∀ i, Module R (M i)]
4543

4644
variable [AddCommMonoid N] [Module R N]
4745

46+
section DecidableEq
47+
48+
variable [DecidableEq ι]
49+
4850
/-- `DFinsupp.mk` as a `LinearMap`. -/
4951
def lmk (s : Finset ι) : (∀ i : (↑s : Set ι), M i) →ₗ[R] Π₀ i, M i where
5052
toFun := mk s
@@ -137,8 +139,7 @@ def lsum [Semiring S] [Module S N] [SMulCommClass R S N] :
137139
(∀ i, M i →ₗ[R] N) ≃ₗ[S] (Π₀ i, M i) →ₗ[R] N where
138140
toFun F :=
139141
{ toFun := sumAddHom fun i => (F i).toAddMonoidHom
140-
-- Porting note: needed to (β := M) hint below
141-
map_add' := (DFinsupp.liftAddHom (β := M) (fun (i : ι) => (F i).toAddMonoidHom)).map_add
142+
map_add' := (DFinsupp.liftAddHom fun (i : ι) => (F i).toAddMonoidHom).map_add
142143
map_smul' := fun c f => by
143144
dsimp
144145
apply DFinsupp.induction f
@@ -174,12 +175,13 @@ theorem lsum_single [Semiring S] [Module S N] [SMulCommClass R S N] (F : ∀ i,
174175

175176
end Lsum
176177

178+
end DecidableEq
179+
177180
/-! ### Bundled versions of `DFinsupp.mapRange`
178181
179182
The names should match the equivalent bundled `Finsupp.mapRange` definitions.
180183
-/
181184

182-
183185
section mapRange
184186

185187
variable {β β₁ β₂ : ι → Type*}
@@ -217,13 +219,10 @@ theorem mapRange.linearMap_comp (f : ∀ i, β₁ i →ₗ[R] β₂ i) (f₂ :
217219
(fun i => (f i).map_zero) (fun i => (f₂ i).map_zero) (by simp)
218220
#align dfinsupp.map_range.linear_map_comp DFinsupp.mapRange.linearMap_comp
219221

220-
theorem sum_mapRange_index.linearMap [∀ (i : ι) (x : β₁ i), Decidable (x ≠ 0)]
221-
[∀ (i : ι) (x : β₂ i), Decidable (x ≠ 0)] {f : ∀ i, β₁ i →ₗ[R] β₂ i} {h : ∀ i, β₂ i →ₗ[R] N}
222-
{l : Π₀ i, β₁ i} :
223-
-- Porting note: Needed to add (M := ...) below
224-
(DFinsupp.lsum ℕ (M := β₂)) h (mapRange.linearMap f l)
225-
= (DFinsupp.lsum ℕ (M := β₁)) (fun i => (h i).comp (f i)) l := by
226-
simpa [DFinsupp.sumAddHom_apply] using sum_mapRange_index fun i => by simp
222+
theorem sum_mapRange_index.linearMap [DecidableEq ι] {f : ∀ i, β₁ i →ₗ[R] β₂ i}
223+
{h : ∀ i, β₂ i →ₗ[R] N} {l : Π₀ i, β₁ i} :
224+
DFinsupp.lsum ℕ h (mapRange.linearMap f l) = DFinsupp.lsum ℕ (fun i => (h i).comp (f i)) l := by
225+
classical simpa [DFinsupp.sumAddHom_apply] using sum_mapRange_index fun i => by simp
227226
#align dfinsupp.sum_map_range_index.linear_map DFinsupp.sum_mapRange_index.linearMap
228227

229228
/-- `DFinsupp.mapRange.linearMap` as a `LinearEquiv`. -/
@@ -289,6 +288,10 @@ variable [Semiring R] [AddCommMonoid N] [Module R N]
289288

290289
open DFinsupp
291290

291+
section DecidableEq
292+
293+
variable [DecidableEq ι]
294+
292295
theorem dfinsupp_sum_mem {β : ι → Type*} [∀ i, Zero (β i)] [∀ (i) (x : β i), Decidable (x ≠ 0)]
293296
(S : Submodule R N) (f : Π₀ i, β i) (g : ∀ i, β i → N) (h : ∀ c, f c ≠ 0 → g c (f c) ∈ S) :
294297
f.sum g ∈ S :=
@@ -356,6 +359,16 @@ theorem mem_iSup_iff_exists_dfinsupp' (p : ι → Submodule R N) [∀ (i) (x : p
356359
LinearMap.toAddMonoidHom_coe, coeSubtype]
357360
#align submodule.mem_supr_iff_exists_dfinsupp' Submodule.mem_iSup_iff_exists_dfinsupp'
358361

362+
theorem mem_biSup_iff_exists_dfinsupp (p : ι → Prop) [DecidablePred p] (S : ι → Submodule R N)
363+
(x : N) :
364+
(x ∈ ⨆ (i) (_ : p i), S i) ↔
365+
∃ f : Π₀ i, S i,
366+
DFinsupp.lsum ℕ (M := fun i ↦ ↥(S i)) (fun i => (S i).subtype) (f.filter p) = x :=
367+
SetLike.ext_iff.mp (biSup_eq_range_dfinsupp_lsum p S) x
368+
#align submodule.mem_bsupr_iff_exists_dfinsupp Submodule.mem_biSup_iff_exists_dfinsupp
369+
370+
end DecidableEq
371+
359372
lemma mem_iSup_iff_exists_finsupp (p : ι → Submodule R N) (x : N) :
360373
x ∈ iSup p ↔ ∃ (f : ι →₀ N), (∀ i, f i ∈ p i) ∧ (f.sum fun _i xi ↦ xi) = x := by
361374
classical
@@ -366,14 +379,6 @@ lemma mem_iSup_iff_exists_finsupp (p : ι → Submodule R N) (x : N) :
366379
· ext; simp
367380
· simp [Finsupp.mem_support_iff.mp hi]
368381

369-
theorem mem_biSup_iff_exists_dfinsupp (p : ι → Prop) [DecidablePred p] (S : ι → Submodule R N)
370-
(x : N) :
371-
(x ∈ ⨆ (i) (_ : p i), S i) ↔
372-
∃ f : Π₀ i, S i,
373-
DFinsupp.lsum ℕ (M := fun i ↦ ↥(S i)) (fun i => (S i).subtype) (f.filter p) = x :=
374-
SetLike.ext_iff.mp (biSup_eq_range_dfinsupp_lsum p S) x
375-
#align submodule.mem_bsupr_iff_exists_dfinsupp Submodule.mem_biSup_iff_exists_dfinsupp
376-
377382
open BigOperators
378383

379384
theorem mem_iSup_finset_iff_exists_sum {s : Finset ι} (p : ι → Submodule R N) (a : N) :
@@ -421,7 +426,7 @@ open DFinsupp
421426

422427
section Semiring
423428

424-
variable [Semiring R] [AddCommMonoid N] [Module R N]
429+
variable [DecidableEq ι] [Semiring R] [AddCommMonoid N] [Module R N]
425430

426431
/-- Independence of a family of submodules can be expressed as a quantifier over `DFinsupp`s.
427432
@@ -477,7 +482,7 @@ end Semiring
477482

478483
section Ring
479484

480-
variable [Ring R] [AddCommGroup N] [Module R N]
485+
variable [DecidableEq ι] [Ring R] [AddCommGroup N] [Module R N]
481486

482487
/- If `DFinsupp.sumAddHom` applied with `AddSubmonoid.subtype` is injective then the additive
483488
subgroups are independent. -/

Mathlib/LinearAlgebra/Matrix/PosDef.lean

Lines changed: 3 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -262,8 +262,9 @@ lemma eigenvalues_self_mul_conjTranspose_nonneg (A : Matrix m n 𝕜) [Decidable
262262
(posSemidef_self_mul_conjTranspose _).eigenvalues_nonneg _
263263

264264
/-- A matrix is positive semidefinite if and only if it has the form `Bᴴ * B` for some `B`. -/
265-
lemma posSemidef_iff_eq_transpose_mul_self [DecidableEq n] {A : Matrix n n 𝕜} :
265+
lemma posSemidef_iff_eq_transpose_mul_self {A : Matrix n n 𝕜} :
266266
PosSemidef A ↔ ∃ (B : Matrix n n 𝕜), A = Bᴴ * B := by
267+
classical
267268
refine ⟨fun hA ↦ ⟨hA.sqrt, ?_⟩, fun ⟨B, hB⟩ ↦ (hB ▸ posSemidef_conjTranspose_mul_self B)⟩
268269
simp_rw [← PosSemidef.sq_sqrt hA, pow_two]
269270
rw [hA.posSemidef_sqrt.1]
@@ -276,7 +277,7 @@ lemma IsHermitian.posSemidef_of_eigenvalues_nonneg [DecidableEq n] {A : Matrix n
276277
simpa using h i
277278

278279
/-- For `A` positive semidefinite, we have `x⋆ A x = 0` iff `A x = 0`. -/
279-
theorem PosSemidef.dotProduct_mulVec_zero_iff [DecidableEq n]
280+
theorem PosSemidef.dotProduct_mulVec_zero_iff
280281
{A : Matrix n n 𝕜} (hA : PosSemidef A) (x : n → 𝕜) :
281282
star x ⬝ᵥ mulVec A x = 0 ↔ mulVec A x = 0 := by
282283
constructor

Mathlib/LinearAlgebra/Matrix/Spectrum.lean

Lines changed: 9 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -24,16 +24,18 @@ spectral theorem, diagonalization theorem
2424

2525
namespace Matrix
2626

27-
variable {𝕜 : Type*} [IsROrC 𝕜] {n : Type*} [Fintype n] [DecidableEq n]
27+
variable {𝕜 : Type*} [IsROrC 𝕜] {n : Type*} [Fintype n]
2828

2929
variable {A : Matrix n n 𝕜}
3030

31-
open scoped Matrix
32-
3331
open scoped BigOperators
3432

3533
namespace IsHermitian
3634

35+
section DecidableEq
36+
37+
variable [DecidableEq n]
38+
3739
variable (hA : A.IsHermitian)
3840

3941
/-- The eigenvalues of a hermitian matrix, indexed by `Fin (Fintype.card n)` where `n` is the index
@@ -172,9 +174,12 @@ lemma mulVec_eigenvectorBasis (i : n) :
172174
convert this using 1
173175
rw [mul_comm, Pi.smul_apply, IsROrC.real_smul_eq_coe_mul, hA.eigenvectorMatrix_apply]
174176

177+
end DecidableEq
178+
175179
/-- A nonzero Hermitian matrix has an eigenvector with nonzero eigenvalue. -/
176-
lemma exists_eigenvector_of_ne_zero (h_ne : A ≠ 0) :
180+
lemma exists_eigenvector_of_ne_zero (hA : IsHermitian A) (h_ne : A ≠ 0) :
177181
∃ (v : n → 𝕜) (t : ℝ), t ≠ 0 ∧ v ≠ 0 ∧ mulVec A v = t • v := by
182+
classical
178183
have : hA.eigenvalues ≠ 0
179184
· contrapose! h_ne
180185
have := hA.spectral_theorem'

Mathlib/Tactic/NormNum/BigOperators.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -171,9 +171,9 @@ def Multiset.ProveZeroOrConsResult.eq_trans {α : Q(Type u)} {s t : Q(Multiset $
171171
| .zero pf => .zero q(Eq.trans $eq $pf)
172172
| .cons a s' pf => .cons a s' q(Eq.trans $eq $pf)
173173

174-
lemma Multiset.insert_eq_cons {α : Type*} [DecidableEq α] (a : α) (s : Multiset α) :
175-
insert a s = Multiset.cons a s := by
176-
ext; simp
174+
lemma Multiset.insert_eq_cons {α : Type*} (a : α) (s : Multiset α) :
175+
insert a s = Multiset.cons a s :=
176+
rfl
177177

178178
lemma Multiset.range_zero' {n : ℕ} (pn : NormNum.IsNat n 0) :
179179
Multiset.range n = 0 := by rw [pn.out, Nat.cast_zero, Multiset.range_zero]

0 commit comments

Comments
 (0)