Skip to content

Commit

Permalink
chore: remove variables (#18193)
Browse files Browse the repository at this point in the history
  • Loading branch information
adomani committed Oct 25, 2024
1 parent f8829a6 commit 9d50218
Show file tree
Hide file tree
Showing 21 changed files with 24 additions and 35 deletions.
5 changes: 1 addition & 4 deletions Mathlib/Algebra/Module/ZLattice/Covolume.lean
Original file line number Diff line number Diff line change
Expand Up @@ -33,10 +33,7 @@ open Submodule MeasureTheory Module MeasureTheory Module ZSpan

section General

variable (K : Type*) [NormedLinearOrderedField K] [HasSolidNorm K] [FloorRing K]
variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E]
variable [ProperSpace E] [MeasurableSpace E]
variable (L : Submodule ℤ E) [DiscreteTopology L] [IsZLattice K L]
variable {E : Type*} [NormedAddCommGroup E] [ProperSpace E] [MeasurableSpace E] (L : Submodule ℤ E)

/-- The covolume of a `ℤ`-lattice is the volume of some fundamental domain; see
`ZLattice.covolume_eq_volume` for the proof that the volume does not depend on the choice of
Expand Down
7 changes: 1 addition & 6 deletions Mathlib/Analysis/Calculus/Deriv/Prod.lean
Original file line number Diff line number Diff line change
Expand Up @@ -29,12 +29,7 @@ open Filter Asymptotics Set

variable {𝕜 : Type u} [NontriviallyNormedField 𝕜]
variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F]
variable {E : Type w} [NormedAddCommGroup E] [NormedSpace 𝕜 E]
variable {f f₀ f₁ g : 𝕜 → F}
variable {f' f₀' f₁' g' : F}
variable {x : 𝕜}
variable {s t : Set 𝕜}
variable {L L₁ L₂ : Filter 𝕜}
variable {f₁ : 𝕜 → F} {f₁' : F} {x : 𝕜} {s : Set 𝕜} {L : Filter 𝕜}

section CartesianProduct

Expand Down
1 change: 0 additions & 1 deletion Mathlib/Analysis/Calculus/FDeriv/Pi.lean
Original file line number Diff line number Diff line change
Expand Up @@ -11,7 +11,6 @@ import Mathlib.Analysis.Calculus.FDeriv.Add

variable {𝕜 ι : Type*} [DecidableEq ι] [Fintype ι] [NontriviallyNormedField 𝕜]
variable {E : ι → Type*} [∀ i, NormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)]
variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F]

@[fun_prop]
theorem hasFDerivAt_update (x : ∀ i, E i) {i : ι} (y : E i) :
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Analysis/Calculus/FDeriv/Star.lean
Original file line number Diff line number Diff line change
Expand Up @@ -27,7 +27,7 @@ variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E]
variable {F : Type*} [NormedAddCommGroup F] [StarAddMonoid F] [NormedSpace 𝕜 F] [StarModule 𝕜 F]
[ContinuousStar F]

