Skip to content

Commit 51b8a38

Browse files
committed
chore(*): fix some Decidable*/Fintype` assumptions (#24986)
1 parent 19af017 commit 51b8a38

File tree

4 files changed

+8
-5
lines changed

4 files changed

+8
-5
lines changed

Mathlib/Algebra/BigOperators/Ring/Nat.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -30,9 +30,10 @@ lemma odd_sum_iff_odd_card_odd {s : Finset ι} (f : ι → ℕ) :
3030
Odd (∑ i ∈ s, f i) ↔ Odd #{x ∈ s | Odd (f x)} := by
3131
simp only [← Nat.not_even_iff_odd, even_sum_iff_even_card_odd]
3232

33-
theorem card_preimage_eq_sum_card_image_eq {M : Type*} [DecidableEq M] {f : ι → M} {s : Finset M}
33+
theorem card_preimage_eq_sum_card_image_eq {M : Type*} {f : ι → M} {s : Finset M}
3434
(hb : ∀ b ∈ s, Set.Finite {a | f a = b}) :
3535
Nat.card (f⁻¹' s) = ∑ b ∈ s, Nat.card {a // f a = b} := by
36+
classical
3637
-- `t = s ∩ Set.range f` as a `Finset`
3738
let t := (Set.finite_coe_iff.mp (Finite.Set.finite_inter_of_left ↑s (Set.range f))).toFinset
3839
rw [show Nat.card (f⁻¹' s) = Nat.card (f⁻¹' t) by simp [t]]

Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Proper.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -191,13 +191,14 @@ theorem valuativeCriterion_existence_aux
191191
{O : Type*} [CommRing O] [IsDomain O] [ValuationRing O]
192192
{K : Type*} [Field K] [Algebra O K] [IsFractionRing O K]
193193
(φ₀ : (𝒜 0) →+* O)
194-
(ι : Type*) [Fintype ι] (x : ι → A) (h2 : Algebra.adjoin (𝒜 0) (Set.range x) = ⊤)
194+
(ι : Type*) [Finite ι] (x : ι → A) (h2 : Algebra.adjoin (𝒜 0) (Set.range x) = ⊤)
195195
(j : ι) (φ : Away 𝒜 (x j) →+* K)
196196
(hcomm : (algebraMap O K).comp φ₀ = φ.comp (fromZeroRingHom 𝒜 _))
197197
(d : ι → ℕ) (hdi : ∀ i, 0 < d i) (hxdi : ∀ i, x i ∈ 𝒜 (d i)) :
198198
∃ (j₀ : ι) (φ' : Away 𝒜 (x j * x j₀) →+* K), φ'.comp (awayMap 𝒜 (hxdi j₀) rfl) = φ ∧
199199
(φ'.comp (awayMap 𝒜 (hxdi j) (mul_comm (x j) (x j₀)))).range ≤ (algebraMap O K).range := by
200200
classical
201+
cases nonempty_fintype ι
201202
let ψ (i : ι) : ValueGroup O K :=
202203
valuation O K ((φ (Away.isLocalizationElem (hxdi j) (hxdi i))) ^ ∏ k ∈ Finset.univ.erase i, d k)
203204
have : Nonempty ι := ⟨j⟩

Mathlib/Combinatorics/SimpleGraph/Tutte.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -52,7 +52,7 @@ private lemma tutte_exists_isAlternating_isCycles {x b a c : V} {M : Subgraph (G
5252
· simp +contextual [hnpxb, edge_adj, hnab.symm]
5353
· simpa [edge_adj, adj_congr_of_sym2 _ ((adj_edge _ _).mp hvw).1.symm] using .inl hgadj
5454

55-
variable [Fintype V]
55+
variable [Finite V]
5656

5757
lemma IsTutteViolator.mono {u : Set V} (h : G ≤ G') (ht : G'.IsTutteViolator u) :
5858
G.IsTutteViolator u := by
@@ -120,6 +120,7 @@ theorem Subgraph.IsPerfectMatching.exists_of_isClique_supp
120120
G.deleteUniversalVerts.coe.IsClique K.supp) :
121121
∃ (M : Subgraph G), M.IsPerfectMatching := by
122122
classical
123+
cases nonempty_fintype V
123124
obtain ⟨M, hM, hsub⟩ := IsMatching.exists_verts_compl_subset_universalVerts h h'
124125
obtain ⟨M', hM'⟩ := ((G.isClique_universalVerts.subset hsub).even_iff_exists_isMatching
125126
(Set.toFinite _)).mp (by simpa [Set.even_ncard_compl_iff hveven, -Set.toFinset_card,
@@ -269,6 +270,7 @@ lemma exists_isTutteViolator (h : ∀ (M : G.Subgraph), ¬M.IsPerfectMatching)
269270
(hvEven : Even (Nat.card V)) :
270271
∃ u, G.IsTutteViolator u := by
271272
classical
273+
cases nonempty_fintype V
272274
-- It suffices to consider the edge-maximal case
273275
obtain ⟨Gmax, hSubgraph, hMatchingFree, hMaximal⟩ := exists_maximal_isMatchingFree h
274276
refine ⟨Gmax.universalVerts, .mono hSubgraph ?_⟩

Mathlib/LinearAlgebra/Dual/Lemmas.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -105,8 +105,7 @@ variable {R : Type uR} {M : Type uM} {K : Type uK} {V : Type uV} {ι : Type uι}
105105

106106
section CommSemiring
107107

108-
variable [CommSemiring R] [AddCommMonoid M] [Module R M] [DecidableEq ι]
109-
variable (b : Basis ι R M)
108+
variable [CommSemiring R] [AddCommMonoid M] [Module R M]
110109

111110
section Finite
112111

0 commit comments

Comments
 (0)