Skip to content

Commit

Permalink
chore: avoid some unused variables (#11583)
Browse files Browse the repository at this point in the history
These will be caught by the linter in a future lean version.
  • Loading branch information
Ruben-VandeVelde authored and utensil committed Mar 26, 2024
1 parent 6595001 commit 7b56865
Show file tree
Hide file tree
Showing 12 changed files with 25 additions and 34 deletions.
10 changes: 5 additions & 5 deletions Mathlib/Algebra/Polynomial/BigOperators.lean
Original file line number Diff line number Diff line change
Expand Up @@ -270,16 +270,16 @@ theorem multiset_prod_X_sub_C_coeff_card_pred (t : Multiset R) (ht : 0 < Multise
(t.map fun x => X - C x).prod.coeff ((Multiset.card t) - 1) = -t.sum := by
nontriviality R
convert multiset_prod_X_sub_C_nextCoeff (by assumption)
rw [nextCoeff]; split_ifs with h
· rw [natDegree_multiset_prod_of_monic] at h
rw [nextCoeff, if_neg]
swap
· rw [natDegree_multiset_prod_of_monic]
swap
· simp only [Multiset.mem_map]
rintro _ ⟨_, _, rfl⟩
apply monic_X_sub_C
simp_rw [Multiset.sum_eq_zero_iff, Multiset.mem_map] at h
contrapose! h
simp_rw [Multiset.sum_eq_zero_iff, Multiset.mem_map]
obtain ⟨x, hx⟩ := card_pos_iff_exists_mem.mp ht
exact ⟨_, ⟨_, ⟨x, hx, rfl⟩, natDegree_X_sub_C _⟩, one_ne_zero
exact fun h => one_ne_zero <| h 1 ⟨_, ⟨x, hx, rfl⟩, natDegree_X_sub_C _⟩
congr; rw [natDegree_multiset_prod_of_monic] <;> · simp [natDegree_X_sub_C, monic_X_sub_C]
set_option linter.uppercaseLean3 false in
#align polynomial.multiset_prod_X_sub_C_coeff_card_pred Polynomial.multiset_prod_X_sub_C_coeff_card_pred
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Combinatorics/SimpleGraph/Connectivity.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1757,7 +1757,7 @@ protected def transfer {u v : V} (p : G.Walk u v)
(H : SimpleGraph V) (h : ∀ e, e ∈ p.edges → e ∈ H.edgeSet) : H.Walk u v :=
match p with
| nil => nil
| cons' u v w a p =>
| cons' u v w _ p =>
cons (h s(u, v) (by simp)) (p.transfer H fun e he => h e (by simp [he]))
#align simple_graph.walk.transfer SimpleGraph.Walk.transfer

Expand Down
3 changes: 1 addition & 2 deletions Mathlib/Data/List/BigOperators/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -202,8 +202,7 @@ theorem prod_take_succ :
@[to_additive "A list with sum not zero must have positive length."]
theorem length_pos_of_prod_ne_one (L : List M) (h : L.prod ≠ 1) : 0 < L.length := by
cases L
· contrapose h
simp
· simp at h
· simp
#align list.length_pos_of_prod_ne_one List.length_pos_of_prod_ne_one
#align list.length_pos_of_sum_ne_zero List.length_pos_of_sum_ne_zero
Expand Down
12 changes: 5 additions & 7 deletions Mathlib/Data/PNat/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -398,13 +398,11 @@ theorem dvd_iff {k m : ℕ+} : k ∣ m ↔ (k : ℕ) ∣ (m : ℕ) := by
· rcases h with ⟨_, rfl⟩
apply dvd_mul_right
· rcases h with ⟨a, h⟩
cases a with
| zero =>
contrapose h
apply ne_zero
| succ n =>
use ⟨n.succ, n.succ_pos⟩
rw [← coe_inj, h, mul_coe, mk_coe]
obtain ⟨n, rfl⟩ := Nat.exists_eq_succ_of_ne_zero (n := a) <| by
rintro rfl
simp only [mul_zero, ne_zero] at h
use ⟨n.succ, n.succ_pos⟩
rw [← coe_inj, h, mul_coe, mk_coe]
#align pnat.dvd_iff PNat.dvd_iff

theorem dvd_iff' {k m : ℕ+} : k ∣ m ↔ mod m k = k := by
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Mathport/Notation.lean
Original file line number Diff line number Diff line change
Expand Up @@ -530,7 +530,7 @@ elab (name := notation3) doc:(docComment)? attrs?:(Parser.Term.attributes)? attr
boundIdents := boundIdents.insert lit.getId lit
boundValues := boundValues.insert lit.getId <| lit.1.mkAntiquotNode `term
boundNames := boundNames.push lit.getId
| stx => throwUnsupportedSyntax
| _stx => throwUnsupportedSyntax
if hasScoped && !hasBindersItem then
throwError "If there is a `scoped` item then there must be a `(...)` item for binders."

Expand Down
3 changes: 1 addition & 2 deletions Mathlib/NumberTheory/Divisors.lean
Original file line number Diff line number Diff line change
Expand Up @@ -410,8 +410,7 @@ theorem sum_properDivisors_dvd (h : (∑ x in n.properDivisors, x) ∣ n) :
cases' n with n
· simp
· cases' n with n
· contrapose! h
simp
· simp at h
· rw [or_iff_not_imp_right]
intro ne_n
have hlt : ∑ x in n.succ.succ.properDivisors, x < n.succ.succ :=
Expand Down
4 changes: 1 addition & 3 deletions Mathlib/RingTheory/HahnSeries/Summable.lean
Original file line number Diff line number Diff line change
Expand Up @@ -423,9 +423,7 @@ def embDomain (s : SummableFamily Γ R α) (f : α ↪ β) : SummableFamily Γ R
· dsimp only at h
rw [dif_pos hb] at h
exact Set.mem_iUnion.2 ⟨Classical.choose hb, h⟩
· contrapose! h
rw [dif_neg hb]
simp
· simp [-Set.mem_range, dif_neg hb] at h
finite_co_support' g :=
((s.finite_co_support g).image f).subset
(by
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/RingTheory/MvPolynomial/Homogeneous.lean
Original file line number Diff line number Diff line change
Expand Up @@ -108,7 +108,7 @@ theorem homogeneousSubmodule_mul [CommSemiring R] (m n : ℕ) :
rw [← hde, ← hφ, ← hψ, Finset.sum_subset Finsupp.support_add, Finset.sum_subset hd',
Finset.sum_subset he', ← Finset.sum_add_distrib]
· congr
all_goals intro i hi; apply Finsupp.not_mem_support_iff.mp
all_goals intro _ _; apply Finsupp.not_mem_support_iff.mp
#align mv_polynomial.homogeneous_submodule_mul MvPolynomial.homogeneousSubmodule_mul

section
Expand Down
4 changes: 1 addition & 3 deletions Mathlib/RingTheory/PowerSeries/Order.lean
Original file line number Diff line number Diff line change
Expand Up @@ -64,9 +64,7 @@ theorem order_finite_iff_ne_zero : (order φ).Dom ↔ φ ≠ 0 := by
simp only [order]
constructor
· split_ifs with h <;> intro H
· contrapose! H
simp only [← Part.eq_none_iff']
rfl
· simp only [PartENat.top_eq_none, Part.not_none_dom] at H
· exact h
· intro h
simp [h]
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/SetTheory/Ordinal/NaturalOps.lean
Original file line number Diff line number Diff line change
Expand Up @@ -292,7 +292,7 @@ theorem blsub_nadd_of_mono {f : ∀ c < a ♯ b, Ordinal.{max u v}}

theorem nadd_assoc (a b c) : a ♯ b ♯ c = a ♯ (b ♯ c) := by
rw [nadd_def a (b ♯ c), nadd_def, blsub_nadd_of_mono, blsub_nadd_of_mono, max_assoc]
· congr <;> ext (d hd) <;> apply nadd_assoc
· congr <;> ext <;> apply nadd_assoc
· exact fun _ _ h => nadd_le_nadd_left h a
· exact fun _ _ h => nadd_le_nadd_right h c
termination_by (a, b, c)
Expand Down
7 changes: 4 additions & 3 deletions Mathlib/Topology/Algebra/Order/Compact.lean
Original file line number Diff line number Diff line change
Expand Up @@ -109,9 +109,10 @@ instance (priority := 100) ConditionallyCompleteLinearOrder.toCompactIccSpace (
rw [diff_subset_iff]
exact Subset.trans Icc_subset_Icc_union_Ioc <| union_subset_union Subset.rfl <|
Ioc_subset_Ioc_left hy.1.le
rcases hc.2.eq_or_lt with (rfl | hlt); · exact hcs.2
contrapose! hf
intro U hU
rcases hc.2.eq_or_lt with (rfl | hlt)
· exact hcs.2
exfalso
refine hf fun U hU => ?_
rcases (mem_nhdsWithin_Ici_iff_exists_mem_Ioc_Ico_subset hlt).1
(mem_nhdsWithin_of_mem_nhds hU) with
⟨y, hxy, hyU⟩
Expand Down
8 changes: 3 additions & 5 deletions Mathlib/Topology/Connected/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1035,11 +1035,9 @@ theorem isPreconnected_iff_subset_of_disjoint {s : Set α} :
have hyu : y ∈ u := or_iff_not_imp_right.mp (hs hys) hyv
exact h ⟨y, hys, hyu⟩ ⟨x, hxs, hxv⟩
· intro u v hu hv hs hsu hsv
rw [nonempty_iff_ne_empty]
intro H
specialize h u v hu hv hs H
contrapose H
apply Nonempty.ne_empty
by_contra H
specialize h u v hu hv hs (Set.not_nonempty_iff_eq_empty.mp H)
apply H
cases' h with h h
· rcases hsv with ⟨x, hxs, hxv⟩
exact ⟨x, hxs, ⟨h hxs, hxv⟩⟩
Expand Down

0 comments on commit 7b56865

Please sign in to comment.