variable {f : E → F} {f' : E →L[𝕜] F} (e : E →L[𝕜] F) {x : E} {s : Set E} {L : Filter E}
variable {f : E → F} {f' : E →L[𝕜] F} {x : E} {s : Set E} {L : Filter E}

@[fun_prop]
theorem HasStrictFDerivAt.star (h : HasStrictFDerivAt f f' x) :
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Analysis/Calculus/Gradient/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -252,7 +252,7 @@ section congr

/-! ### Congruence properties of the Gradient -/

variable {f₀ f₁ : F → 𝕜} {f₀' f₁' : F} {x₀ x₁ : F} {s₀ s₁ t : Set F} {L₀ L₁ : Filter F}
variable {f₀ f₁ : F → 𝕜} {f₀' f₁' : F} {t : Set F}

theorem Filter.EventuallyEq.hasGradientAtFilter_iff (h₀ : f₀ =ᶠ[L] f₁) (hx : f₀ x = f₁ x)
(h₁ : f₀' = f₁') : HasGradientAtFilter f₀ f₀' x L ↔ HasGradientAtFilter f₁ f₁' x L :=
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -52,8 +52,6 @@ noncomputable section
variable {𝕜 : Type*} [NontriviallyNormedField 𝕜]
variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E]
variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F]
variable {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G]
variable {G' : Type*} [NormedAddCommGroup G'] [NormedSpace 𝕜 G']
variable {ε : ℝ}

open Filter Metric Set
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Analysis/Calculus/LineDeriv/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -475,7 +475,7 @@ section CompRight

variable {E : Type*} [AddCommGroup E] [Module 𝕜 E]
{E' : Type*} [AddCommGroup E'] [Module 𝕜 E']
{f : E → F} {f' : F} {x v : E'} {L : E' →ₗ[𝕜] E}
{f : E → F} {f' : F} {x : E'} {L : E' →ₗ[𝕜] E}

theorem HasLineDerivAt.of_comp {v : E'} (hf : HasLineDerivAt 𝕜 (f ∘ L) f' x v) :
HasLineDerivAt 𝕜 f f' (L x) (L v) := by
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/MeasureTheory/Integral/SetIntegral.lean
Original file line number Diff line number Diff line change
Expand Up @@ -66,7 +66,7 @@ variable [MeasurableSpace X]
section NormedAddCommGroup

variable [NormedAddCommGroup E] [NormedSpace ℝ E]
{f g : X → E} {s t : Set X} {μ ν : Measure X} {l l' : Filter X}
{f g : X → E} {s t : Set X} {μ : Measure X}

theorem setIntegral_congr_ae₀ (hs : NullMeasurableSet s μ) (h : ∀ᵐ x ∂μ, x ∈ s → f x = g x) :
∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ :=
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Order/PrimeIdeal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -101,7 +101,7 @@ end Preorder

section SemilatticeInf

variable [SemilatticeInf P] {x y : P} {I : Ideal P}
variable [SemilatticeInf P] {I : Ideal P}

theorem IsPrime.mem_or_mem (hI : IsPrime I) {x y : P} : x ⊓ y ∈ I → x ∈ I ∨ y ∈ I := by
contrapose!
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Probability/Distributions/Uniform.lean
Original file line number Diff line number Diff line change
Expand Up @@ -48,7 +48,7 @@ noncomputable section

namespace MeasureTheory

variable {E : Type*} [MeasurableSpace E] {m : Measure E} {μ : Measure E}
variable {E : Type*} [MeasurableSpace E] {μ : Measure E}

namespace pdf

Expand Down Expand Up @@ -206,7 +206,7 @@ end MeasureTheory

namespace PMF

variableβ γ : Type*}
variable {α : Type*}

open scoped Classical NNReal ENNReal

Expand Down
6 changes: 3 additions & 3 deletions Mathlib/Probability/Kernel/Composition.lean
Original file line number Diff line number Diff line change
Expand Up @@ -65,7 +65,7 @@ namespace ProbabilityTheory

namespace Kernel

variable {α β ι : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β}
variable {α β : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β}

section CompositionProduct

Expand Down Expand Up @@ -335,7 +335,7 @@ end Ae

section Restrict

variable {κ : Kernel α β} [IsSFiniteKernel κ] {η : Kernel (α × β) γ} [IsSFiniteKernel η] {a : α}
variable {κ : Kernel α β} [IsSFiniteKernel κ] {η : Kernel (α × β) γ} [IsSFiniteKernel η]

theorem compProd_restrict {s : Set β} {t : Set γ} (hs : MeasurableSet s) (ht : MeasurableSet t) :
Kernel.restrict κ hs ⊗ₖ Kernel.restrict η ht = Kernel.restrict (κ ⊗ₖ η) (hs.prod ht) := by
Expand Down Expand Up @@ -759,7 +759,7 @@ def prodMkLeft (γ : Type*) [MeasurableSpace γ] (κ : Kernel α β) : Kernel (
def prodMkRight (γ : Type*) [MeasurableSpace γ] (κ : Kernel α β) : Kernel (α × γ) β :=
comap κ Prod.fst measurable_fst

variable {γ : Type*} {mγ : MeasurableSpace γ} {f : β → γ} {g : γ → α}
variable {γ : Type*} {mγ : MeasurableSpace γ}

@[simp]
theorem prodMkLeft_apply (κ : Kernel α β) (ca : γ × α) : prodMkLeft γ κ ca = κ ca.snd :=
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Probability/Kernel/Disintegration/CondCDF.lean
Original file line number Diff line number Diff line change
Expand Up @@ -42,7 +42,7 @@ open scoped NNReal ENNReal MeasureTheory Topology

namespace MeasureTheory.Measure

variableβ : Type*} {mα : MeasurableSpace α} (ρ : Measure (α × ℝ))
variable {α : Type*} {mα : MeasurableSpace α} (ρ : Measure (α × ℝ))

/-- Measure on `α` such that for a measurable set `s`, `ρ.IicSnd r s = ρ (s ×ˢ Iic r)`. -/
noncomputable def IicSnd (r : ℝ) : Measure α :=
Expand Down Expand Up @@ -112,7 +112,7 @@ open MeasureTheory

namespace ProbabilityTheory

variableβ ι : Type*} {mα : MeasurableSpace α}
variable {α : Type*} {mα : MeasurableSpace α}

attribute [local instance] MeasureTheory.Measure.IsFiniteMeasure.IicSnd

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Probability/Kernel/Integral.lean
Original file line number Diff line number Diff line change
Expand Up @@ -91,7 +91,7 @@ end Const

section Restrict

variable {s t : Set β}
variable {s : Set β}

@[simp]
theorem integral_restrict (hs : MeasurableSet s) :
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Probability/Kernel/Invariance.lean
Original file line number Diff line number Diff line change
Expand Up @@ -30,7 +30,7 @@ open scoped MeasureTheory ENNReal ProbabilityTheory

namespace ProbabilityTheory

variable {α β γ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ}
variable {α β : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β}

namespace Kernel

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Probability/Kernel/MeasurableIntegral.lean
Original file line number Diff line number Diff line change
Expand Up @@ -231,7 +231,7 @@ alias _root_.Measurable.set_lintegral_kernel := _root_.Measurable.setLIntegral_k

end Lintegral

variable {E : Type*} [NormedAddCommGroup E] [IsSFiniteKernel κ] [IsSFiniteKernel η]
variable {E : Type*} [NormedAddCommGroup E] [IsSFiniteKernel κ]

theorem measurableSet_kernel_integrable ⦃f : α → β → E⦄ (hf : StronglyMeasurable (uncurry f)) :
MeasurableSet {x | Integrable (f x) (κ x)} := by
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Probability/Martingale/OptionalSampling.lean
Original file line number Diff line number Diff line change
Expand Up @@ -143,11 +143,11 @@ and is a measurable space with the Borel σ-algebra. -/

variable {ι : Type*} [LinearOrder ι] [LocallyFiniteOrder ι] [OrderBot ι] [TopologicalSpace ι]
[DiscreteTopology ι] [MeasurableSpace ι] [BorelSpace ι] [MeasurableSpace E] [BorelSpace E]
[SecondCountableTopology E] {ℱ : Filtration ι m} {τ σ : Ω → ι} {f : ι → Ω → E} {i n : ι}
[SecondCountableTopology E] {ℱ : Filtration ι m} {τ σ : Ω → ι} {f : ι → Ω → E} {i : ι}

theorem condexp_stoppedValue_stopping_time_ae_eq_restrict_le (h : Martingale f ℱ μ)
(hτ : IsStoppingTime ℱ τ) (hσ : IsStoppingTime ℱ σ) [SigmaFinite (μ.trim hσ.measurableSpace_le)]
(hτ_le : ∀ x, τ x ≤ n) :
(hτ_le : ∀ x, τ x ≤ i) :
μ[stoppedValue f τ|hσ.measurableSpace] =ᵐ[μ.restrict {x : Ω | τ x ≤ σ x}] stoppedValue f τ := by
rw [ae_eq_restrict_iff_indicator_ae_eq
(hτ.measurableSpace_le _ (hτ.measurableSet_le_stopping_time hσ))]
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/RingTheory/Flat/Stability.lean
Original file line number Diff line number Diff line change
Expand Up @@ -153,7 +153,7 @@ section Localization

variable {R : Type u} {M Mp : Type*} (Rp : Type v)
[CommRing R] [AddCommGroup M] [Module R M] [CommRing Rp] [Algebra R Rp]
[AddCommGroup Mp] [Module R Mp] [Module Rp Mp] [IsScalarTower R Rp Mp] (f : M →ₗ[R] Mp)
[AddCommGroup Mp] [Module R Mp] [Module Rp Mp] [IsScalarTower R Rp Mp]

instance localizedModule [Module.Flat R M] (S : Submonoid R) : Module.Flat (Localization S)
(LocalizedModule S M) := by
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/RingTheory/HahnSeries/Summable.lean
Original file line number Diff line number Diff line change
Expand Up @@ -42,7 +42,7 @@ open Pointwise

noncomputable section

variableΓ' R V α β : Type*}
variableR α β : Type*}

namespace HahnSeries

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/RingTheory/LocalRing/Module.lean
Original file line number Diff line number Diff line change
Expand Up @@ -28,7 +28,7 @@ This file gathers various results about finite modules over a local ring `(R,
`l` is a split injection if and only if `k ⊗ l` is a (split) injection.
-/

variable {R S} [CommRing R] [CommRing S] [Algebra R S]
variable {R} [CommRing R]

section

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/RingTheory/Regular/IsSMulRegular.lean
Original file line number Diff line number Diff line change
Expand Up @@ -136,7 +136,7 @@ open Submodule Pointwise

variable (M) [CommRing R] [AddCommGroup M] [Module R M]
[AddCommGroup M'] [Module R M'] [AddCommGroup M''] [Module R M'']
(I : Ideal R) (N : Submodule R M) (r : R)
(N : Submodule R M) (r : R)

variable (R) in
lemma biUnion_associatedPrimes_eq_compl_regular [IsNoetherianRing R] :
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Topology/Order/UpperLowerSetTopology.lean
Original file line number Diff line number Diff line change
Expand Up @@ -95,7 +95,7 @@ protected def rec {β : WithUpperSet α → Sort*} (h : ∀ a, β (toUpperSet a)
instance [Nonempty α] : Nonempty (WithUpperSet α) := ‹Nonempty α›
instance [Inhabited α] : Inhabited (WithUpperSet α) := ‹Inhabited α›

variable [Preorder α] [Preorder β] [Preorder γ]
variable [Preorder α] [Preorder β]

instance : Preorder (WithUpperSet α) := ‹Preorder α›
instance : TopologicalSpace (WithUpperSet α) := upperSet α
Expand Down

0 comments on commit 9d50218

Please sign in to comment.