Skip to content

Commit

Permalink
chore: remove more unused variables (#18425)
Browse files Browse the repository at this point in the history
  • Loading branch information
adomani committed Oct 30, 2024
1 parent 3f7722b commit 250bf30
Show file tree
Hide file tree
Showing 22 changed files with 30 additions and 43 deletions.
2 changes: 1 addition & 1 deletion Mathlib/Combinatorics/Digraph/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -96,7 +96,7 @@ Any bipartite digraph may be regarded as a subgraph of one of these.
def completeBipartiteGraph (V W : Type*) : Digraph (Sum V W) where
Adj v w := v.isLeft ∧ w.isRight ∨ v.isRight ∧ w.isLeft

variable {ι : Sort*} {V W X : Type*} (G : Digraph V) (G' : Digraph W) {a b c u v w : V}
variable {ι : Sort*} {V : Type*} (G : Digraph V) {a b : V}

theorem adj_injective : Injective (Adj : Digraph V → V → V → Prop) := fun _ _ ↦ Digraph.ext

Expand Down
7 changes: 3 additions & 4 deletions Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean
Original file line number Diff line number Diff line change
Expand Up @@ -100,8 +100,7 @@ variable {α β : Type*}
/-! ### Truncated supremum, truncated infimum -/

section SemilatticeSup
variable [SemilatticeSup α] [SemilatticeSup β]
[BoundedOrder β] {s t : Finset α} {a b : α}
variable [SemilatticeSup α] [SemilatticeSup β] [BoundedOrder β] {s t : Finset α} {a : α}

private lemma sup_aux [@DecidableRel α (· ≤ ·)] : a ∈ lowerClosure s → {b ∈ s | a ≤ b}.Nonempty :=
fun ⟨b, hb, hab⟩ ↦ ⟨b, mem_filter.2 ⟨hb, hab⟩⟩
Expand Down Expand Up @@ -283,7 +282,7 @@ lemma truncatedInf_sups_of_not_mem (ha : a ∉ upperClosure s ⊔ upperClosure t
end DistribLattice

section BooleanAlgebra
variable [BooleanAlgebra α] [@DecidableRel α (· ≤ ·)] {s : Finset α} {a : α}
variable [BooleanAlgebra α] [@DecidableRel α (· ≤ ·)]

@[simp] lemma compl_truncatedSup (s : Finset α) (a : α) :
(truncatedSup s a)ᶜ = truncatedInf sᶜˢ aᶜ := map_truncatedSup (OrderIso.compl α) _ _
Expand Down Expand Up @@ -329,7 +328,7 @@ open Finset hiding card
open Fintype Nat

namespace AhlswedeZhang
variable {α : Type*} [Fintype α] [DecidableEq α] {𝒜 : Finset (Finset α)} {s : Finset α}
variable {α : Type*} [Fintype α] [DecidableEq α] {𝒜 : Finset (Finset α)} {s : Finset α}

/-- Weighted sum of the size of the truncated infima of a set family. Relevant to the
Ahlswede-Zhang identity. -/
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Combinatorics/SetFamily/Compression/UV.lean
Original file line number Diff line number Diff line change
Expand Up @@ -70,7 +70,7 @@ namespace UV
section GeneralizedBooleanAlgebra

variable [GeneralizedBooleanAlgebra α] [DecidableRel (@Disjoint α _ _)]
[DecidableRel ((· ≤ ·) : α → α → Prop)] {s : Finset α} {u v a b : α}
[DecidableRel ((· ≤ ·) : α → α → Prop)] {s : Finset α} {u v a : α}

/-- UV-compressing `a` means removing `v` from it and adding `u` if `a` and `u` are disjoint and
`v ≤ a` (it replaces the `v` part of `a` by the `u` part). Else, UV-compressing `a` doesn't do
Expand Down Expand Up @@ -265,7 +265,7 @@ end GeneralizedBooleanAlgebra

open FinsetFamily

variable [DecidableEq α] {𝒜 : Finset (Finset α)} {u v a : Finset α} {r : ℕ}
variable [DecidableEq α] {𝒜 : Finset (Finset α)} {u v : Finset α} {r : ℕ}

/-- Compressing a finset doesn't change its size. -/
theorem card_compress (huv : #u = #v) (a : Finset α) : #(compress u v a) = #a := by
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Combinatorics/SetFamily/FourFunctions.lean
Original file line number Diff line number Diff line change
Expand Up @@ -58,8 +58,8 @@ open scoped FinsetFamily
variable {α β : Type*}

section Finset
variable [DecidableEq α] [LinearOrderedCommSemiring β] {𝒜 : Finset (Finset α)}
{a : α} {f f₁ f₂ f₃ f₄ g μ : Finset α → β} {s t u : Finset α}
variable [DecidableEq α] [LinearOrderedCommSemiring β] {𝒜 : Finset (Finset α)}
{a : α} {f f₁ f₂ f₃ f₄ : Finset α → β} {s t u : Finset α}

/-- The `n = 1` case of the Ahlswede-Daykin inequality. Note that we can't just expand everything
out and bound termwise since `c₀ * d₁` appears twice on the RHS of the assumptions while `c₁ * d₀`
Expand Down
10 changes: 5 additions & 5 deletions Mathlib/Data/Finset/Sups.lean
Original file line number Diff line number Diff line change
Expand Up @@ -356,7 +356,7 @@ end DistribLattice

section Finset
variable [DecidableEq α]
variable {𝒜 ℬ : Finset (Finset α)} {s t : Finset α} {a : α}
variable {𝒜 ℬ : Finset (Finset α)} {s t : Finset α}

@[simp] lemma powerset_union (s t : Finset α) : (s ∪ t).powerset = s.powerset ⊻ t.powerset := by
ext u
Expand Down Expand Up @@ -508,7 +508,7 @@ theorem disjSups_disjSups_disjSups_comm : s ○ t ○ (u ○ v) = s ○ u ○ (t
end DistribLattice
section Diffs
variable [DecidableEq α]
variable [GeneralizedBooleanAlgebra α] (s s₁ s₂ t t₁ t₂ u v : Finset α)
variable [GeneralizedBooleanAlgebra α] (s s₁ s₂ t t₁ t₂ u : Finset α)

/-- `s \\ t` is the finset of elements of the form `a \ b` where `a ∈ s`, `b ∈ t`. -/
def diffs : Finset α → Finset α → Finset α := image₂ (· \ ·)
Expand Down Expand Up @@ -592,7 +592,7 @@ lemma diffs_right_comm : s \\ t \\ u = s \\ u \\ t := image₂_right_comm sdiff_
end Diffs

section Compls
variable [BooleanAlgebra α] (s s₁ s₂ t t₁ t₂ u v : Finset α)
variable [BooleanAlgebra α] (s s₁ s₂ t : Finset α)

/-- `sᶜˢ` is the finset of elements of the form `aᶜ` where `a ∈ s`. -/
def compls : Finset α → Finset α := map ⟨compl, compl_injective⟩
Expand All @@ -602,7 +602,7 @@ scoped[FinsetFamily] postfix:max "ᶜˢ" => Finset.compls

open FinsetFamily

variable {s t} {a b c : α}
variable {s t} {a : α}

@[simp] lemma mem_compls : a ∈ sᶜˢ ↔ aᶜ ∈ s := by
rw [Iff.comm, ← mem_map' ⟨compl, compl_injective⟩, Embedding.coeFn_mk, compl_compl, compls]
Expand All @@ -615,7 +615,7 @@ variable (s t)

@[simp] lemma card_compls : #sᶜˢ = #s := card_map _

variable {s s₁ s₂ t t₁ t₂ u}
variable {s s₁ s₂ t}

lemma compl_mem_compls : a ∈ s → aᶜ ∈ sᶜˢ := mem_map_of_mem _
@[simp] lemma compls_subset_compls : s₁ᶜˢ ⊆ s₂ᶜˢ ↔ s₁ ⊆ s₂ := map_subset_map
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/GroupTheory/GroupAction/Blocks.lean
Original file line number Diff line number Diff line change
Expand Up @@ -348,7 +348,7 @@ theorem IsBlockSystem.of_normal {N : Subgroup G} [N.Normal] :
exact .orbit_of_normal a

section Group
variable {S H : Type*} [Group H] [SetLike S H] [SubgroupClass S H] {s : S} {a b : G}
variable {S H : Type*} [Group H] [SetLike S H] [SubgroupClass S H] {s : S} {a : G}

/-!
Annoyingly, it seems like the following two lemmas cannot be unified.
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/GroupTheory/GroupAction/Support.lean
Original file line number Diff line number Diff line change
Expand Up @@ -44,7 +44,7 @@ theorem Supports.mono (h : s ⊆ t) (hs : Supports G s b) : Supports G t b := fu
end SMul

variable [Group H] [SMul G α] [SMul G β] [MulAction H α] [SMul H β] [SMulCommClass G H β]
[SMulCommClass G H α] {s t : Set α} {b : β}
[SMulCommClass G H α] {s : Set α} {b : β}

-- TODO: This should work without `SMulCommClass`
@[to_additive]
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/GroupTheory/Order/Min.lean
Original file line number Diff line number Diff line change
Expand Up @@ -49,7 +49,7 @@ lemma minOrder_le_orderOf (ha : a ≠ 1) (ha' : IsOfFinOrder a) : minOrder α

end Monoid

variable [Group α] {s : Subgroup α} {n : ℕ}
variable [Group α] {s : Subgroup α}

@[to_additive]
lemma le_minOrder_iff_forall_subgroup {n : ℕ∞} :
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/GroupTheory/Perm/Cycle/Concrete.lean
Original file line number Diff line number Diff line change
Expand Up @@ -118,7 +118,7 @@ end List

namespace Cycle

variable [DecidableEq α] (s s' : Cycle α)
variable [DecidableEq α] (s : Cycle α)

/-- A cycle `s : Cycle α`, given `Nodup s` can be interpreted as an `Equiv.Perm α`
where each element in the list is permuted to the next one, defined as `formPerm`.
Expand Down
5 changes: 2 additions & 3 deletions Mathlib/LinearAlgebra/BilinearForm/TensorProduct.lean
Original file line number Diff line number Diff line change
Expand Up @@ -21,10 +21,9 @@ import Mathlib.LinearAlgebra.BilinearForm.Hom

suppress_compilation

universe u v w uR uA uM₁ uM₂ uN₁ uN₂
universe u v w uR uA uM₁ uM₂ uN₁ uN₂

variable {ι : Type uι} {R : Type uR} {A : Type uA} {M₁ : Type uM₁} {M₂ : Type uM₂} {N₁ : Type uN₁}
{N₂ : Type uN₂}
variable {R : Type uR} {A : Type uA} {M₁ : Type uM₁} {M₂ : Type uM₂} {N₁ : Type uN₁} {N₂ : Type uN₂}

open TensorProduct

Expand Down
1 change: 0 additions & 1 deletion Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -46,7 +46,6 @@ This file is almost identical to `Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.le
variable {R : Type*} [CommRing R]
variable {M : Type*} [AddCommGroup M] [Module R M]
variable (Q : QuadraticForm R M)
variable {n : ℕ}

namespace CliffordAlgebra

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/LinearAlgebra/CliffordAlgebra/Equivs.lean
Original file line number Diff line number Diff line change
Expand Up @@ -354,7 +354,7 @@ open scoped DualNumber

open DualNumber TrivSqZeroExt

variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M]
variable {R : Type*} [CommRing R]

theorem ι_mul_ι (r₁ r₂) : ι (0 : QuadraticForm R R) r₁ * ι (0 : QuadraticForm R R) r₂ = 0 := by
rw [← mul_one r₁, ← mul_one r₂, ← smul_eq_mul R, ← smul_eq_mul R, LinearMap.map_smul,
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/LinearAlgebra/QuadraticForm/Isometry.lean
Original file line number Diff line number Diff line change
Expand Up @@ -17,7 +17,7 @@ import Mathlib.LinearAlgebra.QuadraticForm.Basic
`Q₁ →qᵢ Q₂` is notation for `Q₁.Isometry Q₂`.
-/

variable {ι R M M₁ M₂ M₃ M₄ N : Type*}
variable {R M M₁ M₂ M₃ M₄ N : Type*}

namespace QuadraticMap

Expand Down
4 changes: 1 addition & 3 deletions Mathlib/LinearAlgebra/QuadraticForm/Prod.lean
Original file line number Diff line number Diff line change
Expand Up @@ -346,9 +346,7 @@ end Semiring
namespace Ring

variable [CommRing R]
variable [∀ i, AddCommGroup (Mᵢ i)] [∀ i, AddCommGroup (Nᵢ i)] [AddCommGroup P]
variable [∀ i, Module R (Mᵢ i)] [∀ i, Module R (Nᵢ i)] [Module R P]
variable [Fintype ι]
variable [∀ i, AddCommGroup (Mᵢ i)] [AddCommGroup P] [∀ i, Module R (Mᵢ i)] [Module R P] [Fintype ι]

@[simp] theorem polar_pi (Q : ∀ i, QuadraticMap R (Mᵢ i) P) (x y : ∀ i, Mᵢ i) :
polar (pi Q) x y = ∑ i, polar (Q i) (x i) (y i) := by
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean
Original file line number Diff line number Diff line change
Expand Up @@ -47,7 +47,7 @@ suppress_compilation

open scoped TensorProduct DirectSum

variable {R ι A B : Type*}
variable {R ι : Type*}

namespace TensorProduct

Expand Down
4 changes: 1 addition & 3 deletions Mathlib/Order/Filter/Cocardinal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -22,9 +22,7 @@ In this file we define `Filter.cocardinal hc`: the filter of sets with cardinali
open Set Filter Cardinal

universe u
variable {ι : Type u} {α β : Type u}
variable {c : Cardinal.{u}} {hreg : c.IsRegular}
variable {l : Filter α}
variable {α : Type u} {c : Cardinal.{u}} {hreg : c.IsRegular}

namespace Filter

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Order/KonigLemma.lean
Original file line number Diff line number Diff line change
Expand Up @@ -51,7 +51,7 @@ Formulate the lemma as a statement about graphs.
open Set
section Sequence

variable {α : Type*} [PartialOrder α] [IsStronglyAtomic α] {a b : α}
variable {α : Type*} [PartialOrder α] [IsStronglyAtomic α] {b : α}

/-- **Kőnig's infinity lemma** : if each element in a strongly atomic order
is covered by only finitely many others, and `b` is an element with infinitely many things above it,
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Order/PFilter.lean
Original file line number Diff line number Diff line change
Expand Up @@ -143,7 +143,7 @@ end SemilatticeInf

section CompleteSemilatticeInf

variable [CompleteSemilatticeInf P] {F : PFilter P}
variable [CompleteSemilatticeInf P]

theorem sInf_gc :
GaloisConnection (fun x => toDual (principal x)) fun F => sInf (ofDual F : PFilter P) :=
Expand Down
2 changes: 0 additions & 2 deletions Mathlib/RingTheory/WittVector/Defs.lean
Original file line number Diff line number Diff line change
Expand Up @@ -156,8 +156,6 @@ evaluating this at `(x₀, x₁)` gives us the sum of two Witt vectors `x₀ + x
def eval {k : ℕ} (φ : ℕ → MvPolynomial (Fin k × ℕ) ℤ) (x : Fin k → 𝕎 R) : 𝕎 R :=
mk p fun n => peval (φ n) fun i => (x i).coeff

variable (R) [Fact p.Prime]

instance : Zero (𝕎 R) :=
⟨eval (wittZero p) ![]⟩

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Tactic/Widget/StringDiagram.lean
Original file line number Diff line number Diff line change
Expand Up @@ -182,7 +182,7 @@ def WhiskerLeft.nodes (v h₁ h₂ : ℕ) : WhiskerLeft → List Node
let ss := η.nodes v (h₁ + 1) (h₂ + 1)
s :: ss

variable {ρ : Type} [Context ρ] [MonadMor₁ (CoherenceM ρ)]
variable {ρ : Type} [MonadMor₁ (CoherenceM ρ)]

/-- The list of nodes at the top of a string diagram. -/
def topNodes (η : WhiskerLeft) : CoherenceM ρ (List Node) := do
Expand Down
2 changes: 0 additions & 2 deletions Mathlib/Topology/ContinuousMap/CompactlySupported.lean
Original file line number Diff line number Diff line change
Expand Up @@ -447,8 +447,6 @@ theorem zero_comp (g : β →co γ) : (0 : C_c(γ, δ)).comp g = 0 :=

end

variable [T2Space γ]

/-- Composition as an additive monoid homomorphism. -/
def compAddMonoidHom [AddMonoid δ] [ContinuousAdd δ] (g : β →co γ) : C_c(γ, δ) →+ C_c(β, δ) where
toFun f := f.comp g
Expand Down
8 changes: 3 additions & 5 deletions Mathlib/Topology/FiberBundle/Trivialization.lean
Original file line number Diff line number Diff line change
Expand Up @@ -50,7 +50,7 @@ type of linear trivializations is not even particularly well-behaved.
open TopologicalSpace Filter Set Bundle Function
open scoped Topology

variable {ι : Type*} {B : Type*} (F : Type*) {E : B → Type*}
variable {B : Type*} (F : Type*) {E : B → Type*}
variable {Z : Type*} [TopologicalSpace B] [TopologicalSpace F] {proj : Z → B}

/-- This structure contains the information left for a local trivialization (which is implemented
Expand Down Expand Up @@ -194,7 +194,7 @@ theorem symm_trans_target_eq (e e' : Pretrivialization F proj) :
(e.toPartialEquiv.symm.trans e'.toPartialEquiv).target = (e.baseSet ∩ e'.baseSet) ×ˢ univ := by
rw [← PartialEquiv.symm_source, symm_trans_symm, symm_trans_source_eq, inter_comm]

variable (e' : Pretrivialization F (π F E)) {x' : TotalSpace F E} {b : B} {y : E b}
variable (e' : Pretrivialization F (π F E)) {b : B} {y : E b}

@[simp]
theorem coe_mem_source : ↑y ∈ e'.source ↔ b ∈ e'.baseSet :=
Expand Down Expand Up @@ -502,7 +502,7 @@ theorem continuousAt_of_comp_left {X : Type*} [TopologicalSpace X] {f : X → Z}
rw [e.source_eq, ← preimage_comp]
exact hf_proj.preimage_mem_nhds (e.open_baseSet.mem_nhds he)

variable (e' : Trivialization F (π F E)) {x' : TotalSpace F E} {b : B} {y : E b}
variable (e' : Trivialization F (π F E)) {b : B} {y : E b}

protected theorem continuousOn : ContinuousOn e' e'.source :=
e'.continuousOn_toFun
Expand Down Expand Up @@ -643,8 +643,6 @@ theorem coordChangeHomeomorph_coe (e₁ e₂ : Trivialization F proj) {b : B} (h
(h₂ : b ∈ e₂.baseSet) : ⇑(e₁.coordChangeHomeomorph e₂ h₁ h₂) = e₁.coordChange e₂ b :=
rfl

variable {B' : Type*} [TopologicalSpace B']

theorem isImage_preimage_prod (e : Trivialization F proj) (s : Set B) :
e.toPartialHomeomorph.IsImage (proj ⁻¹' s) (s ×ˢ univ) := fun x hx => by simp [e.coe_fst', hx]

Expand Down

0 comments on commit 250bf30

Please sign in to comment.