From 4d4f0f25d08d3790512250371049b68988a0bfcc Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Fri, 30 Jun 2023 01:46:32 +0000 Subject: [PATCH] fix: change compl precedence (#5586) Co-authored-by: Yury G. Kudryashov --- Archive/Wiedijk100Theorems/BallotProblem.lean | 4 +- Mathlib/Algebra/Algebra/Spectrum.lean | 2 +- Mathlib/Algebra/BigOperators/Basic.lean | 2 +- Mathlib/Algebra/IndicatorFunction.lean | 12 ++--- Mathlib/Algebra/Ring/BooleanRing.lean | 2 +- Mathlib/Algebra/Ring/Idempotents.lean | 2 +- Mathlib/Algebra/Star/Pointwise.lean | 2 +- Mathlib/Algebra/Support.lean | 2 +- .../PrimeSpectrum/Basic.lean | 6 +-- .../PrimeSpectrum/IsOpenComapC.lean | 2 +- .../ProjectiveSpectrum/Topology.lean | 4 +- Mathlib/Analysis/Calculus/ContDiff.lean | 2 +- Mathlib/Analysis/Calculus/Deriv/Slope.lean | 8 ++-- Mathlib/Analysis/Calculus/Dslope.lean | 4 +- Mathlib/Analysis/Convex/Integral.lean | 6 +-- Mathlib/Analysis/Convex/StoneSeparation.lean | 2 +- .../Analysis/NormedSpace/FiniteDimension.lean | 4 +- Mathlib/Analysis/NormedSpace/MStructure.lean | 28 +++++------ Mathlib/Analysis/NormedSpace/Multilinear.lean | 2 +- Mathlib/Analysis/NormedSpace/PiLp.lean | 2 +- Mathlib/Analysis/NormedSpace/RieszLemma.lean | 4 +- Mathlib/Analysis/NormedSpace/Units.lean | 2 +- .../Analysis/SpecialFunctions/Log/Basic.lean | 2 +- .../Analysis/SpecialFunctions/Log/Deriv.lean | 6 +-- .../SpecialFunctions/Pow/Continuity.lean | 2 +- .../Trigonometric/InverseDeriv.lean | 12 ++--- .../Limits/Shapes/Biproducts.lean | 12 ++--- .../CategoryTheory/Limits/Shapes/Types.lean | 2 +- Mathlib/Combinatorics/Configuration.lean | 4 +- Mathlib/Combinatorics/SimpleGraph/Basic.lean | 4 +- .../Combinatorics/SimpleGraph/Ends/Defs.lean | 6 +-- .../SimpleGraph/StronglyRegular.lean | 8 ++-- Mathlib/Data/Fin/Basic.lean | 2 +- Mathlib/Data/Finset/Basic.lean | 2 +- Mathlib/Data/Finset/Lattice.lean | 4 +- Mathlib/Data/Finset/Preimage.lean | 2 +- Mathlib/Data/Fintype/Basic.lean | 8 ++-- Mathlib/Data/Fintype/Card.lean | 6 +-- Mathlib/Data/Set/Basic.lean | 4 +- Mathlib/Data/Set/Function.lean | 22 ++++----- Mathlib/Data/Set/Image.lean | 10 ++-- Mathlib/Data/Set/Intervals/Basic.lean | 8 ++-- .../Set/Intervals/OrdConnectedComponent.lean | 2 +- Mathlib/Data/Set/Lattice.lean | 12 ++--- Mathlib/Data/Set/Pointwise/Basic.lean | 2 +- Mathlib/Data/Set/Prod.lean | 8 ++-- Mathlib/Dynamics/Ergodic/Ergodic.lean | 2 +- Mathlib/Dynamics/OmegaLimit.lean | 6 +-- Mathlib/Geometry/Euclidean/MongePoint.lean | 2 +- Mathlib/GroupTheory/FreeProduct.lean | 24 +++++----- Mathlib/LinearAlgebra/Finsupp.lean | 4 +- Mathlib/Logic/Embedding/Set.lean | 2 +- Mathlib/Logic/Equiv/Embedding.lean | 4 +- Mathlib/Logic/Equiv/LocalEquiv.lean | 2 +- .../Constructions/BorelSpace/Basic.lean | 12 ++--- .../MeasureTheory/Constructions/Polish.lean | 4 +- .../Covering/Differentiation.lean | 30 ++++++------ .../MeasureTheory/Covering/LiminfLimsup.lean | 2 +- .../MeasureTheory/Decomposition/Jordan.lean | 8 ++-- .../MeasureTheory/Decomposition/Lebesgue.lean | 12 ++--- .../Decomposition/RadonNikodym.lean | 2 +- .../Function/AEEqOfIntegral.lean | 4 +- .../Function/AEMeasurableSequence.lean | 6 +-- .../ConditionalExpectation/AEMeasurable.lean | 2 +- .../ConditionalExpectation/Indicator.lean | 10 ++-- Mathlib/MeasureTheory/Function/Egorov.lean | 2 +- .../MeasureTheory/Function/SimpleFunc.lean | 4 +- .../Function/SimpleFuncDenseLp.lean | 4 +- .../Function/StronglyMeasurable/Basic.lean | 10 ++-- .../Function/UniformIntegrable.lean | 2 +- Mathlib/MeasureTheory/Integral/Average.lean | 2 +- Mathlib/MeasureTheory/Integral/Lebesgue.lean | 4 +- .../MeasureTheory/Integral/SetIntegral.lean | 2 +- Mathlib/MeasureTheory/MeasurableSpace.lean | 4 +- Mathlib/MeasureTheory/MeasurableSpaceDef.lean | 12 ++--- Mathlib/MeasureTheory/Measure/AEDisjoint.lean | 4 +- .../MeasureTheory/Measure/AEMeasurable.lean | 8 ++-- .../Measure/Lebesgue/EqHaar.lean | 4 +- .../MeasureTheory/Measure/MeasureSpace.lean | 46 +++++++++---------- .../Measure/MeasureSpaceDef.lean | 6 +-- .../Measure/MutuallySingular.lean | 2 +- .../MeasureTheory/Measure/NullMeasurable.lean | 12 ++--- Mathlib/MeasureTheory/Measure/OpenPos.lean | 2 +- .../MeasureTheory/Measure/OuterMeasure.lean | 8 ++-- .../MeasureTheory/Measure/Portmanteau.lean | 20 ++++---- Mathlib/MeasureTheory/Measure/Sub.lean | 2 +- Mathlib/MeasureTheory/PiSystem.lean | 10 ++-- Mathlib/ModelTheory/Definability.lean | 4 +- Mathlib/ModelTheory/Syntax.lean | 3 +- Mathlib/Order/Antichain.lean | 4 +- Mathlib/Order/Atoms.lean | 2 +- Mathlib/Order/Basic.lean | 12 ++--- Mathlib/Order/BooleanAlgebra.lean | 8 ++-- Mathlib/Order/Circular.lean | 4 +- Mathlib/Order/CompleteBooleanAlgebra.lean | 14 +++--- Mathlib/Order/Disjointed.lean | 5 +- Mathlib/Order/Filter/Bases.lean | 12 ++--- Mathlib/Order/Filter/Basic.lean | 15 +++--- Mathlib/Order/Filter/Cofinite.lean | 2 +- Mathlib/Order/Filter/Prod.lean | 2 +- Mathlib/Order/Filter/Ultrafilter.lean | 2 +- Mathlib/Order/Heyting/Basic.lean | 24 +++++----- Mathlib/Order/Heyting/Hom.lean | 2 +- Mathlib/Order/Heyting/Regular.lean | 10 ++-- Mathlib/Order/Hom/Lattice.lean | 2 +- Mathlib/Order/LiminfLimsup.lean | 4 +- Mathlib/Order/PrimeIdeal.lean | 4 +- Mathlib/Order/UpperLower/Basic.lean | 12 ++--- Mathlib/Probability/CondCount.lean | 4 +- .../Probability/ConditionalProbability.lean | 2 +- Mathlib/Probability/Independence/Basic.lean | 2 +- Mathlib/Probability/Kernel/CondCdf.lean | 4 +- .../Probability/Kernel/Disintegration.lean | 12 ++--- Mathlib/RingTheory/Polynomial/Basic.lean | 8 ++-- Mathlib/SetTheory/Cardinal/Basic.lean | 4 +- Mathlib/SetTheory/Cardinal/Cofinality.lean | 2 +- Mathlib/SetTheory/Cardinal/Ordinal.lean | 4 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 4 +- Mathlib/Topology/Algebra/GroupWithZero.lean | 4 +- .../Topology/Algebra/InfiniteSum/Basic.lean | 18 ++++---- .../Topology/Algebra/InfiniteSum/Order.lean | 6 +-- .../Algebra/Module/FiniteDimension.lean | 6 +-- Mathlib/Topology/Algebra/Order/T5.lean | 2 +- Mathlib/Topology/Basic.lean | 42 ++++++++--------- Mathlib/Topology/Bornology/Basic.lean | 10 ++-- Mathlib/Topology/Category/Compactum.lean | 2 +- .../Category/TopCat/Limits/Products.lean | 4 +- .../Topology/Compactification/OnePoint.lean | 16 +++---- Mathlib/Topology/Connected.lean | 4 +- Mathlib/Topology/Constructions.lean | 4 +- .../Topology/ContinuousFunction/Bounded.lean | 8 ++-- .../Topology/ContinuousFunction/Ideals.lean | 4 +- Mathlib/Topology/ContinuousOn.lean | 12 ++--- Mathlib/Topology/DenseEmbedding.lean | 2 +- Mathlib/Topology/GDelta.lean | 6 +-- Mathlib/Topology/Inseparable.lean | 2 +- Mathlib/Topology/Instances/RatLemmas.lean | 2 +- Mathlib/Topology/LocalAtTarget.lean | 2 +- Mathlib/Topology/LocalHomeomorph.lean | 2 +- Mathlib/Topology/LocallyConstant/Basic.lean | 4 +- Mathlib/Topology/LocallyFinite.lean | 2 +- Mathlib/Topology/MetricSpace/Baire.lean | 2 +- .../MetricSpace/HausdorffDimension.lean | 6 +-- .../MetricSpace/HausdorffDistance.lean | 8 ++-- Mathlib/Topology/MetricSpace/Metrizable.lean | 4 +- Mathlib/Topology/MetricSpace/PiNat.lean | 8 ++-- Mathlib/Topology/MetricSpace/Polish.lean | 36 +++++++-------- Mathlib/Topology/Order/LowerTopology.lean | 16 +++---- Mathlib/Topology/Paracompact.lean | 10 ++-- Mathlib/Topology/Separation.lean | 6 +-- Mathlib/Topology/Sets/Closeds.lean | 2 +- Mathlib/Topology/Sets/Compacts.lean | 2 +- Mathlib/Topology/ShrinkingLemma.lean | 4 +- Mathlib/Topology/StoneCech.lean | 2 +- Mathlib/Topology/SubsetProperties.lean | 18 ++++---- Mathlib/Topology/Support.lean | 2 +- Mathlib/Topology/UniformSpace/Compact.lean | 6 +-- Mathlib/Topology/UrysohnsLemma.lean | 2 +- 158 files changed, 528 insertions(+), 525 deletions(-) diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index f997f0da5a552..252e5ebff765c 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -261,9 +261,9 @@ theorem headI_mem_of_nonempty {α : Type _} [Inhabited α] : ∀ {l : List α} ( #align ballot.head_mem_of_nonempty Ballot.headI_mem_of_nonempty theorem first_vote_neg (p q : ℕ) (h : 0 < p + q) : - condCount (countedSequence p q) ({l | l.headI = 1}ᶜ) = q / (p + q) := by + condCount (countedSequence p q) {l | l.headI = 1}ᶜ = q / (p + q) := by have := condCount_compl - ({l : List ℤ | l.headI = 1}ᶜ) (countedSequence_finite p q) (countedSequence_nonempty p q) + {l : List ℤ | l.headI = 1}ᶜ (countedSequence_finite p q) (countedSequence_nonempty p q) rw [compl_compl, first_vote_pos _ _ h] at this rw [(_ : (q / (p + q) : ENNReal) = 1 - p / (p + q)), ← this, ENNReal.add_sub_cancel_right] · simp only [Ne.def, ENNReal.div_eq_top, Nat.cast_eq_zero, add_eq_zero_iff, ENNReal.nat_ne_top, diff --git a/Mathlib/Algebra/Algebra/Spectrum.lean b/Mathlib/Algebra/Algebra/Spectrum.lean index 6566d752d7091..bbe431a36c5a9 100644 --- a/Mathlib/Algebra/Algebra/Spectrum.lean +++ b/Mathlib/Algebra/Algebra/Spectrum.lean @@ -70,7 +70,7 @@ algebra `A`. The spectrum is simply the complement of the resolvent set. -/ def spectrum (a : A) : Set R := - resolventSet R aᶜ + (resolventSet R a)ᶜ #align spectrum spectrum variable {R} diff --git a/Mathlib/Algebra/BigOperators/Basic.lean b/Mathlib/Algebra/BigOperators/Basic.lean index 1ecc4252b4641..2719620afd007 100644 --- a/Mathlib/Algebra/BigOperators/Basic.lean +++ b/Mathlib/Algebra/BigOperators/Basic.lean @@ -2026,7 +2026,7 @@ theorem prod_subtype_mul_prod_subtype {α β : Type _} [Fintype α] [CommMonoid ((∏ i : { x // p x }, f i) * ∏ i : { x // ¬p x }, f i) = ∏ i, f i := by classical let s := { x | p x }.toFinset - rw [← Finset.prod_subtype s, ← Finset.prod_subtype (sᶜ)] + rw [← Finset.prod_subtype s, ← Finset.prod_subtype sᶜ] · exact Finset.prod_mul_prod_compl _ _ · simp · simp diff --git a/Mathlib/Algebra/IndicatorFunction.lean b/Mathlib/Algebra/IndicatorFunction.lean index 82808a07cd005..6cde5c43bca77 100644 --- a/Mathlib/Algebra/IndicatorFunction.lean +++ b/Mathlib/Algebra/IndicatorFunction.lean @@ -415,28 +415,28 @@ theorem mulIndicator_mul' (s : Set α) (f g : α → M) : @[to_additive (attr := simp)] theorem mulIndicator_compl_mul_self_apply (s : Set α) (f : α → M) (a : α) : - mulIndicator (sᶜ) f a * mulIndicator s f a = f a := + mulIndicator sᶜ f a * mulIndicator s f a = f a := by_cases (fun ha : a ∈ s => by simp [ha]) fun ha => by simp [ha] #align set.mul_indicator_compl_mul_self_apply Set.mulIndicator_compl_mul_self_apply #align set.indicator_compl_add_self_apply Set.indicator_compl_add_self_apply @[to_additive (attr := simp)] theorem mulIndicator_compl_mul_self (s : Set α) (f : α → M) : - mulIndicator (sᶜ) f * mulIndicator s f = f := + mulIndicator sᶜ f * mulIndicator s f = f := funext <| mulIndicator_compl_mul_self_apply s f #align set.mul_indicator_compl_mul_self Set.mulIndicator_compl_mul_self #align set.indicator_compl_add_self Set.indicator_compl_add_self @[to_additive (attr := simp)] theorem mulIndicator_self_mul_compl_apply (s : Set α) (f : α → M) (a : α) : - mulIndicator s f a * mulIndicator (sᶜ) f a = f a := + mulIndicator s f a * mulIndicator sᶜ f a = f a := by_cases (fun ha : a ∈ s => by simp [ha]) fun ha => by simp [ha] #align set.mul_indicator_self_mul_compl_apply Set.mulIndicator_self_mul_compl_apply #align set.indicator_self_add_compl_apply Set.indicator_self_add_compl_apply @[to_additive (attr := simp)] theorem mulIndicator_self_mul_compl (s : Set α) (f : α → M) : - mulIndicator s f * mulIndicator (sᶜ) f = f := + mulIndicator s f * mulIndicator sᶜ f = f := funext <| mulIndicator_self_mul_compl_apply s f #align set.mul_indicator_self_mul_compl Set.mulIndicator_self_mul_compl #align set.indicator_self_add_compl Set.indicator_self_add_compl @@ -546,13 +546,13 @@ theorem mulIndicator_div' (s : Set α) (f g : α → G) : @[to_additive indicator_compl'] theorem mulIndicator_compl (s : Set α) (f : α → G) : - mulIndicator (sᶜ) f = f * (mulIndicator s f)⁻¹ := + mulIndicator sᶜ f = f * (mulIndicator s f)⁻¹ := eq_mul_inv_of_mul_eq <| s.mulIndicator_compl_mul_self f #align set.mul_indicator_compl Set.mulIndicator_compl #align set.indicator_compl' Set.indicator_compl' theorem indicator_compl {G} [AddGroup G] (s : Set α) (f : α → G) : - indicator (sᶜ) f = f - indicator s f := by rw [sub_eq_add_neg, indicator_compl'] + indicator sᶜ f = f - indicator s f := by rw [sub_eq_add_neg, indicator_compl'] #align set.indicator_compl Set.indicator_compl @[to_additive indicator_diff'] diff --git a/Mathlib/Algebra/Ring/BooleanRing.lean b/Mathlib/Algebra/Ring/BooleanRing.lean index 5189c02950c0b..4065b973c987b 100644 --- a/Mathlib/Algebra/Ring/BooleanRing.lean +++ b/Mathlib/Algebra/Ring/BooleanRing.lean @@ -300,7 +300,7 @@ theorem ofBoolAlg_inf (a b : AsBoolAlg α) : ofBoolAlg (a ⊓ b) = ofBoolAlg a * #align of_boolalg_inf ofBoolAlg_inf @[simp] -theorem ofBoolAlg_compl (a : AsBoolAlg α) : ofBoolAlg (aᶜ) = 1 + ofBoolAlg a := +theorem ofBoolAlg_compl (a : AsBoolAlg α) : ofBoolAlg aᶜ = 1 + ofBoolAlg a := rfl #align of_boolalg_compl ofBoolAlg_compl diff --git a/Mathlib/Algebra/Ring/Idempotents.lean b/Mathlib/Algebra/Ring/Idempotents.lean index 8e0c3784c47c3..63dc291f32d2a 100644 --- a/Mathlib/Algebra/Ring/Idempotents.lean +++ b/Mathlib/Algebra/Ring/Idempotents.lean @@ -122,7 +122,7 @@ instance : HasCompl { p : R // IsIdempotentElem p } := ⟨fun p => ⟨1 - p, p.prop.one_sub⟩⟩ @[simp] -theorem coe_compl (p : { p : R // IsIdempotentElem p }) : ↑(pᶜ) = (1 : R) - ↑p := +theorem coe_compl (p : { p : R // IsIdempotentElem p }) : ↑pᶜ = (1 : R) - ↑p := rfl #align is_idempotent_elem.coe_compl IsIdempotentElem.coe_compl diff --git a/Mathlib/Algebra/Star/Pointwise.lean b/Mathlib/Algebra/Star/Pointwise.lean index eb3ee5ced861a..d1b69a3939680 100644 --- a/Mathlib/Algebra/Star/Pointwise.lean +++ b/Mathlib/Algebra/Star/Pointwise.lean @@ -94,7 +94,7 @@ theorem iUnion_star {ι : Sort _} [Star α] (s : ι → Set α) : (⋃ i, s i) #align set.Union_star Set.iUnion_star @[simp] -theorem compl_star [Star α] : (sᶜ)⋆ = s⋆ᶜ := preimage_compl +theorem compl_star [Star α] : sᶜ⋆ = s⋆ᶜ := preimage_compl #align set.compl_star Set.compl_star @[simp] diff --git a/Mathlib/Algebra/Support.lean b/Mathlib/Algebra/Support.lean index 78efab91a2ef2..6d3a3ebd2ea79 100644 --- a/Mathlib/Algebra/Support.lean +++ b/Mathlib/Algebra/Support.lean @@ -60,7 +60,7 @@ theorem nmem_mulSupport {f : α → M} {x : α} : x ∉ mulSupport f ↔ f x = 1 #align function.nmem_support Function.nmem_support @[to_additive] -theorem compl_mulSupport {f : α → M} : mulSupport fᶜ = { x | f x = 1 } := +theorem compl_mulSupport {f : α → M} : (mulSupport f)ᶜ = { x | f x = 1 } := ext fun _ => nmem_mulSupport #align function.compl_mul_support Function.compl_mulSupport #align function.compl_support Function.compl_support diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index f2a1d18982c91..03545092c8822 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -415,7 +415,7 @@ instance zariskiTopology : TopologicalSpace (PrimeSpectrum R) := #align prime_spectrum.zariski_topology PrimeSpectrum.zariskiTopology theorem isOpen_iff (U : Set (PrimeSpectrum R)) : IsOpen U ↔ ∃ s, Uᶜ = zeroLocus s := by - simp only [@eq_comm _ (Uᶜ)] ; rfl + simp only [@eq_comm _ Uᶜ] ; rfl #align prime_spectrum.is_open_iff PrimeSpectrum.isOpen_iff theorem isClosed_iff_zeroLocus (Z : Set (PrimeSpectrum R)) : IsClosed Z ↔ ∃ s, Z = zeroLocus s := by @@ -778,7 +778,7 @@ theorem isOpen_basicOpen {a : R} : IsOpen (basicOpen a : Set (PrimeSpectrum R)) @[simp] theorem basicOpen_eq_zeroLocus_compl (r : R) : - (basicOpen r : Set (PrimeSpectrum R)) = zeroLocus {r}ᶜ := + (basicOpen r : Set (PrimeSpectrum R)) = (zeroLocus {r})ᶜ := Set.ext fun x => by simp only [SetLike.mem_coe, mem_basicOpen, Set.mem_compl_iff, mem_zeroLocus, Set.singleton_subset_iff] #align prime_spectrum.basic_open_eq_zero_locus_compl PrimeSpectrum.basicOpen_eq_zeroLocus_compl @@ -853,7 +853,7 @@ theorem isCompact_basicOpen (f : R) : IsCompact (basicOpen f : Set (PrimeSpectru rcases Submodule.exists_finset_of_mem_iSup I hn with ⟨s, hs⟩ use s -- Using simp_rw here, because `hI` and `zero_locus_supr` need to be applied underneath binders - simp_rw [basicOpen_eq_zeroLocus_compl f, Set.inter_comm (zeroLocus {f}ᶜ), ← Set.diff_eq, + simp_rw [basicOpen_eq_zeroLocus_compl f, Set.inter_comm (zeroLocus {f})ᶜ, ← Set.diff_eq, Set.diff_eq_empty] rw [show (Set.iInter fun i => Set.iInter fun (_ : i ∈ s) => Z i) = Set.iInter fun i => Set.iInter fun (_ : i ∈ s) => zeroLocus (I i) from congr_arg _ diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/IsOpenComapC.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/IsOpenComapC.lean index 8741ec2dca0d8..b7637ea664646 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/IsOpenComapC.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/IsOpenComapC.lean @@ -55,7 +55,7 @@ theorem comap_C_mem_imageOfDf {I : PrimeSpectrum R[X]} /-- The open set `imageOfDf f` coincides with the image of `basicOpen f` under the morphism `C⁺ : Spec R[x] → Spec R`. -/ theorem imageOfDf_eq_comap_C_compl_zeroLocus : - imageOfDf f = PrimeSpectrum.comap (C : R →+* R[X]) '' zeroLocus {f}ᶜ := by + imageOfDf f = PrimeSpectrum.comap (C : R →+* R[X]) '' (zeroLocus {f})ᶜ := by ext x refine' ⟨fun hx => ⟨⟨map C x.asIdeal, isPrime_map_C_of_isPrime x.IsPrime⟩, ⟨_, _⟩⟩, _⟩ · rw [mem_compl_iff, mem_zeroLocus, singleton_subset_iff] diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean index 45fdc8e1ef3d9..5a4ddeec74053 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean @@ -348,7 +348,7 @@ set_option linter.uppercaseLean3 false in #align projective_spectrum.Top ProjectiveSpectrum.top theorem isOpen_iff (U : Set (ProjectiveSpectrum 𝒜)) : IsOpen U ↔ ∃ s, Uᶜ = zeroLocus 𝒜 s := by - simp only [@eq_comm _ (Uᶜ)] ; rfl + simp only [@eq_comm _ Uᶜ] ; rfl #align projective_spectrum.is_open_iff ProjectiveSpectrum.isOpen_iff theorem isClosed_iff_zeroLocus (Z : Set (ProjectiveSpectrum 𝒜)) : @@ -405,7 +405,7 @@ theorem isOpen_basicOpen {a : A} : IsOpen (basicOpen 𝒜 a : Set (ProjectiveSpe @[simp] theorem basicOpen_eq_zeroLocus_compl (r : A) : - (basicOpen 𝒜 r : Set (ProjectiveSpectrum 𝒜)) = zeroLocus 𝒜 {r}ᶜ := + (basicOpen 𝒜 r : Set (ProjectiveSpectrum 𝒜)) = (zeroLocus 𝒜 {r})ᶜ := Set.ext fun x => by simp only [Set.mem_compl_iff, mem_zeroLocus, Set.singleton_subset_iff] ; rfl #align projective_spectrum.basic_open_eq_zero_locus_compl ProjectiveSpectrum.basicOpen_eq_zeroLocus_compl diff --git a/Mathlib/Analysis/Calculus/ContDiff.lean b/Mathlib/Analysis/Calculus/ContDiff.lean index 450e0bef1f91d..54d5b7eca95b2 100644 --- a/Mathlib/Analysis/Calculus/ContDiff.lean +++ b/Mathlib/Analysis/Calculus/ContDiff.lean @@ -1688,7 +1688,7 @@ theorem contDiffAt_inv {x : 𝕜'} (hx : x ≠ 0) {n} : ContDiffAt 𝕜 n Inv.in simpa only [Ring.inverse_eq_inv'] using contDiffAt_ring_inverse 𝕜 (Units.mk0 x hx) #align cont_diff_at_inv contDiffAt_inv -theorem contDiffOn_inv {n} : ContDiffOn 𝕜 n (Inv.inv : 𝕜' → 𝕜') ({0}ᶜ) := fun _ hx => +theorem contDiffOn_inv {n} : ContDiffOn 𝕜 n (Inv.inv : 𝕜' → 𝕜') {0}ᶜ := fun _ hx => (contDiffAt_inv 𝕜 hx).contDiffWithinAt #align cont_diff_on_inv contDiffOn_inv diff --git a/Mathlib/Analysis/Calculus/Deriv/Slope.lean b/Mathlib/Analysis/Calculus/Deriv/Slope.lean index 38626b2feead6..aab69558c517a 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Slope.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Slope.lean @@ -59,17 +59,17 @@ variable {L L₁ L₂ : Filter 𝕜} definition with a limit. In this version we have to take the limit along the subset `-{x}`, because for `y=x` the slope equals zero due to the convention `0⁻¹=0`. -/ theorem hasDerivAtFilter_iff_tendsto_slope {x : 𝕜} {L : Filter 𝕜} : - HasDerivAtFilter f f' x L ↔ Tendsto (slope f x) (L ⊓ 𝓟 ({x}ᶜ)) (𝓝 f') := + HasDerivAtFilter f f' x L ↔ Tendsto (slope f x) (L ⊓ 𝓟 {x}ᶜ) (𝓝 f') := calc HasDerivAtFilter f f' x L ↔ Tendsto (fun y ↦ slope f x y - (y - x)⁻¹ • (y - x) • f') L (𝓝 0) := by simp only [hasDerivAtFilter_iff_tendsto, ← norm_inv, ← norm_smul, ← tendsto_zero_iff_norm_tendsto_zero, slope_def_module, smul_sub] - _ ↔ Tendsto (fun y ↦ slope f x y - (y - x)⁻¹ • (y - x) • f') (L ⊓ 𝓟 ({x}ᶜ)) (𝓝 0) := + _ ↔ Tendsto (fun y ↦ slope f x y - (y - x)⁻¹ • (y - x) • f') (L ⊓ 𝓟 {x}ᶜ) (𝓝 0) := .symm <| tendsto_inf_principal_nhds_iff_of_forall_eq <| by simp - _ ↔ Tendsto (fun y ↦ slope f x y - f') (L ⊓ 𝓟 ({x}ᶜ)) (𝓝 0) := tendsto_congr' <| by + _ ↔ Tendsto (fun y ↦ slope f x y - f') (L ⊓ 𝓟 {x}ᶜ) (𝓝 0) := tendsto_congr' <| by refine (EqOn.eventuallyEq fun y hy ↦ ?_).filter_mono inf_le_right rw [inv_smul_smul₀ (sub_ne_zero.2 hy) f'] - _ ↔ Tendsto (slope f x) (L ⊓ 𝓟 ({x}ᶜ)) (𝓝 f') := + _ ↔ Tendsto (slope f x) (L ⊓ 𝓟 {x}ᶜ) (𝓝 f') := by rw [← nhds_translation_sub f', tendsto_comap_iff]; rfl #align has_deriv_at_filter_iff_tendsto_slope hasDerivAtFilter_iff_tendsto_slope diff --git a/Mathlib/Analysis/Calculus/Dslope.lean b/Mathlib/Analysis/Calculus/Dslope.lean index d56283ff9952c..2b5af13602a7d 100644 --- a/Mathlib/Analysis/Calculus/Dslope.lean +++ b/Mathlib/Analysis/Calculus/Dslope.lean @@ -55,7 +55,7 @@ theorem ContinuousLinearMap.dslope_comp {F : Type _} [NormedAddCommGroup F] [Nor · simpa only [dslope_of_ne _ hne] using f.toLinearMap.slope_comp g a b #align continuous_linear_map.dslope_comp ContinuousLinearMap.dslope_comp -theorem eqOn_dslope_slope (f : 𝕜 → E) (a : 𝕜) : EqOn (dslope f a) (slope f a) ({a}ᶜ) := fun _ => +theorem eqOn_dslope_slope (f : 𝕜 → E) (a : 𝕜) : EqOn (dslope f a) (slope f a) {a}ᶜ := fun _ => dslope_of_ne f #align eq_on_dslope_slope eqOn_dslope_slope @@ -78,7 +78,7 @@ theorem dslope_sub_smul_of_ne (f : 𝕜 → E) (h : b ≠ a) : #align dslope_sub_smul_of_ne dslope_sub_smul_of_ne theorem eqOn_dslope_sub_smul (f : 𝕜 → E) (a : 𝕜) : - EqOn (dslope (fun x => (x - a) • f x) a) f ({a}ᶜ) := fun _ => dslope_sub_smul_of_ne f + EqOn (dslope (fun x => (x - a) • f x) a) f {a}ᶜ := fun _ => dslope_sub_smul_of_ne f #align eq_on_dslope_sub_smul eqOn_dslope_sub_smul theorem dslope_sub_smul [DecidableEq 𝕜] (f : 𝕜 → E) (a : 𝕜) : diff --git a/Mathlib/Analysis/Convex/Integral.lean b/Mathlib/Analysis/Convex/Integral.lean index fd1643b0e8b93..26440bbcb1abe 100644 --- a/Mathlib/Analysis/Convex/Integral.lean +++ b/Mathlib/Analysis/Convex/Integral.lean @@ -237,13 +237,13 @@ theorem ConcaveOn.le_map_integral [IsProbabilityMeasure μ] (hg : ConcaveOn ℝ values of `f` over `t` and `tᶜ` are different. -/ theorem ae_eq_const_or_exists_average_ne_compl [IsFiniteMeasure μ] (hfi : Integrable f μ) : f =ᵐ[μ] const α (⨍ x, f x ∂μ) ∨ - ∃ t, MeasurableSet t ∧ μ t ≠ 0 ∧ μ (tᶜ) ≠ 0 ∧ (⨍ x in t, f x ∂μ) ≠ ⨍ x in tᶜ, f x ∂μ := by + ∃ t, MeasurableSet t ∧ μ t ≠ 0 ∧ μ tᶜ ≠ 0 ∧ (⨍ x in t, f x ∂μ) ≠ ⨍ x in tᶜ, f x ∂μ := by refine' or_iff_not_imp_right.mpr fun H => _; push_neg at H refine' hfi.ae_eq_of_forall_set_integral_eq _ _ (integrable_const _) fun t ht ht' => _; clear ht' simp only [const_apply, set_integral_const] by_cases h₀ : μ t = 0 · rw [restrict_eq_zero.2 h₀, integral_zero_measure, h₀, ENNReal.zero_toReal, zero_smul] - by_cases h₀' : μ (tᶜ) = 0 + by_cases h₀' : μ tᶜ = 0 · rw [← ae_eq_univ] at h₀' rw [restrict_congr_set h₀', restrict_univ, measure_congr h₀', measure_smul_average] have := average_mem_openSegment_compl_self ht.nullMeasurableSet h₀ h₀' hfi @@ -258,7 +258,7 @@ theorem Convex.average_mem_interior_of_set [IsFiniteMeasure μ] (hs : Convex ℝ (hfs : ∀ᵐ x ∂μ, f x ∈ s) (hfi : Integrable f μ) (ht : (⨍ x in t, f x ∂μ) ∈ interior s) : (⨍ x, f x ∂μ) ∈ interior s := by rw [← measure_toMeasurable] at h0 ; rw [← restrict_toMeasurable (measure_ne_top μ t)] at ht - by_cases h0' : μ (toMeasurable μ tᶜ) = 0 + by_cases h0' : μ (toMeasurable μ t)ᶜ = 0 · rw [← ae_eq_univ] at h0' rwa [restrict_congr_set h0', restrict_univ] at ht exact diff --git a/Mathlib/Analysis/Convex/StoneSeparation.lean b/Mathlib/Analysis/Convex/StoneSeparation.lean index 8e14a07a86792..9b25ee4f28001 100644 --- a/Mathlib/Analysis/Convex/StoneSeparation.lean +++ b/Mathlib/Analysis/Convex/StoneSeparation.lean @@ -84,7 +84,7 @@ theorem not_disjoint_segment_convexHull_triple {p q u v x y z : E} (hz : z ∈ s /-- **Stone's Separation Theorem** -/ theorem exists_convex_convex_compl_subset (hs : Convex 𝕜 s) (ht : Convex 𝕜 t) (hst : Disjoint s t) : - ∃ C : Set E, Convex 𝕜 C ∧ Convex 𝕜 (Cᶜ) ∧ s ⊆ C ∧ t ⊆ Cᶜ := by + ∃ C : Set E, Convex 𝕜 C ∧ Convex 𝕜 Cᶜ ∧ s ⊆ C ∧ t ⊆ Cᶜ := by let S : Set (Set E) := { C | Convex 𝕜 C ∧ Disjoint C t } obtain ⟨C, hC, hsC, hCmax⟩ := zorn_subset_nonempty S diff --git a/Mathlib/Analysis/NormedSpace/FiniteDimension.lean b/Mathlib/Analysis/NormedSpace/FiniteDimension.lean index 8a7f4c4f1a375..0c63c5d9cc200 100644 --- a/Mathlib/Analysis/NormedSpace/FiniteDimension.lean +++ b/Mathlib/Analysis/NormedSpace/FiniteDimension.lean @@ -621,7 +621,7 @@ instance (priority := 900) FiniteDimensional.proper_real (E : Type u) [NormedAdd `IsCompact.exists_mem_frontier_infDist_compl_eq_dist`. -/ theorem exists_mem_frontier_infDist_compl_eq_dist {E : Type _} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] {x : E} {s : Set E} (hx : x ∈ s) (hs : s ≠ univ) : - ∃ y ∈ frontier s, Metric.infDist x (sᶜ) = dist x y := by + ∃ y ∈ frontier s, Metric.infDist x sᶜ = dist x y := by rcases Metric.exists_mem_closure_infDist_eq_dist (nonempty_compl.2 hs) x with ⟨y, hys, hyd⟩ rw [closure_compl] at hys refine' @@ -638,7 +638,7 @@ theorem exists_mem_frontier_infDist_compl_eq_dist {E : Type _} [NormedAddCommGro nonrec theorem IsCompact.exists_mem_frontier_infDist_compl_eq_dist {E : Type _} [NormedAddCommGroup E] [NormedSpace ℝ E] [Nontrivial E] {x : E} {K : Set E} (hK : IsCompact K) (hx : x ∈ K) : - ∃ y ∈ frontier K, Metric.infDist x (Kᶜ) = dist x y := by + ∃ y ∈ frontier K, Metric.infDist x Kᶜ = dist x y := by obtain hx' | hx' : x ∈ interior K ∪ frontier K := by rw [← closure_eq_interior_union_frontier] exact subset_closure hx diff --git a/Mathlib/Analysis/NormedSpace/MStructure.lean b/Mathlib/Analysis/NormedSpace/MStructure.lean index bbc5b4d760c9c..f558b2a7fc81e 100644 --- a/Mathlib/Analysis/NormedSpace/MStructure.lean +++ b/Mathlib/Analysis/NormedSpace/MStructure.lean @@ -177,7 +177,7 @@ instance Subtype.hasCompl : HasCompl { f : M // IsLprojection X f } := ⟨fun P => ⟨1 - P, P.prop.Lcomplement⟩⟩ @[simp] -theorem coe_compl (P : { P : M // IsLprojection X P }) : ↑(Pᶜ) = (1 : M) - ↑P := +theorem coe_compl (P : { P : M // IsLprojection X P }) : ↑Pᶜ = (1 : M) - ↑P := rfl #align is_Lprojection.coe_compl IsLprojection.coe_compl @@ -260,22 +260,22 @@ theorem coe_top [FaithfulSMul M X] : rfl #align is_Lprojection.coe_top IsLprojection.coe_top -theorem compl_mul {P : { P : M // IsLprojection X P }} {Q : M} : ↑(Pᶜ) * Q = Q - ↑P * Q := by +theorem compl_mul {P : { P : M // IsLprojection X P }} {Q : M} : ↑Pᶜ * Q = Q - ↑P * Q := by rw [coe_compl, sub_mul, one_mul] #align is_Lprojection.compl_mul IsLprojection.compl_mul -theorem mul_compl_self {P : { P : M // IsLprojection X P }} : (↑P : M) * ↑(Pᶜ) = 0 := by +theorem mul_compl_self {P : { P : M // IsLprojection X P }} : (↑P : M) * ↑Pᶜ = 0 := by rw [coe_compl, mul_sub, mul_one, P.prop.proj.eq, sub_self] #align is_Lprojection.mul_compl_self IsLprojection.mul_compl_self theorem distrib_lattice_lemma [FaithfulSMul M X] {P Q R : { P : M // IsLprojection X P }} : - ((↑P : M) + ↑(Pᶜ) * R) * (↑P + ↑Q * ↑R * ↑(Pᶜ)) = ↑P + ↑Q * ↑R * ↑(Pᶜ) := by - rw [add_mul, mul_add, mul_add, (mul_assoc _ (R : M) (↑Q * ↑R * ↑(Pᶜ))), + ((↑P : M) + ↑Pᶜ * R) * (↑P + ↑Q * ↑R * ↑Pᶜ) = ↑P + ↑Q * ↑R * ↑Pᶜ := by + rw [add_mul, mul_add, mul_add, (mul_assoc _ (R : M) (↑Q * ↑R * ↑Pᶜ)), ← mul_assoc (R : M) (↑Q * ↑R) _, ← coe_inf Q, (Pᶜ.prop.commute R.prop).eq, - ((Q ⊓ R).prop.commute (Pᶜ).prop).eq, (R.prop.commute (Q ⊓ R).prop).eq, coe_inf Q, + ((Q ⊓ R).prop.commute Pᶜ.prop).eq, (R.prop.commute (Q ⊓ R).prop).eq, coe_inf Q, mul_assoc (Q : M), ←mul_assoc, mul_assoc (R : M), (Pᶜ.prop.commute P.prop).eq, mul_compl_self, MulZeroClass.zero_mul, MulZeroClass.mul_zero, zero_add, add_zero, ← mul_assoc, P.prop.proj.eq, - R.prop.proj.eq, ←coe_inf Q, mul_assoc, ((Q ⊓ R).prop.commute (Pᶜ).prop).eq, ← mul_assoc, + R.prop.proj.eq, ←coe_inf Q, mul_assoc, ((Q ⊓ R).prop.commute Pᶜ.prop).eq, ← mul_assoc, Pᶜ.prop.proj.eq] #align is_Lprojection.distrib_lattice_lemma IsLprojection.distrib_lattice_lemma @@ -307,15 +307,15 @@ instance [FaithfulSMul M X] : Lattice { P : M // IsLprojection X P } where instance Subtype.distribLattice [FaithfulSMul M X] : DistribLattice { P : M // IsLprojection X P } where le_sup_inf P Q R := by - have e₁ : ↑((P ⊔ Q) ⊓ (P ⊔ R)) = ↑P + ↑Q * (R : M) * ↑(Pᶜ) := by + have e₁ : ↑((P ⊔ Q) ⊓ (P ⊔ R)) = ↑P + ↑Q * (R : M) * ↑Pᶜ := by rw [coe_inf, coe_sup, coe_sup, ← add_sub, ← add_sub, ← compl_mul, ← compl_mul, add_mul, - mul_add, ((Pᶜ).prop.commute Q.prop).eq, mul_add, ← mul_assoc, mul_assoc (Q: M), - ((Pᶜ).prop.commute P.prop).eq, mul_compl_self, MulZeroClass.zero_mul, MulZeroClass.mul_zero, - zero_add, add_zero, ← mul_assoc, mul_assoc (Q : M), P.prop.proj.eq, (Pᶜ).prop.proj.eq, - mul_assoc, ((Pᶜ).prop.commute R.prop).eq, ← mul_assoc] - have e₂ : ↑((P ⊔ Q) ⊓ (P ⊔ R)) * ↑(P ⊔ Q ⊓ R) = (P : M) + ↑Q * ↑R * ↑(Pᶜ) := by + mul_add, (Pᶜ.prop.commute Q.prop).eq, mul_add, ← mul_assoc, mul_assoc (Q: M), + (Pᶜ.prop.commute P.prop).eq, mul_compl_self, MulZeroClass.zero_mul, MulZeroClass.mul_zero, + zero_add, add_zero, ← mul_assoc, mul_assoc (Q : M), P.prop.proj.eq, Pᶜ.prop.proj.eq, + mul_assoc, (Pᶜ.prop.commute R.prop).eq, ← mul_assoc] + have e₂ : ↑((P ⊔ Q) ⊓ (P ⊔ R)) * ↑(P ⊔ Q ⊓ R) = (P : M) + ↑Q * ↑R * ↑Pᶜ := by rw [coe_inf, coe_sup, coe_sup, coe_sup, ← add_sub, ← add_sub, ← add_sub, ← compl_mul, ← - compl_mul, ← compl_mul, ((Pᶜ).prop.commute (Q ⊓ R).prop).eq, coe_inf, mul_assoc, + compl_mul, ← compl_mul, (Pᶜ.prop.commute (Q ⊓ R).prop).eq, coe_inf, mul_assoc, distrib_lattice_lemma, (Q.prop.commute R.prop).eq, distrib_lattice_lemma] rw [le_def, e₁, coe_inf, e₂] diff --git a/Mathlib/Analysis/NormedSpace/Multilinear.lean b/Mathlib/Analysis/NormedSpace/Multilinear.lean index 2a06f70e40bb0..17d4643315439 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear.lean @@ -270,7 +270,7 @@ theorem restr_norm_le {k n : ℕ} (f : (MultilinearMap 𝕜 (fun _ : Fin n => G) rw [mul_right_comm, mul_assoc] convert H _ using 2 simp only [apply_dite norm, Fintype.prod_dite, prod_const ‖z‖, Finset.card_univ, - Fintype.card_of_subtype (sᶜ) fun _ => mem_compl, card_compl, Fintype.card_fin, hk, mk_coe, ← + Fintype.card_of_subtype sᶜ fun _ => mem_compl, card_compl, Fintype.card_fin, hk, mk_coe, ← (s.orderIsoOfFin hk).symm.bijective.prod_comp fun x => ‖v x‖] convert rfl #align multilinear_map.restr_norm_le MultilinearMap.restr_norm_le diff --git a/Mathlib/Analysis/NormedSpace/PiLp.lean b/Mathlib/Analysis/NormedSpace/PiLp.lean index 9d466cd112407..cb19bf412de59 100644 --- a/Mathlib/Analysis/NormedSpace/PiLp.lean +++ b/Mathlib/Analysis/NormedSpace/PiLp.lean @@ -490,7 +490,7 @@ instance [∀ i, EMetricSpace (α i)] : EMetricSpace (PiLp p α) := `L^p` distance, and having as uniformity the product uniformity. -/ instance [∀ i, PseudoMetricSpace (β i)] : PseudoMetricSpace (PiLp p β) := ((pseudoMetricAux p β).replaceUniformity (aux_uniformity_eq p β).symm).replaceBornology fun s => - Filter.ext_iff.1 (aux_cobounded_eq p β).symm (sᶜ) + Filter.ext_iff.1 (aux_cobounded_eq p β).symm sᶜ /-- metric space instance on the product of finitely many metric spaces, using the `L^p` distance, and having as uniformity the product uniformity. -/ diff --git a/Mathlib/Analysis/NormedSpace/RieszLemma.lean b/Mathlib/Analysis/NormedSpace/RieszLemma.lean index 0718872a6e75e..d45e93fbb0de3 100644 --- a/Mathlib/Analysis/NormedSpace/RieszLemma.lean +++ b/Mathlib/Analysis/NormedSpace/RieszLemma.lean @@ -111,8 +111,8 @@ theorem riesz_lemma_of_norm_lt {c : 𝕜} (hc : 1 < ‖c‖) {R : ℝ} (hR : ‖ #align riesz_lemma_of_norm_lt riesz_lemma_of_norm_lt theorem Metric.closedBall_infDist_compl_subset_closure {x : F} {s : Set F} (hx : x ∈ s) : - closedBall x (infDist x (sᶜ)) ⊆ closure s := by - cases' eq_or_ne (infDist x (sᶜ)) 0 with h₀ h₀ + closedBall x (infDist x sᶜ) ⊆ closure s := by + cases' eq_or_ne (infDist x sᶜ) 0 with h₀ h₀ · rw [h₀, closedBall_zero'] exact closure_mono (singleton_subset_iff.2 hx) · rw [← closure_ball x h₀] diff --git a/Mathlib/Analysis/NormedSpace/Units.lean b/Mathlib/Analysis/NormedSpace/Units.lean index 305880fb9839d..7a3149c4bc170 100644 --- a/Mathlib/Analysis/NormedSpace/Units.lean +++ b/Mathlib/Analysis/NormedSpace/Units.lean @@ -96,7 +96,7 @@ namespace nonunits /-- The `nonunits` in a complete normed ring are contained in the complement of the ball of radius `1` centered at `1 : R`. -/ -theorem subset_compl_ball : nonunits R ⊆ Metric.ball (1 : R) 1ᶜ := fun x hx h₁ ↦ hx <| +theorem subset_compl_ball : nonunits R ⊆ (Metric.ball (1 : R) 1)ᶜ := fun x hx h₁ ↦ hx <| sub_sub_self 1 x ▸ (Units.oneSub (1 - x) (by rwa [mem_ball_iff_norm'] at h₁)).isUnit #align nonunits.subset_compl_ball nonunits.subset_compl_ball diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index 033a8d4d29095..ae96a6ea92aa6 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -309,7 +309,7 @@ theorem tendsto_log_nhdsWithin_zero : Tendsto log (𝓝[≠] 0) atBot := by simpa [← tendsto_comp_exp_atBot] using tendsto_id #align real.tendsto_log_nhds_within_zero Real.tendsto_log_nhdsWithin_zero -theorem continuousOn_log : ContinuousOn log ({0}ᶜ) := by +theorem continuousOn_log : ContinuousOn log {0}ᶜ := by simp only [continuousOn_iff_continuous_restrict, restrict] conv in log _ => rw [log_of_ne_zero (show (x : ℝ) ≠ 0 from x.2)] exact expOrderIso.symm.continuous.comp (continuous_subtype_val.norm.subtype_mk _) diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean index 3b679b4fc71d4..74b8218697a98 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean @@ -58,7 +58,7 @@ theorem differentiableAt_log (hx : x ≠ 0) : DifferentiableAt ℝ log x := (hasDerivAt_log hx).differentiableAt #align real.differentiable_at_log Real.differentiableAt_log -theorem differentiableOn_log : DifferentiableOn ℝ log ({0}ᶜ) := fun _x hx => +theorem differentiableOn_log : DifferentiableOn ℝ log {0}ᶜ := fun _x hx => (differentiableAt_log hx).differentiableWithinAt #align real.differentiable_on_log Real.differentiableOn_log @@ -78,8 +78,8 @@ theorem deriv_log' : deriv log = Inv.inv := funext deriv_log #align real.deriv_log' Real.deriv_log' -theorem contDiffOn_log {n : ℕ∞} : ContDiffOn ℝ n log ({0}ᶜ) := by - suffices : ContDiffOn ℝ ⊤ log ({0}ᶜ); exact this.of_le le_top +theorem contDiffOn_log {n : ℕ∞} : ContDiffOn ℝ n log {0}ᶜ := by + suffices : ContDiffOn ℝ ⊤ log {0}ᶜ; exact this.of_le le_top refine' (contDiffOn_top_iff_deriv_of_open isOpen_compl_singleton).2 _ simp [differentiableOn_log, contDiffOn_inv] #align real.cont_diff_on_log Real.contDiffOn_log diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean index 1cf72afb70b8a..bc6fe8cac0ade 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean @@ -64,7 +64,7 @@ theorem cpow_eq_nhds' {p : ℂ × ℂ} (hp_fst : p.fst ≠ 0) : dsimp only rw [cpow_def_of_ne_zero hx] refine' IsOpen.eventually_mem _ hp_fst - change IsOpen ({ x : ℂ × ℂ | x.1 = 0 }ᶜ) + change IsOpen { x : ℂ × ℂ | x.1 = 0 }ᶜ rw [isOpen_compl_iff] exact isClosed_eq continuous_fst continuous_const #align cpow_eq_nhds' cpow_eq_nhds' diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean index 292224d23e304..9258c569d4c1d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean @@ -96,7 +96,7 @@ theorem differentiableWithinAt_arcsin_Ici {x : ℝ} : theorem differentiableWithinAt_arcsin_Iic {x : ℝ} : DifferentiableWithinAt ℝ arcsin (Iic x) x ↔ x ≠ 1 := by refine' ⟨fun h => _, fun h => (hasDerivWithinAt_arcsin_Iic h).differentiableWithinAt⟩ - rw [← neg_neg x, ← image_neg_Ici] at h + rw [← neg_neg x, ← image_neg_Ici] at h have := (h.comp (-x) differentiableWithinAt_id.neg (mapsTo_image _ _)).neg simpa [(· ∘ ·), differentiableWithinAt_arcsin_Ici] using this #align real.differentiable_within_at_arcsin_Iic Real.differentiableWithinAt_arcsin_Iic @@ -113,16 +113,16 @@ theorem deriv_arcsin : deriv arcsin = fun x => 1 / sqrt (1 - x ^ 2) := by by_cases h : x ≠ -1 ∧ x ≠ 1 · exact (hasDerivAt_arcsin h.1 h.2).deriv · rw [deriv_zero_of_not_differentiableAt (mt differentiableAt_arcsin.1 h)] - simp only [not_and_or, Ne.def, Classical.not_not] at h + simp only [not_and_or, Ne.def, Classical.not_not] at h rcases h with (rfl | rfl) <;> simp #align real.deriv_arcsin Real.deriv_arcsin -theorem differentiableOn_arcsin : DifferentiableOn ℝ arcsin ({-1, 1}ᶜ) := fun _x hx => +theorem differentiableOn_arcsin : DifferentiableOn ℝ arcsin {-1, 1}ᶜ := fun _x hx => (differentiableAt_arcsin.2 ⟨fun h => hx (Or.inl h), fun h => hx (Or.inr h)⟩).differentiableWithinAt #align real.differentiable_on_arcsin Real.differentiableOn_arcsin -theorem contDiffOn_arcsin {n : ℕ∞} : ContDiffOn ℝ n arcsin ({-1, 1}ᶜ) := fun _x hx => +theorem contDiffOn_arcsin {n : ℕ∞} : ContDiffOn ℝ n arcsin {-1, 1}ᶜ := fun _x hx => (contDiffAt_arcsin (mt Or.inl hx) (mt Or.inr hx)).contDiffWithinAt #align real.cont_diff_on_arcsin Real.contDiffOn_arcsin @@ -180,11 +180,11 @@ theorem deriv_arccos : deriv arccos = fun x => -(1 / sqrt (1 - x ^ 2)) := funext fun x => (deriv_const_sub _).trans <| by simp only [deriv_arcsin] #align real.deriv_arccos Real.deriv_arccos -theorem differentiableOn_arccos : DifferentiableOn ℝ arccos ({-1, 1}ᶜ) := +theorem differentiableOn_arccos : DifferentiableOn ℝ arccos {-1, 1}ᶜ := differentiableOn_arcsin.const_sub _ #align real.differentiable_on_arccos Real.differentiableOn_arccos -theorem contDiffOn_arccos {n : ℕ∞} : ContDiffOn ℝ n arccos ({-1, 1}ᶜ) := +theorem contDiffOn_arccos {n : ℕ∞} : ContDiffOn ℝ n arccos {-1, 1}ᶜ := contDiffOn_const.sub contDiffOn_arcsin #align real.cont_diff_on_arccos Real.contDiffOn_arccos diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index c0be6d13e0dd2..3e1291297ce9e 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -739,7 +739,7 @@ variable {K : Type} [Fintype K] [HasFiniteBiproducts C] (f : K → C) def kernelForkBiproductToSubtype (p : Set K) : LimitCone (parallelPair (biproduct.toSubtype f p) 0) where cone := - KernelFork.ofι (biproduct.fromSubtype f (pᶜ)) + KernelFork.ofι (biproduct.fromSubtype f pᶜ) (by ext j k simp only [Category.assoc, biproduct.ι_fromSubtype_assoc, biproduct.ι_toSubtype_assoc, @@ -747,7 +747,7 @@ def kernelForkBiproductToSubtype (p : Set K) : LimitCone (parallelPair (biproduc erw [dif_neg k.2] simp only [zero_comp]) isLimit := - KernelFork.IsLimit.ofι _ _ (fun {W} g _ => g ≫ biproduct.toSubtype f (pᶜ)) + KernelFork.IsLimit.ofι _ _ (fun {W} g _ => g ≫ biproduct.toSubtype f pᶜ) (by intro W' g' w ext j @@ -766,7 +766,7 @@ instance (p : Set K) : HasKernel (biproduct.toSubtype f p) := /-- The kernel of `biproduct.toSubtype f p` is `⨁ Subtype.restrict pᶜ f`. -/ @[simps!] def kernelBiproductToSubtypeIso (p : Set K) : - kernel (biproduct.toSubtype f p) ≅ ⨁ Subtype.restrict (pᶜ) f := + kernel (biproduct.toSubtype f p) ≅ ⨁ Subtype.restrict pᶜ f := limit.isoLimitCone (kernelForkBiproductToSubtype f p) #align category_theory.limits.kernel_biproduct_to_subtype_iso CategoryTheory.Limits.kernelBiproductToSubtypeIso @@ -776,7 +776,7 @@ def kernelBiproductToSubtypeIso (p : Set K) : def cokernelCoforkBiproductFromSubtype (p : Set K) : ColimitCocone (parallelPair (biproduct.fromSubtype f p) 0) where cocone := - CokernelCofork.ofπ (biproduct.toSubtype f (pᶜ)) + CokernelCofork.ofπ (biproduct.toSubtype f pᶜ) (by ext j k simp only [Category.assoc, Pi.compl_apply, biproduct.ι_fromSubtype_assoc, @@ -785,7 +785,7 @@ def cokernelCoforkBiproductFromSubtype (p : Set K) : simp only [zero_comp] exact not_not.mpr k.2) isColimit := - CokernelCofork.IsColimit.ofπ _ _ (fun {W} g _ => biproduct.fromSubtype f (pᶜ) ≫ g) + CokernelCofork.IsColimit.ofπ _ _ (fun {W} g _ => biproduct.fromSubtype f pᶜ ≫ g) (by intro W g' w ext j @@ -803,7 +803,7 @@ instance (p : Set K) : HasCokernel (biproduct.fromSubtype f p) := /-- The cokernel of `biproduct.fromSubtype f p` is `⨁ Subtype.restrict pᶜ f`. -/ @[simps!] def cokernelBiproductFromSubtypeIso (p : Set K) : - cokernel (biproduct.fromSubtype f p) ≅ ⨁ Subtype.restrict (pᶜ) f := + cokernel (biproduct.fromSubtype f p) ≅ ⨁ Subtype.restrict pᶜ f := colimit.isoColimitCocone (cokernelCoforkBiproductFromSubtype f p) #align category_theory.limits.cokernel_biproduct_from_subtype_iso CategoryTheory.Limits.cokernelBiproductFromSubtypeIso diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean index a22353e451d3a..5a8e83d71a6a2 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean @@ -330,7 +330,7 @@ theorem binaryCofan_isColimit_iff {X Y : Type u} (c : BinaryCofan X Y) : /-- Any monomorphism in `Type` is a coproduct injection. -/ noncomputable def isCoprodOfMono {X Y : Type u} (f : X ⟶ Y) [Mono f] : - IsColimit (BinaryCofan.mk f (Subtype.val : ↑(Set.range fᶜ) → Y)) := by + IsColimit (BinaryCofan.mk f (Subtype.val : ↑(Set.range f)ᶜ → Y)) := by apply Nonempty.some rw [binaryCofan_isColimit_iff] refine' ⟨(mono_iff_injective f).mp inferInstance, Subtype.val_injective, _⟩ diff --git a/Mathlib/Combinatorics/Configuration.lean b/Mathlib/Combinatorics/Configuration.lean index 8bfcd0fcaf2be..5914871701b7f 100644 --- a/Mathlib/Combinatorics/Configuration.lean +++ b/Mathlib/Combinatorics/Configuration.lean @@ -143,13 +143,13 @@ theorem Nondegenerate.exists_injective_of_card_le [Nondegenerate P L] [Fintype P obtain ⟨p, hl⟩ := exists_point l rw [Finset.card_singleton, Finset.singleton_biUnion, Nat.one_le_iff_ne_zero] exact Finset.card_ne_zero_of_mem (Set.mem_toFinset.mpr hl) - suffices s.biUnion tᶜ.card ≤ sᶜ.card by + suffices (s.biUnion t)ᶜ.card ≤ sᶜ.card by -- Rephrase in terms of complements (uses `h`) rw [Finset.card_compl, Finset.card_compl, tsub_le_iff_left] at this replace := h.trans this rwa [← add_tsub_assoc_of_le s.card_le_univ, le_tsub_iff_left (le_add_left s.card_le_univ), add_le_add_iff_right] at this - have hs₂ : s.biUnion tᶜ.card ≤ 1 := by + have hs₂ : (s.biUnion t)ᶜ.card ≤ 1 := by -- At most one line through two points of `s` refine' Finset.card_le_one_iff.mpr @fun p₁ p₂ hp₁ hp₂ => _ simp_rw [Finset.mem_compl, Finset.mem_biUnion, exists_prop, not_exists, not_and, diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index 657fdfa111e69..ca2fa01fa222c 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -998,7 +998,7 @@ theorem card_neighborSet_union_compl_neighborSet [Fintype V] (G : SimpleGraph V) #align simple_graph.card_neighbor_set_union_compl_neighbor_set SimpleGraph.card_neighborSet_union_compl_neighborSet theorem neighborSet_compl (G : SimpleGraph V) (v : V) : - Gᶜ.neighborSet v = G.neighborSet vᶜ \ {v} := by + Gᶜ.neighborSet v = (G.neighborSet v)ᶜ \ {v} := by ext w simp [and_comm, eq_comm] #align simple_graph.neighbor_set_compl SimpleGraph.neighborSet_compl @@ -1480,7 +1480,7 @@ theorem neighborFinset_eq_filter {v : V} [DecidableRel G.Adj] : #align simple_graph.neighbor_finset_eq_filter SimpleGraph.neighborFinset_eq_filter theorem neighborFinset_compl [DecidableEq V] [DecidableRel G.Adj] (v : V) : - Gᶜ.neighborFinset v = G.neighborFinset vᶜ \ {v} := by + Gᶜ.neighborFinset v = (G.neighborFinset v)ᶜ \ {v} := by simp only [neighborFinset, neighborSet_compl, Set.toFinset_diff, Set.toFinset_compl, Set.toFinset_singleton] #align simple_graph.neighbor_finset_compl SimpleGraph.neighborFinset_compl diff --git a/Mathlib/Combinatorics/SimpleGraph/Ends/Defs.lean b/Mathlib/Combinatorics/SimpleGraph/Ends/Defs.lean index e640db7739eda..edef85c90431a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Ends/Defs.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Ends/Defs.lean @@ -29,7 +29,7 @@ namespace SimpleGraph /-- The components outside a given set of vertices `K` -/ @[reducible] def ComponentCompl := - (G.induce (Kᶜ)).ConnectedComponent + (G.induce Kᶜ).ConnectedComponent #align simple_graph.component_compl SimpleGraph.ComponentCompl variable {G} {K L M} @@ -37,7 +37,7 @@ variable {G} {K L M} /-- The connected component of `v` in `G.induce Kᶜ`. -/ @[reducible] def componentComplMk (G : SimpleGraph V) {v : V} (vK : v ∉ K) : G.ComponentCompl K := - connectedComponentMk (G.induce (Kᶜ)) ⟨v, vK⟩ + connectedComponentMk (G.induce Kᶜ) ⟨v, vK⟩ #align simple_graph.component_compl_mk SimpleGraph.componentComplMk /-- The set of vertices of `G` making up the connected component `C` -/ @@ -207,7 +207,7 @@ theorem hom_eq_iff_not_disjoint (C : G.ComponentCompl L) (h : K ⊆ L) (D : G.Co theorem hom_refl (C : G.ComponentCompl L) : C.hom (subset_refl L) = C := by change C.map _ = C - erw [induceHom_id G (Lᶜ), ConnectedComponent.map_id] + erw [induceHom_id G Lᶜ, ConnectedComponent.map_id] #align simple_graph.component_compl.hom_refl SimpleGraph.ComponentCompl.hom_refl theorem hom_trans (C : G.ComponentCompl L) (h : K ⊆ L) (h' : M ⊆ K) : diff --git a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean index d77ed35dd57c7..38ef924b6db6d 100644 --- a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean +++ b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean @@ -115,16 +115,16 @@ set_option linter.uppercaseLean3 false in #align simple_graph.is_SRG_with.card_neighbor_finset_union_of_adj SimpleGraph.IsSRGWith.card_neighborFinset_union_of_adj theorem compl_neighborFinset_sdiff_inter_eq {v w : V} : - G.neighborFinset vᶜ \ {v} ∩ (G.neighborFinset wᶜ \ {w}) = - (G.neighborFinset vᶜ ∩ G.neighborFinset wᶜ) \ ({w} ∪ {v}) := by + (G.neighborFinset v)ᶜ \ {v} ∩ ((G.neighborFinset w)ᶜ \ {w}) = + ((G.neighborFinset v)ᶜ ∩ (G.neighborFinset w)ᶜ) \ ({w} ∪ {v}) := by ext rw [← not_iff_not] simp [imp_iff_not_or, or_assoc, or_comm, or_left_comm] #align simple_graph.compl_neighbor_finset_sdiff_inter_eq SimpleGraph.compl_neighborFinset_sdiff_inter_eq theorem sdiff_compl_neighborFinset_inter_eq {v w : V} (h : G.Adj v w) : - (G.neighborFinset vᶜ ∩ G.neighborFinset wᶜ) \ ({w} ∪ {v}) = - G.neighborFinset vᶜ ∩ G.neighborFinset wᶜ := by + ((G.neighborFinset v)ᶜ ∩ (G.neighborFinset w)ᶜ) \ ({w} ∪ {v}) = + (G.neighborFinset v)ᶜ ∩ (G.neighborFinset w)ᶜ := by ext simp only [and_imp, mem_union, mem_sdiff, mem_compl, and_iff_left_iff_imp, mem_neighborFinset, mem_inter, mem_singleton] diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 8b5cf650acf82..a2b92498ec4a8 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -2193,7 +2193,7 @@ theorem succAbove_right_inj {x : Fin (n + 1)} : x.succAbove a = x.succAbove b /-- `succAbove` is injective at the pivot -/ theorem succAbove_left_injective : Injective (@succAbove n) := fun _ _ h => by - simpa [range_succAbove] using congr_arg (fun f : Fin n ↪o Fin (n + 1) => Set.range fᶜ) h + simpa [range_succAbove] using congr_arg (fun f : Fin n ↪o Fin (n + 1) => (Set.range f)ᶜ) h #align fin.succ_above_left_injective Fin.succAbove_left_injective /-- `succAbove` is injective at the pivot -/ diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 765f5dcafbf89..d606de49fb8c5 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -1182,7 +1182,7 @@ theorem insert_inj (ha : a ∉ s) : insert a s = insert b s ↔ a = b := ⟨fun h => eq_of_not_mem_of_mem_insert (h.subst <| mem_insert_self _ _) ha, congr_arg (insert · s)⟩ #align finset.insert_inj Finset.insert_inj -theorem insert_inj_on (s : Finset α) : Set.InjOn (fun a => insert a s) (sᶜ) := fun _ h _ _ => +theorem insert_inj_on (s : Finset α) : Set.InjOn (fun a => insert a s) sᶜ := fun _ h _ _ => (insert_inj h).1 #align finset.insert_inj_on Finset.insert_inj_on diff --git a/Mathlib/Data/Finset/Lattice.lean b/Mathlib/Data/Finset/Lattice.lean index 0889d500c5c13..c282cf4636a8e 100644 --- a/Mathlib/Data/Finset/Lattice.lean +++ b/Mathlib/Data/Finset/Lattice.lean @@ -651,12 +651,12 @@ theorem sup_himp_left (hs : s.Nonempty) (f : ι → α) (a : α) : #align finset.sup_himp_left Finset.sup_himp_left @[simp] -protected theorem compl_sup (s : Finset ι) (f : ι → α) : s.sup fᶜ = s.inf fun i => f iᶜ := +protected theorem compl_sup (s : Finset ι) (f : ι → α) : (s.sup f)ᶜ = s.inf fun i => (f i)ᶜ := map_finset_sup (OrderIso.compl α) _ _ #align finset.compl_sup Finset.compl_sup @[simp] -protected theorem compl_inf (s : Finset ι) (f : ι → α) : s.inf fᶜ = s.sup fun i => f iᶜ := +protected theorem compl_inf (s : Finset ι) (f : ι → α) : (s.inf f)ᶜ = s.sup fun i => (f i)ᶜ := map_finset_inf (OrderIso.compl α) _ _ #align finset.compl_inf Finset.compl_inf diff --git a/Mathlib/Data/Finset/Preimage.lean b/Mathlib/Data/Finset/Preimage.lean index 2916d0f4d38d3..7dcde84762c95 100644 --- a/Mathlib/Data/Finset/Preimage.lean +++ b/Mathlib/Data/Finset/Preimage.lean @@ -75,7 +75,7 @@ theorem preimage_union [DecidableEq α] [DecidableEq β] {f : α → β} {s t : @[simp, nolint simpNF] -- Porting note: linter complains that LHS doesn't simplify theorem preimage_compl [DecidableEq α] [DecidableEq β] [Fintype α] [Fintype β] {f : α → β} (s : Finset β) (hf : Function.Injective f) : - preimage (sᶜ) f (hf.injOn _) = preimage s f (hf.injOn _)ᶜ := + preimage sᶜ f (hf.injOn _) = (preimage s f (hf.injOn _))ᶜ := Finset.coe_injective (by simp) #align finset.preimage_compl Finset.preimage_compl diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 02bb11c8eaf31..102c11c2009eb 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -174,7 +174,7 @@ theorem not_mem_compl : a ∉ sᶜ ↔ a ∈ s := by rw [mem_compl, not_not] #align finset.not_mem_compl Finset.not_mem_compl @[simp, norm_cast] -theorem coe_compl (s : Finset α) : ↑(sᶜ) = (↑s : Set α)ᶜ := +theorem coe_compl (s : Finset α) : ↑sᶜ = (↑s : Set α)ᶜ := Set.ext fun _ => mem_compl #align finset.coe_compl Finset.coe_compl @@ -219,13 +219,13 @@ theorem compl_inter (s t : Finset α) : (s ∩ t)ᶜ = sᶜ ∪ tᶜ := #align finset.compl_inter Finset.compl_inter @[simp] -theorem compl_erase : s.erase aᶜ = insert a (sᶜ) := by +theorem compl_erase : (s.erase a)ᶜ = insert a sᶜ := by ext simp only [or_iff_not_imp_left, mem_insert, not_and, mem_compl, mem_erase] #align finset.compl_erase Finset.compl_erase @[simp] -theorem compl_insert : insert a sᶜ = sᶜ.erase a := by +theorem compl_insert : (insert a s)ᶜ = sᶜ.erase a := by ext simp only [not_or, mem_insert, iff_self_iff, mem_compl, mem_erase] #align finset.compl_insert Finset.compl_insert @@ -237,7 +237,7 @@ theorem insert_compl_self (x : α) : insert x ({x}ᶜ : Finset α) = univ := by @[simp] theorem compl_filter (p : α → Prop) [DecidablePred p] [∀ x, Decidable ¬p x] : - univ.filter pᶜ = univ.filter fun x => ¬p x := + (univ.filter p)ᶜ = univ.filter fun x => ¬p x := ext <| by simp #align finset.compl_filter Finset.compl_filter diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index cc3e433be9fd2..0ea6a6f4bf160 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -292,8 +292,8 @@ theorem Finset.card_compl [DecidableEq α] [Fintype α] (s : Finset α) : Finset.card_univ_diff s #align finset.card_compl Finset.card_compl -theorem Fintype.card_compl_set [Fintype α] (s : Set α) [Fintype s] [Fintype (↥(sᶜ) : Sort _)] : - Fintype.card (↥(sᶜ) : Sort _) = Fintype.card α - Fintype.card s := by +theorem Fintype.card_compl_set [Fintype α] (s : Set α) [Fintype s] [Fintype (↥sᶜ : Sort _)] : + Fintype.card (↥sᶜ : Sort _) = Fintype.card α - Fintype.card s := by classical rw [← Set.toFinset_card, ← Set.toFinset_card, ← Finset.card_compl, Set.toFinset_compl] #align fintype.card_compl_set Fintype.card_compl_set @@ -850,7 +850,7 @@ theorem Fintype.card_subtype_compl [Fintype α] (p : α → Prop) [Fintype { x / [Fintype { x // ¬p x }] : Fintype.card { x // ¬p x } = Fintype.card α - Fintype.card { x // p x } := by classical - rw [Fintype.card_of_subtype (Set.toFinset (pᶜ)), Set.toFinset_compl p, + rw [Fintype.card_of_subtype (Set.toFinset pᶜ), Set.toFinset_compl p, Finset.card_compl, Fintype.card_of_subtype (Set.toFinset p)] <;> · intro simp only [Set.mem_toFinset, Set.mem_compl_iff] diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 67d93cf93a5d2..845c4261696e9 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -1748,11 +1748,11 @@ theorem subset_compl_iff_disjoint_right : s ⊆ tᶜ ↔ Disjoint s t := @le_compl_iff_disjoint_right (Set α) _ _ _ #align set.subset_compl_iff_disjoint_right Set.subset_compl_iff_disjoint_right -theorem disjoint_compl_left_iff_subset : Disjoint (sᶜ) t ↔ t ⊆ s := +theorem disjoint_compl_left_iff_subset : Disjoint sᶜ t ↔ t ⊆ s := disjoint_compl_left_iff #align set.disjoint_compl_left_iff_subset Set.disjoint_compl_left_iff_subset -theorem disjoint_compl_right_iff_subset : Disjoint s (tᶜ) ↔ s ⊆ t := +theorem disjoint_compl_right_iff_subset : Disjoint s tᶜ ↔ s ⊆ t := disjoint_compl_right_iff #align set.disjoint_compl_right_iff_subset Set.disjoint_compl_right_iff_subset diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index e14f56165dc1a..364cda7455a5e 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -125,13 +125,13 @@ theorem restrict_extend_range (f : α → β) (g : α → γ) (g' : β → γ) : @[simp] theorem restrict_extend_compl_range (f : α → β) (g : α → γ) (g' : β → γ) : - range fᶜ.restrict (extend f g g') = g' ∘ Subtype.val := by + (range f)ᶜ.restrict (extend f g g') = g' ∘ Subtype.val := by classical exact restrict_dite_compl _ _ #align set.restrict_extend_compl_range Set.restrict_extend_compl_range theorem range_extend_subset (f : α → β) (g : α → γ) (g' : β → γ) : - range (extend f g g') ⊆ range g ∪ g' '' range fᶜ := by + range (extend f g g') ⊆ range g ∪ g' '' (range f)ᶜ := by classical rintro _ ⟨y, rfl⟩ rw [extend_def] @@ -140,7 +140,7 @@ theorem range_extend_subset (f : α → β) (g : α → γ) (g' : β → γ) : #align set.range_extend_subset Set.range_extend_subset theorem range_extend {f : α → β} (hf : Injective f) (g : α → γ) (g' : β → γ) : - range (extend f g g') = range g ∪ g' '' range fᶜ := by + range (extend f g g') = range g ∪ g' '' (range f)ᶜ := by refine' (range_extend_subset _ _ _).antisymm _ rintro z (⟨x, rfl⟩ | ⟨y, hy, rfl⟩) exacts [⟨f x, hf.extend_apply _ _ _⟩, ⟨y, extend_apply' _ _ _ hy⟩] @@ -538,7 +538,7 @@ theorem surjective_mapsTo_image_restrict (f : α → β) (s : Set α) : ⟨⟨x, hs⟩, Subtype.ext hxy⟩ #align set.surjective_maps_to_image_restrict Set.surjective_mapsTo_image_restrict -theorem MapsTo.mem_iff (h : MapsTo f s t) (hc : MapsTo f (sᶜ) (tᶜ)) {x} : f x ∈ t ↔ x ∈ s := +theorem MapsTo.mem_iff (h : MapsTo f s t) (hc : MapsTo f sᶜ tᶜ) {x} : f x ∈ t ↔ x ∈ s := ⟨fun ht => by_contra fun hs => hc hs ht, fun hx => h hx⟩ #align set.maps_to.mem_iff Set.MapsTo.mem_iff @@ -871,13 +871,13 @@ theorem image_eq_iff_surjOn_mapsTo : f '' s = t ↔ s.SurjOn f t ∧ s.MapsTo f exact ⟨s.surjOn_image f, s.mapsTo_image f⟩ #align set.image_eq_iff_surj_on_maps_to Set.image_eq_iff_surjOn_mapsTo -theorem SurjOn.mapsTo_compl (h : SurjOn f s t) (h' : Injective f) : MapsTo f (sᶜ) (tᶜ) := +theorem SurjOn.mapsTo_compl (h : SurjOn f s t) (h' : Injective f) : MapsTo f sᶜ tᶜ := fun _ hs ht => let ⟨_, hx', HEq⟩ := h ht hs <| h' HEq ▸ hx' #align set.surj_on.maps_to_compl Set.SurjOn.mapsTo_compl -theorem MapsTo.surjOn_compl (h : MapsTo f s t) (h' : Surjective f) : SurjOn f (sᶜ) (tᶜ) := +theorem MapsTo.surjOn_compl (h : MapsTo f s t) (h' : Surjective f) : SurjOn f sᶜ tᶜ := h'.forall.2 fun _ ht => (mem_image_of_mem _) fun hs => ht (h hs) #align set.maps_to.surj_on_compl Set.MapsTo.surjOn_compl @@ -1010,7 +1010,7 @@ theorem bijective_iff_bijOn_univ : Bijective f ↔ BijOn f univ univ := alias bijective_iff_bijOn_univ ↔ _root_.Function.Bijective.bijOn_univ _ #align function.bijective.bij_on_univ Function.Bijective.bijOn_univ -theorem BijOn.compl (hst : BijOn f s t) (hf : Bijective f) : BijOn f (sᶜ) (tᶜ) := +theorem BijOn.compl (hst : BijOn f s t) (hf : Bijective f) : BijOn f sᶜ tᶜ := ⟨hst.surjOn.mapsTo_compl hf.1, hf.1.injOn _, hst.mapsTo.surjOn_compl hf.2⟩ #align set.bij_on.compl Set.BijOn.compl @@ -1290,7 +1290,7 @@ theorem surjOn_iff_exists_bijOn_subset : SurjOn f s t ↔ ∃ (s' : _) (_ : s' #align set.surj_on_iff_exists_bij_on_subset Set.surjOn_iff_exists_bijOn_subset theorem preimage_invFun_of_mem [n : Nonempty α] {f : α → β} (hf : Injective f) {s : Set α} - (h : Classical.choice n ∈ s) : invFun f ⁻¹' s = f '' s ∪ range fᶜ := by + (h : Classical.choice n ∈ s) : invFun f ⁻¹' s = f '' s ∪ (range f)ᶜ := by ext x rcases em (x ∈ range f) with (⟨a, rfl⟩ | hx) · simp only [mem_preimage, mem_union, mem_compl_iff, mem_range_self, not_true, or_false, @@ -1403,7 +1403,7 @@ theorem piecewise_eqOn (f g : α → β) : EqOn (s.piecewise f g) f s := fun _ = piecewise_eq_of_mem _ _ _ #align set.piecewise_eq_on Set.piecewise_eqOn -theorem piecewise_eqOn_compl (f g : α → β) : EqOn (s.piecewise f g) g (sᶜ) := fun _ => +theorem piecewise_eqOn_compl (f g : α → β) : EqOn (s.piecewise f g) g sᶜ := fun _ => piecewise_eq_of_not_mem _ _ _ #align set.piecewise_eq_on_compl Set.piecewise_eqOn_compl @@ -1508,7 +1508,7 @@ theorem range_piecewise (f g : α → β) : range (s.piecewise f g) = f '' s ∪ theorem injective_piecewise_iff {f g : α → β} : Injective (s.piecewise f g) ↔ - InjOn f s ∧ InjOn g (sᶜ) ∧ ∀ x ∈ s, ∀ (y) (_ : y ∉ s), f x ≠ g y := by + InjOn f s ∧ InjOn g sᶜ ∧ ∀ x ∈ s, ∀ (y) (_ : y ∉ s), f x ≠ g y := by rw [injective_iff_injOn_univ, ← union_compl_self s, injOn_union (@disjoint_compl_right _ _ s), (piecewise_eqOn s f g).injOn_iff, (piecewise_eqOn_compl s f g).injOn_iff] refine' and_congr Iff.rfl (and_congr Iff.rfl <| forall₄_congr fun x hx y hy => _) @@ -1529,7 +1529,7 @@ theorem pi_piecewise {ι : Type _} {α : ι → Type _} (s s' : Set ι) (t t' : -- porting note: new lemma theorem univ_pi_piecewise {ι : Type _} {α : ι → Type _} (s : Set ι) (t t' : ∀ i, Set (α i)) - [∀ x, Decidable (x ∈ s)] : pi univ (s.piecewise t t') = pi s t ∩ pi (sᶜ) t' := by + [∀ x, Decidable (x ∈ s)] : pi univ (s.piecewise t t') = pi s t ∩ pi sᶜ t' := by simp [compl_eq_univ_diff] theorem univ_pi_piecewise_univ {ι : Type _} {α : ι → Type _} (s : Set ι) (t : ∀ i, Set (α i)) diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index ba6d75ce2892c..3192bd655a62f 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -530,7 +530,7 @@ theorem compl_image : image (compl : Set α → Set α) = preimage compl := image_eq_preimage_of_inverse compl_compl compl_compl #align set.compl_image Set.compl_image -theorem compl_image_set_of {p : Set α → Prop} : compl '' { s | p s } = { s | p (sᶜ) } := +theorem compl_image_set_of {p : Set α → Prop} : compl '' { s | p s } = { s | p sᶜ } := congr_fun compl_image p #align set.compl_image_set_of Set.compl_image_set_of @@ -937,12 +937,12 @@ theorem preimage_inr_range_inl : Sum.inr ⁻¹' range (Sum.inl : α → Sum α #align set.preimage_inr_range_inl Set.preimage_inr_range_inl @[simp] -theorem compl_range_inl : range (Sum.inl : α → Sum α β)ᶜ = range (Sum.inr : β → Sum α β) := +theorem compl_range_inl : (range (Sum.inl : α → Sum α β))ᶜ = range (Sum.inr : β → Sum α β) := IsCompl.compl_eq isCompl_range_inl_range_inr #align set.compl_range_inl Set.compl_range_inl @[simp] -theorem compl_range_inr : range (Sum.inr : β → Sum α β)ᶜ = range (Sum.inl : α → Sum α β) := +theorem compl_range_inr : (range (Sum.inr : β → Sum α β))ᶜ = range (Sum.inl : α → Sum α β) := IsCompl.compl_eq isCompl_range_inl_range_inr.symm #align set.compl_range_inr Set.compl_range_inr @@ -1188,7 +1188,7 @@ theorem isCompl_range_some_none (α : Type _) : IsCompl (range (some : α → Op #align set.is_compl_range_some_none Set.isCompl_range_some_none @[simp] -theorem compl_range_some (α : Type _) : range (some : α → Option α)ᶜ = {none} := +theorem compl_range_some (α : Type _) : (range (some : α → Option α))ᶜ = {none} := (isCompl_range_some_none α).compl_eq #align set.compl_range_some Set.compl_range_some @@ -1340,7 +1340,7 @@ theorem Injective.exists_unique_of_mem_range (hf : Injective f) {b : β} (hb : b #align function.injective.exists_unique_of_mem_range Function.Injective.exists_unique_of_mem_range theorem Injective.compl_image_eq (hf : Injective f) (s : Set α) : - (f '' s)ᶜ = f '' sᶜ ∪ range fᶜ := by + (f '' s)ᶜ = f '' sᶜ ∪ (range f)ᶜ := by ext y rcases em (y ∈ range f) with (⟨x, rfl⟩ | hx) · simp [hf.eq_iff] diff --git a/Mathlib/Data/Set/Intervals/Basic.lean b/Mathlib/Data/Set/Intervals/Basic.lean index 75e07d3a43c80..0d77fdb923168 100644 --- a/Mathlib/Data/Set/Intervals/Basic.lean +++ b/Mathlib/Data/Set/Intervals/Basic.lean @@ -1057,22 +1057,22 @@ theorem not_mem_Iio : c ∉ Iio b ↔ b ≤ c := #align set.not_mem_Iio Set.not_mem_Iio @[simp] -theorem compl_Iic : Iic aᶜ = Ioi a := +theorem compl_Iic : (Iic a)ᶜ = Ioi a := ext fun _ => not_le #align set.compl_Iic Set.compl_Iic @[simp] -theorem compl_Ici : Ici aᶜ = Iio a := +theorem compl_Ici : (Ici a)ᶜ = Iio a := ext fun _ => not_le #align set.compl_Ici Set.compl_Ici @[simp] -theorem compl_Iio : Iio aᶜ = Ici a := +theorem compl_Iio : (Iio a)ᶜ = Ici a := ext fun _ => not_lt #align set.compl_Iio Set.compl_Iio @[simp] -theorem compl_Ioi : Ioi aᶜ = Iic a := +theorem compl_Ioi : (Ioi a)ᶜ = Iic a := ext fun _ => not_lt #align set.compl_Ioi Set.compl_Ioi diff --git a/Mathlib/Data/Set/Intervals/OrdConnectedComponent.lean b/Mathlib/Data/Set/Intervals/OrdConnectedComponent.lean index ac435e387fb0a..27b6404bdb2f0 100644 --- a/Mathlib/Data/Set/Intervals/OrdConnectedComponent.lean +++ b/Mathlib/Data/Set/Intervals/OrdConnectedComponent.lean @@ -168,7 +168,7 @@ belong both to some `Set.ordConnectedComponent tᶜ x`, `x ∈ s`, and to some `Set.ordConnectedComponent sᶜ x`, `x ∈ t`. In the case of two disjoint closed sets, this is the union of all open intervals $(a, b)$ such that their endpoints belong to different sets. -/ def ordSeparatingSet (s t : Set α) : Set α := - (⋃ x ∈ s, ordConnectedComponent (tᶜ) x) ∩ ⋃ x ∈ t, ordConnectedComponent (sᶜ) x + (⋃ x ∈ s, ordConnectedComponent tᶜ x) ∩ ⋃ x ∈ t, ordConnectedComponent sᶜ x #align set.ord_separating_set Set.ordSeparatingSet theorem ordSeparatingSet_comm (s t : Set α) : ordSeparatingSet s t = ordSeparatingSet t s := diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index cf952e7703c59..1671220548a9a 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -487,34 +487,34 @@ lemma iInter_eq_const (hf : ∀ i, f i = s) : (⋂ i, f i) = s := end Nonempty @[simp] -theorem compl_iUnion (s : ι → Set β) : (⋃ i, s i)ᶜ = ⋂ i, s iᶜ := +theorem compl_iUnion (s : ι → Set β) : (⋃ i, s i)ᶜ = ⋂ i, (s i)ᶜ := compl_iSup #align set.compl_Union Set.compl_iUnion /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ -theorem compl_iUnion₂ (s : ∀ i, κ i → Set α) : (⋃ (i) (j), s i j)ᶜ = ⋂ (i) (j), s i jᶜ := by +theorem compl_iUnion₂ (s : ∀ i, κ i → Set α) : (⋃ (i) (j), s i j)ᶜ = ⋂ (i) (j), (s i j)ᶜ := by simp_rw [compl_iUnion] #align set.compl_Union₂ Set.compl_iUnion₂ @[simp] -theorem compl_iInter (s : ι → Set β) : (⋂ i, s i)ᶜ = ⋃ i, s iᶜ := +theorem compl_iInter (s : ι → Set β) : (⋂ i, s i)ᶜ = ⋃ i, (s i)ᶜ := compl_iInf #align set.compl_Inter Set.compl_iInter /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ -theorem compl_iInter₂ (s : ∀ i, κ i → Set α) : (⋂ (i) (j), s i j)ᶜ = ⋃ (i) (j), s i jᶜ := by +theorem compl_iInter₂ (s : ∀ i, κ i → Set α) : (⋂ (i) (j), s i j)ᶜ = ⋃ (i) (j), (s i j)ᶜ := by simp_rw [compl_iInter] #align set.compl_Inter₂ Set.compl_iInter₂ -- classical -- complete_boolean_algebra -theorem iUnion_eq_compl_iInter_compl (s : ι → Set β) : (⋃ i, s i) = (⋂ i, s iᶜ)ᶜ := by +theorem iUnion_eq_compl_iInter_compl (s : ι → Set β) : (⋃ i, s i) = (⋂ i, (s i)ᶜ)ᶜ := by simp only [compl_iInter, compl_compl] #align set.Union_eq_compl_Inter_compl Set.iUnion_eq_compl_iInter_compl -- classical -- complete_boolean_algebra -theorem iInter_eq_compl_iUnion_compl (s : ι → Set β) : (⋂ i, s i) = (⋃ i, s iᶜ)ᶜ := by +theorem iInter_eq_compl_iUnion_compl (s : ι → Set β) : (⋂ i, s i) = (⋃ i, (s i)ᶜ)ᶜ := by simp only [compl_iUnion, compl_compl] #align set.Inter_eq_compl_Union_compl Set.iInter_eq_compl_iUnion_compl diff --git a/Mathlib/Data/Set/Pointwise/Basic.lean b/Mathlib/Data/Set/Pointwise/Basic.lean index e45ed11713e90..1c5a759058663 100644 --- a/Mathlib/Data/Set/Pointwise/Basic.lean +++ b/Mathlib/Data/Set/Pointwise/Basic.lean @@ -226,7 +226,7 @@ theorem iUnion_inv (s : ι → Set α) : (⋃ i, s i)⁻¹ = ⋃ i, (s i)⁻¹ : #align set.Union_neg Set.iUnion_neg @[to_additive (attr := simp)] -theorem compl_inv : (sᶜ)⁻¹ = s⁻¹ᶜ := +theorem compl_inv : sᶜ⁻¹ = s⁻¹ᶜ := preimage_compl #align set.compl_inv Set.compl_inv #align set.compl_neg Set.compl_neg diff --git a/Mathlib/Data/Set/Prod.lean b/Mathlib/Data/Set/Prod.lean index ca0c9e66aa629..f4a32eb14097d 100644 --- a/Mathlib/Data/Set/Prod.lean +++ b/Mathlib/Data/Set/Prod.lean @@ -513,7 +513,7 @@ theorem diagonal_subset_iff {s} : diagonal α ⊆ s ↔ ∀ x, (x, x) ∈ s := b #align set.diagonal_subset_iff Set.diagonal_subset_iff @[simp] -theorem prod_subset_compl_diagonal_iff_disjoint : s ×ˢ t ⊆ diagonal αᶜ ↔ Disjoint s t := +theorem prod_subset_compl_diagonal_iff_disjoint : s ×ˢ t ⊆ (diagonal α)ᶜ ↔ Disjoint s t := prod_subset_iff.trans disjoint_iff_forall_ne.symm #align set.prod_subset_compl_diagonal_iff_disjoint Set.prod_subset_compl_diagonal_iff_disjoint @@ -590,7 +590,7 @@ theorem offDiag_singleton (a : α) : ({a} : Set α).offDiag = ∅ := by simp #align set.off_diag_singleton Set.offDiag_singleton @[simp] -theorem offDiag_univ : (univ : Set α).offDiag = diagonal αᶜ := +theorem offDiag_univ : (univ : Set α).offDiag = (diagonal α)ᶜ := ext <| by simp #align set.off_diag_univ Set.offDiag_univ @@ -792,7 +792,7 @@ theorem union_pi : (s₁ ∪ s₂).pi t = s₁.pi t ∩ s₂.pi t := by #align set.union_pi Set.union_pi @[simp] -theorem pi_inter_compl (s : Set ι) : pi s t ∩ pi (sᶜ) t = pi univ t := by +theorem pi_inter_compl (s : Set ι) : pi s t ∩ pi sᶜ t = pi univ t := by rw [← union_pi, union_compl_self] #align set.pi_inter_compl Set.pi_inter_compl @@ -815,7 +815,7 @@ theorem pi_update_of_mem [DecidableEq ι] (hi : i ∈ s) (f : ∀ j, α j) (a : theorem univ_pi_update [DecidableEq ι] {β : ∀ _, Type _} (i : ι) (f : ∀ j, α j) (a : α i) (t : ∀ j, α j → Set (β j)) : - (pi univ fun j => t j (update f i a j)) = { x | x i ∈ t i a } ∩ pi ({i}ᶜ) fun j => t j (f j) := + (pi univ fun j => t j (update f i a j)) = { x | x i ∈ t i a } ∩ pi {i}ᶜ fun j => t j (f j) := by rw [compl_eq_univ_diff, ← pi_update_of_mem (mem_univ _)] #align set.univ_pi_update Set.univ_pi_update diff --git a/Mathlib/Dynamics/Ergodic/Ergodic.lean b/Mathlib/Dynamics/Ergodic/Ergodic.lean index b8faf71e653c8..8056a5a02d82a 100644 --- a/Mathlib/Dynamics/Ergodic/Ergodic.lean +++ b/Mathlib/Dynamics/Ergodic/Ergodic.lean @@ -65,7 +65,7 @@ variable {f} {μ : Measure α} namespace PreErgodic theorem measure_self_or_compl_eq_zero (hf : PreErgodic f μ) (hs : MeasurableSet s) - (hs' : f ⁻¹' s = s) : μ s = 0 ∨ μ (sᶜ) = 0 := by + (hs' : f ⁻¹' s = s) : μ s = 0 ∨ μ sᶜ = 0 := by simpa using hf.ae_empty_or_univ hs hs' #align pre_ergodic.measure_self_or_compl_eq_zero PreErgodic.measure_self_or_compl_eq_zero diff --git a/Mathlib/Dynamics/OmegaLimit.lean b/Mathlib/Dynamics/OmegaLimit.lean index cb2fe4b5143d8..4541282227e16 100644 --- a/Mathlib/Dynamics/OmegaLimit.lean +++ b/Mathlib/Dynamics/OmegaLimit.lean @@ -241,7 +241,7 @@ theorem eventually_closure_subset_of_isCompact_absorbing_of_isOpen_of_omegaLimit let k := closure (image2 ϕ v s) have hk : IsCompact (k \ n) := IsCompact.diff (isCompact_of_isClosed_subset hc₁ isClosed_closure hv₂) hn₁ - let j u := closure (image2 ϕ (u ∩ v) s)ᶜ + let j u := (closure (image2 ϕ (u ∩ v) s))ᶜ have hj₁ : ∀ u ∈ f, IsOpen (j u) := fun _ _ ↦ isOpen_compl_iff.mpr isClosed_closure have hj₂ : k \ n ⊆ ⋃ u ∈ f, j u := by have : (⋃ u ∈ f, j u) = ⋃ u : (↥f.sets), j u := biUnion_eq_iUnion _ _ @@ -253,14 +253,14 @@ theorem eventually_closure_subset_of_isCompact_absorbing_of_isOpen_of_omegaLimit rcases hk.elim_finite_subcover_image hj₁ hj₂ with ⟨g, hg₁ : ∀ u ∈ g, u ∈ f, hg₂, hg₃⟩ let w := (⋂ u ∈ g, u) ∩ v have hw₂ : w ∈ f := by simpa [*] - have hw₃ : k \ n ⊆ closure (image2 ϕ w s)ᶜ := by + have hw₃ : k \ n ⊆ (closure (image2 ϕ w s))ᶜ := by apply Subset.trans hg₃ simp only [iUnion_subset_iff, compl_subset_compl] intros u hu mono refine' iInter_subset_of_subset u (iInter_subset_of_subset hu _) all_goals exact Subset.rfl - have hw₄ : kᶜ ⊆ closure (image2 ϕ w s)ᶜ := by + have hw₄ : kᶜ ⊆ (closure (image2 ϕ w s))ᶜ := by simp only [compl_subset_compl] exact closure_mono (image2_subset (inter_subset_right _ _) Subset.rfl) have hnc : nᶜ ⊆ k \ n ∪ kᶜ := by rw [union_comm, ← inter_subset, diff_eq, inter_comm] diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index a06a4f52a48aa..9dfa8aa94bad7 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -174,7 +174,7 @@ result of subtracting `centroidWeightsWithCircumcenter` from theorem mongePointVSubFaceCentroidWeightsWithCircumcenter_eq_sub {n : ℕ} {i₁ i₂ : Fin (n + 3)} (h : i₁ ≠ i₂) : mongePointVSubFaceCentroidWeightsWithCircumcenter i₁ i₂ = - mongePointWeightsWithCircumcenter n - centroidWeightsWithCircumcenter ({i₁, i₂}ᶜ) := by + mongePointWeightsWithCircumcenter n - centroidWeightsWithCircumcenter {i₁, i₂}ᶜ := by ext i cases' i with i · rw [Pi.sub_apply, mongePointWeightsWithCircumcenter, centroidWeightsWithCircumcenter, diff --git a/Mathlib/GroupTheory/FreeProduct.lean b/Mathlib/GroupTheory/FreeProduct.lean index c2e36f91e16b3..57402745a8b87 100644 --- a/Mathlib/GroupTheory/FreeProduct.lean +++ b/Mathlib/GroupTheory/FreeProduct.lean @@ -884,9 +884,9 @@ variable (hYdisj : Pairwise fun i j => Disjoint (Y i) (Y j)) variable (hXYdisj : ∀ i j, Disjoint (X i) (Y j)) -variable (hX : ∀ i, a i • Y iᶜ ⊆ X i) +variable (hX : ∀ i, a i • (Y i)ᶜ ⊆ X i) -variable (hY : ∀ i, a⁻¹ i • X iᶜ ⊆ Y i) +variable (hY : ∀ i, a⁻¹ i • (X i)ᶜ ⊆ Y i) --include hXnonempty hXdisj hYdisj hXYdisj hX hY Porting note: commented out @@ -945,41 +945,41 @@ theorem _root_.FreeGroup.injective_lift_of_ping_pong : Function.Injective (FreeG cases' (lt_or_gt_of_ne hnne0).symm with hlt hgt · have h1n : 1 ≤ n := hlt calc - a i ^ n • X' j ⊆ a i ^ n • Y iᶜ := + a i ^ n • X' j ⊆ a i ^ n • (Y i)ᶜ := smul_set_mono ((hXYdisj j i).union_left <| hYdisj hij.symm).subset_compl_right _ ⊆ X i := by clear hnne0 hlt - refine Int.le_induction (P := fun n => a i ^ n • Y iᶜ ⊆ X i) ?_ ?_ n h1n + refine Int.le_induction (P := fun n => a i ^ n • (Y i)ᶜ ⊆ X i) ?_ ?_ n h1n . dsimp rw [zpow_one] exact hX i . dsimp intro n _hle hi calc - a i ^ (n + 1) • Y iᶜ = (a i ^ n * a i) • Y iᶜ := by rw [zpow_add, zpow_one] - _ = a i ^ n • a i • Y iᶜ := (MulAction.mul_smul _ _ _) + a i ^ (n + 1) • (Y i)ᶜ = (a i ^ n * a i) • (Y i)ᶜ := by rw [zpow_add, zpow_one] + _ = a i ^ n • a i • (Y i)ᶜ := (MulAction.mul_smul _ _ _) _ ⊆ a i ^ n • X i := (smul_set_mono <| hX i) - _ ⊆ a i ^ n • Y iᶜ := (smul_set_mono (hXYdisj i i).subset_compl_right) + _ ⊆ a i ^ n • (Y i)ᶜ := (smul_set_mono (hXYdisj i i).subset_compl_right) _ ⊆ X i := hi _ ⊆ X' i := Set.subset_union_left _ _ · have h1n : n ≤ -1 := by apply Int.le_of_lt_add_one simpa using hgt calc - a i ^ n • X' j ⊆ a i ^ n • X iᶜ := + a i ^ n • X' j ⊆ a i ^ n • (X i)ᶜ := smul_set_mono ((hXdisj hij.symm).union_left (hXYdisj i j).symm).subset_compl_right _ ⊆ Y i := by - refine' Int.le_induction_down (P := fun n => a i ^ n • X iᶜ ⊆ Y i) _ _ _ h1n + refine' Int.le_induction_down (P := fun n => a i ^ n • (X i)ᶜ ⊆ Y i) _ _ _ h1n · dsimp rw [zpow_neg, zpow_one] exact hY i · dsimp intro n _ hi calc - a i ^ (n - 1) • X iᶜ = (a i ^ n * (a i)⁻¹) • X iᶜ := by rw [zpow_sub, zpow_one] - _ = a i ^ n • (a i)⁻¹ • X iᶜ := (MulAction.mul_smul _ _ _) + a i ^ (n - 1) • (X i)ᶜ = (a i ^ n * (a i)⁻¹) • (X i)ᶜ := by rw [zpow_sub, zpow_one] + _ = a i ^ n • (a i)⁻¹ • (X i)ᶜ := (MulAction.mul_smul _ _ _) _ ⊆ a i ^ n • Y i := (smul_set_mono <| hY i) - _ ⊆ a i ^ n • X iᶜ := (smul_set_mono (hXYdisj i i).symm.subset_compl_right) + _ ⊆ a i ^ n • (X i)ᶜ := (smul_set_mono (hXYdisj i i).symm.subset_compl_right) _ ⊆ Y i := hi _ ⊆ X' i := Set.subset_union_right _ _ show _ ∨ ∃ i, 3 ≤ (#H i) diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index 368a3cf555278..b7a92841bf40f 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -160,8 +160,8 @@ theorem disjoint_lsingle_lsingle (s t : Set α) (hs : Disjoint s t) : -- Porting note: 2 placeholders are added to prevent timeout. refine' (Disjoint.mono - (lsingle_range_le_ker_lapply s (sᶜ) _) - (lsingle_range_le_ker_lapply t (tᶜ) _)) + (lsingle_range_le_ker_lapply s sᶜ _) + (lsingle_range_le_ker_lapply t tᶜ _)) _ · apply disjoint_compl_right · apply disjoint_compl_right diff --git a/Mathlib/Logic/Embedding/Set.lean b/Mathlib/Logic/Embedding/Set.lean index d447f91b0cdca..5a279d4626468 100644 --- a/Mathlib/Logic/Embedding/Set.lean +++ b/Mathlib/Logic/Embedding/Set.lean @@ -52,7 +52,7 @@ def optionElim {α β} (f : α ↪ β) (x : β) (h : x ∉ Set.range f) : Option /-- Equivalence between embeddings of `Option α` and a sigma type over the embeddings of `α`. -/ @[simps] -def optionEmbeddingEquiv (α β) : (Option α ↪ β) ≃ Σ f : α ↪ β, ↥(Set.range fᶜ) where +def optionEmbeddingEquiv (α β) : (Option α ↪ β) ≃ Σ f : α ↪ β, ↥(Set.range f)ᶜ where toFun f := ⟨coeWithTop.trans f, f none, fun ⟨x, hx⟩ ↦ Option.some_ne_none x <| f.injective hx⟩ invFun f := f.1.optionElim f.2 f.2.2 left_inv f := ext <| by rintro (_ | _) <;> simp [Option.coe_def] ; rfl diff --git a/Mathlib/Logic/Equiv/Embedding.lean b/Mathlib/Logic/Equiv/Embedding.lean index 4175a7faac4b4..d1b9da8910033 100644 --- a/Mathlib/Logic/Equiv/Embedding.lean +++ b/Mathlib/Logic/Equiv/Embedding.lean @@ -71,7 +71,7 @@ def codRestrict (α : Type _) {β : Type _} (bs : Set β) : in which the second embedding cannot take values in the range of the first. -/ def prodEmbeddingDisjointEquivSigmaEmbeddingRestricted {α β γ : Type _} : { f : (α ↪ γ) × (β ↪ γ) // Disjoint (Set.range f.1) (Set.range f.2) } ≃ - Σf : α ↪ γ, β ↪ ↥(Set.range fᶜ) := + Σf : α ↪ γ, β ↪ ↥(Set.range f)ᶜ := (subtypeProdEquivSigmaSubtype fun (a : α ↪ γ) (b : β ↪ _) => Disjoint (Set.range a) (Set.range b)).trans <| Equiv.sigmaCongrRight fun a => @@ -85,7 +85,7 @@ def prodEmbeddingDisjointEquivSigmaEmbeddingRestricted {α β γ : Type _} : into two dependent embeddings, the second of which avoids any members of the range of the first. This is helpful for constructing larger embeddings out of smaller ones. -/ def sumEmbeddingEquivSigmaEmbeddingRestricted {α β γ : Type _} : - (Sum α β ↪ γ) ≃ Σf : α ↪ γ, β ↪ ↥(Set.range fᶜ) := + (Sum α β ↪ γ) ≃ Σf : α ↪ γ, β ↪ ↥(Set.range f)ᶜ := Equiv.trans sumEmbeddingEquivProdEmbeddingDisjoint prodEmbeddingDisjointEquivSigmaEmbeddingRestricted #align equiv.sum_embedding_equiv_sigma_embedding_restricted Equiv.sumEmbeddingEquivSigmaEmbeddingRestricted diff --git a/Mathlib/Logic/Equiv/LocalEquiv.lean b/Mathlib/Logic/Equiv/LocalEquiv.lean index 3fb04a7de6d25..7ed0f9c4ccbf7 100644 --- a/Mathlib/Logic/Equiv/LocalEquiv.lean +++ b/Mathlib/Logic/Equiv/LocalEquiv.lean @@ -433,7 +433,7 @@ theorem of_symm_image_eq (h : e.symm '' (e.target ∩ t) = e.source ∩ s) : e.I of_preimage_eq <| Eq.trans (iff_preimage_eq.2 rfl).symm_image_eq.symm h #align local_equiv.is_image.of_symm_image_eq LocalEquiv.IsImage.of_symm_image_eq -protected theorem compl (h : e.IsImage s t) : e.IsImage (sᶜ) (tᶜ) := fun _ hx => not_congr (h hx) +protected theorem compl (h : e.IsImage s t) : e.IsImage sᶜ tᶜ := fun _ hx => not_congr (h hx) #align local_equiv.is_image.compl LocalEquiv.IsImage.compl protected theorem inter {s' t'} (h : e.IsImage s t) (h' : e.IsImage s' t') : diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index fdb3e02efbafd..31b250cac7734 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -134,13 +134,13 @@ theorem borel_eq_generateFrom_Iio : borel α = .generateFrom (range Iio) := by rintro _ ⟨a, rfl | rfl⟩ <;> [skip; apply H] by_cases h : ∃ a', ∀ b, a < b ↔ a' ≤ b · rcases h with ⟨a', ha'⟩ - rw [(_ : Ioi a = Iio a'ᶜ)] + rw [(_ : Ioi a = (Iio a')ᶜ)] · exact (H _).compl simp [Set.ext_iff, ha'] · rcases isOpen_iUnion_countable (fun a' : { a' : α // a < a' } => { b | a'.1 < b }) fun a' => isOpen_lt' _ with ⟨v, ⟨hv⟩, vu⟩ simp [Set.ext_iff] at vu - have : Ioi a = ⋃ x : v, Iio x.1.1ᶜ := by + have : Ioi a = ⋃ x : v, (Iio x.1.1)ᶜ := by simp [Set.ext_iff] refine' fun x => ⟨fun ax => _, fun ⟨a', ⟨h, _⟩, ax⟩ => lt_of_lt_of_le h ax⟩ rcases (vu x).2 (by @@ -293,7 +293,7 @@ instance (priority := 100) BorelSpace.countablyGenerated {α : Type _} [Topologi theorem MeasurableSet.induction_on_open [TopologicalSpace α] [MeasurableSpace α] [BorelSpace α] {C : Set α → Prop} (h_open : ∀ U, IsOpen U → C U) - (h_compl : ∀ t, MeasurableSet t → C t → C (tᶜ)) + (h_compl : ∀ t, MeasurableSet t → C t → C tᶜ) (h_union : ∀ f : ℕ → Set α, Pairwise (Disjoint on f) → (∀ i, MeasurableSet (f i)) → (∀ i, C (f i)) → C (⋃ i, f i)) : @@ -846,7 +846,7 @@ theorem Continuous.isOpenPosMeasure_map {f : β → γ} (hf : Continuous f) /-- If a function is defined piecewise in terms of functions which are continuous on their respective pieces, then it is measurable. -/ theorem ContinuousOn.measurable_piecewise {f g : α → γ} {s : Set α} [∀ j : α, Decidable (j ∈ s)] - (hf : ContinuousOn f s) (hg : ContinuousOn g (sᶜ)) (hs : MeasurableSet s) : + (hf : ContinuousOn f s) (hg : ContinuousOn g sᶜ) (hs : MeasurableSet s) : Measurable (s.piecewise f g) := by refine' measurable_of_isOpen fun t ht => _ rw [piecewise_preimage, Set.ite] @@ -946,7 +946,7 @@ theorem ContinuousMap.measurable (f : C(α, γ)) : Measurable f := #align continuous_map.measurable ContinuousMap.measurable theorem measurable_of_continuousOn_compl_singleton [T1Space α] {f : α → γ} (a : α) - (hf : ContinuousOn f ({a}ᶜ)) : Measurable f := + (hf : ContinuousOn f {a}ᶜ) : Measurable f := measurable_of_measurable_on_compl_singleton a (continuousOn_iff_continuous_restrict.1 hf).measurable #align measurable_of_continuous_on_compl_singleton measurable_of_continuousOn_compl_singleton @@ -1732,7 +1732,7 @@ theorem borel_eq_generateFrom_Iio_rat : borel ℝ = .generateFrom (⋃ a : ℚ, refine' fun _ => ⟨fun h => _, fun ⟨i, hai, hix⟩ => (Rat.cast_lt.2 hai).trans_le hix⟩ rcases exists_rat_btwn h with ⟨c, ac, cx⟩ exact ⟨c, Rat.cast_lt.1 ac, cx.le⟩ - : Ioo (a : ℝ) b = (⋃ c > a, Iio (c : ℝ)ᶜ) ∩ Iio (b : ℝ))] + : Ioo (a : ℝ) b = (⋃ c > a, (Iio (c : ℝ))ᶜ) ∩ Iio (b : ℝ))] · have hg : ∀ q : ℚ, MeasurableSet[g] (Iio (q : ℝ)) := fun q => GenerateMeasurable.basic (Iio (q : ℝ)) (by simp) refine' @MeasurableSet.inter _ g _ _ _ (hg _) diff --git a/Mathlib/MeasureTheory/Constructions/Polish.lean b/Mathlib/MeasureTheory/Constructions/Polish.lean index c106d414417ed..fd287ae32ed1c 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish.lean @@ -432,7 +432,7 @@ theorem AnalyticSet.measurablySeparable [T2Space α] [MeasurableSpace α] [Opens /-- **Suslin's Theorem**: in a Hausdorff topological space, an analytic set with an analytic complement is measurable. -/ theorem AnalyticSet.measurableSet_of_compl [T2Space α] [MeasurableSpace α] [OpensMeasurableSpace α] - {s : Set α} (hs : AnalyticSet s) (hsc : AnalyticSet (sᶜ)) : MeasurableSet s := by + {s : Set α} (hs : AnalyticSet s) (hsc : AnalyticSet sᶜ) : MeasurableSet s := by rcases hs.measurablySeparable hsc disjoint_compl_right with ⟨u, hsu, hdu, hmu⟩ obtain rfl : s = u := hsu.antisymm (disjoint_compl_left_iff_subset.1 hdu) exact hmu @@ -461,7 +461,7 @@ theorem measurableSet_preimage_iff_of_surjective [SecondCountableTopology Y] {f apply AnalyticSet.measurableSet_of_compl · rw [← image_preimage_eq s hsurj] exact h.analyticSet_image hf - · rw [← image_preimage_eq (sᶜ) hsurj] + · rw [← image_preimage_eq sᶜ hsurj] exact h.compl.analyticSet_image hf #align measurable.measurable_set_preimage_iff_of_surjective Measurable.measurableSet_preimage_iff_of_surjective diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index 078ee013786d3..67b38c549992d 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -171,13 +171,13 @@ theorem ae_eventually_measure_zero_of_singular (hρ : ρ ⟂ₘ μ) : intro ε εpos set s := {x | ¬∀ᶠ a in v.filterAt x, ρ a < ε * μ a} with hs change μ s = 0 - obtain ⟨o, _, ρo, μo⟩ : ∃ o : Set α, MeasurableSet o ∧ ρ o = 0 ∧ μ (oᶜ) = 0 := hρ + obtain ⟨o, _, ρo, μo⟩ : ∃ o : Set α, MeasurableSet o ∧ ρ o = 0 ∧ μ oᶜ = 0 := hρ apply le_antisymm _ bot_le calc μ s ≤ μ (s ∩ o ∪ oᶜ) := by conv_lhs => rw [← inter_union_compl s o] exact measure_mono (union_subset_union_right _ (inter_subset_right _ _)) - _ ≤ μ (s ∩ o) + μ (oᶜ) := (measure_union_le _ _) + _ ≤ μ (s ∩ o) + μ oᶜ := (measure_union_le _ _) _ = μ (s ∩ o) := by rw [μo, add_zero] _ = (ε : ℝ≥0∞)⁻¹ * (ε • μ) (s ∩ o) := by simp only [coe_nnreal_smul_apply, ← mul_assoc, mul_comm _ (ε : ℝ≥0∞)] @@ -304,8 +304,8 @@ theorem exists_measurable_supersets_limRatio {p q : ℝ≥0} (hpq : p < q) : -- a finite measure part `o n`, taking a measurable superset here, and then taking the union over -- `n`. refine' - ⟨toMeasurable μ (sᶜ) ∪ ⋃ n, toMeasurable (ρ + μ) (u n), - toMeasurable μ (sᶜ) ∪ ⋃ n, toMeasurable (ρ + μ) (w n), _, _, _, _, _⟩ + ⟨toMeasurable μ sᶜ ∪ ⋃ n, toMeasurable (ρ + μ) (u n), + toMeasurable μ sᶜ ∪ ⋃ n, toMeasurable (ρ + μ) (w n), _, _, _, _, _⟩ -- check that these sets are measurable supersets as required · exact (measurableSet_toMeasurable _ _).union @@ -317,19 +317,19 @@ theorem exists_measurable_supersets_limRatio {p q : ℝ≥0} (hpq : p < q) : by_cases h : x ∈ s · refine' Or.inr (mem_iUnion.2 ⟨spanningSetsIndex (ρ + μ) x, _⟩) exact subset_toMeasurable _ _ ⟨⟨h, hx⟩, mem_spanningSetsIndex _ _⟩ - · exact Or.inl (subset_toMeasurable μ (sᶜ) h) + · exact Or.inl (subset_toMeasurable μ sᶜ h) · intro x hx by_cases h : x ∈ s · refine' Or.inr (mem_iUnion.2 ⟨spanningSetsIndex (ρ + μ) x, _⟩) exact subset_toMeasurable _ _ ⟨⟨h, hx⟩, mem_spanningSetsIndex _ _⟩ - · exact Or.inl (subset_toMeasurable μ (sᶜ) h) + · exact Or.inl (subset_toMeasurable μ sᶜ h) -- it remains to check the nontrivial part that these sets have zero measure intersection. -- it suffices to do it for fixed `m` and `n`, as one is taking countable unions. suffices H : ∀ m n : ℕ, μ (toMeasurable (ρ + μ) (u m) ∩ toMeasurable (ρ + μ) (w n)) = 0 · have A : - (toMeasurable μ (sᶜ) ∪ ⋃ n, toMeasurable (ρ + μ) (u n)) ∩ - (toMeasurable μ (sᶜ) ∪ ⋃ n, toMeasurable (ρ + μ) (w n)) ⊆ - toMeasurable μ (sᶜ) ∪ + (toMeasurable μ sᶜ ∪ ⋃ n, toMeasurable (ρ + μ) (u n)) ∩ + (toMeasurable μ sᶜ ∪ ⋃ n, toMeasurable (ρ + μ) (w n)) ⊆ + toMeasurable μ sᶜ ∪ ⋃ (m) (n), toMeasurable (ρ + μ) (u m) ∩ toMeasurable (ρ + μ) (w n) := by simp only [inter_distrib_left, inter_distrib_right, true_and_iff, subset_union_left, union_subset_iff, inter_self] @@ -339,13 +339,13 @@ theorem exists_measurable_supersets_limRatio {p q : ℝ≥0} (hpq : p < q) : · simp_rw [iUnion_inter, inter_iUnion]; exact subset_union_right _ _ refine' le_antisymm ((measure_mono A).trans _) bot_le calc - μ (toMeasurable μ (sᶜ) ∪ + μ (toMeasurable μ sᶜ ∪ ⋃ (m) (n), toMeasurable (ρ + μ) (u m) ∩ toMeasurable (ρ + μ) (w n)) ≤ - μ (toMeasurable μ (sᶜ)) + + μ (toMeasurable μ sᶜ) + μ (⋃ (m) (n), toMeasurable (ρ + μ) (u m) ∩ toMeasurable (ρ + μ) (w n)) := measure_union_le _ _ _ = μ (⋃ (m) (n), toMeasurable (ρ + μ) (u m) ∩ toMeasurable (ρ + μ) (w n)) := by - have : μ (sᶜ) = 0 := v.ae_tendsto_div hρ; rw [measure_toMeasurable, this, zero_add] + have : μ sᶜ = 0 := v.ae_tendsto_div hρ; rw [measure_toMeasurable, this, zero_add] _ ≤ ∑' (m) (n), μ (toMeasurable (ρ + μ) (u m) ∩ toMeasurable (ρ + μ) (w n)) := ((measure_iUnion_le _).trans (ENNReal.tsum_le_tsum fun m => measure_iUnion_le _)) _ = 0 := by simp only [H, tsum_zero] @@ -448,7 +448,7 @@ proved in `measure_le_of_frequently_le`. Since `ρ a / μ a` tends almost everyw theorem measure_le_mul_of_subset_limRatioMeas_lt {p : ℝ≥0} {s : Set α} (h : s ⊆ {x | v.limRatioMeas hρ x < p}) : ρ s ≤ p * μ s := by let t := {x : α | Tendsto (fun a => ρ a / μ a) (v.filterAt x) (𝓝 (v.limRatioMeas hρ x))} - have A : μ (tᶜ) = 0 := v.ae_tendsto_limRatioMeas hρ + have A : μ tᶜ = 0 := v.ae_tendsto_limRatioMeas hρ suffices H : ρ (s ∩ t) ≤ (p • μ) (s ∩ t); exact calc @@ -472,13 +472,13 @@ proved in `measure_le_of_frequently_le`. Since `ρ a / μ a` tends almost everyw theorem mul_measure_le_of_subset_lt_limRatioMeas {q : ℝ≥0} {s : Set α} (h : s ⊆ {x | (q : ℝ≥0∞) < v.limRatioMeas hρ x}) : (q : ℝ≥0∞) * μ s ≤ ρ s := by let t := {x : α | Tendsto (fun a => ρ a / μ a) (v.filterAt x) (𝓝 (v.limRatioMeas hρ x))} - have A : μ (tᶜ) = 0 := v.ae_tendsto_limRatioMeas hρ + have A : μ tᶜ = 0 := v.ae_tendsto_limRatioMeas hρ suffices H : (q • μ) (s ∩ t) ≤ ρ (s ∩ t); exact calc (q • μ) s = (q • μ) (s ∩ t ∪ s ∩ tᶜ) := by rw [inter_union_compl] _ ≤ (q • μ) (s ∩ t) + (q • μ) (s ∩ tᶜ) := (measure_union_le _ _) - _ ≤ ρ (s ∩ t) + q * μ (tᶜ) := by + _ ≤ ρ (s ∩ t) + q * μ tᶜ := by apply add_le_add H rw [coe_nnreal_smul_apply] exact mul_le_mul_left' (measure_mono (inter_subset_right _ _)) _ diff --git a/Mathlib/MeasureTheory/Covering/LiminfLimsup.lean b/Mathlib/MeasureTheory/Covering/LiminfLimsup.lean index 4eae17425c8d7..8d83ecbaf81c4 100644 --- a/Mathlib/MeasureTheory/Covering/LiminfLimsup.lean +++ b/Mathlib/MeasureTheory/Covering/LiminfLimsup.lean @@ -125,7 +125,7 @@ theorem blimsup_cthickening_ae_le_of_eventually_mul_le_aux (p : ℕ → Prop) {s rw [eventually_atTop] refine' ⟨i, fun j hj hj' => Disjoint.inf_right (B j) <| Disjoint.inf_right' (blimsup Y₁ atTop p) _⟩ - change Disjoint (b j) (Z iᶜ) + change Disjoint (b j) (Z i)ᶜ rw [disjoint_compl_right_iff_subset] refine' (closedBall_subset_cthickening (hw j) (M * r₁ (f j))).trans ((cthickening_mono hj' _).trans fun a ha => _) diff --git a/Mathlib/MeasureTheory/Decomposition/Jordan.lean b/Mathlib/MeasureTheory/Decomposition/Jordan.lean index db51d450f03a4..5e4dc0fb404b4 100644 --- a/Mathlib/MeasureTheory/Decomposition/Jordan.lean +++ b/Mathlib/MeasureTheory/Decomposition/Jordan.lean @@ -201,7 +201,7 @@ theorem exists_compl_positive_negative : ∃ S : Set α, MeasurableSet S ∧ j.toSignedMeasure ≤[S] 0 ∧ - 0 ≤[Sᶜ] j.toSignedMeasure ∧ j.posPart S = 0 ∧ j.negPart (Sᶜ) = 0 := by + 0 ≤[Sᶜ] j.toSignedMeasure ∧ j.posPart S = 0 ∧ j.negPart Sᶜ = 0 := by obtain ⟨S, hS₁, hS₂, hS₃⟩ := j.mutuallySingular refine' ⟨S, hS₁, _, _, hS₂, hS₃⟩ · refine' restrict_le_restrict_of_subset_le _ _ fun A hA hA₁ => _ @@ -232,7 +232,7 @@ def toJordanDecomposition (s : SignedMeasure α) : JordanDecomposition α := let i := choose s.exists_compl_positive_negative let hi := choose_spec s.exists_compl_positive_negative { posPart := s.toMeasureOfZeroLE i hi.1 hi.2.1 - negPart := s.toMeasureOfLEZero (iᶜ) hi.1.compl hi.2.2 + negPart := s.toMeasureOfLEZero iᶜ hi.1.compl hi.2.2 posPart_finite := inferInstance negPart_finite := inferInstance mutuallySingular := by @@ -245,7 +245,7 @@ def toJordanDecomposition (s : SignedMeasure α) : JordanDecomposition α := theorem toJordanDecomposition_spec (s : SignedMeasure α) : ∃ (i : Set α) (hi₁ : MeasurableSet i) (hi₂ : 0 ≤[i] s) (hi₃ : s ≤[iᶜ] 0), s.toJordanDecomposition.posPart = s.toMeasureOfZeroLE i hi₁ hi₂ ∧ - s.toJordanDecomposition.negPart = s.toMeasureOfLEZero (iᶜ) hi₁.compl hi₃ := by + s.toJordanDecomposition.negPart = s.toMeasureOfLEZero iᶜ hi₁.compl hi₃ := by set i := choose s.exists_compl_positive_negative obtain ⟨hi₁, hi₂, hi₃⟩ := choose_spec s.exists_compl_positive_negative exact ⟨i, hi₁, hi₂, hi₃, rfl, rfl⟩ @@ -269,7 +269,7 @@ theorem toSignedMeasure_toJordanDecomposition (s : SignedMeasure α) : toMeasureOfLEZero_apply _ hi₃ hi₁.compl hk] simp only [ENNReal.coe_toReal, NNReal.coe_mk, ENNReal.some_eq_coe, sub_neg_eq_add] rw [← of_union _ (MeasurableSet.inter hi₁ hk) (MeasurableSet.inter hi₁.compl hk), - Set.inter_comm i, Set.inter_comm (iᶜ), Set.inter_union_compl _ _] + Set.inter_comm i, Set.inter_comm iᶜ, Set.inter_union_compl _ _] exact (disjoint_compl_right.inf_left _).inf_right _ #align measure_theory.signed_measure.to_signed_measure_to_jordan_decomposition MeasureTheory.SignedMeasure.toSignedMeasure_toJordanDecomposition diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index f95b8393201bc..00060dc78ed4b 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -234,11 +234,11 @@ theorem eq_singularPart {s : Measure α} {f : α → ℝ≥0∞} (hf : Measurabl obtain ⟨hmeas, hsing, hadd'⟩ := haveLebesgueDecomposition_spec μ ν obtain ⟨⟨S, hS₁, hS₂, hS₃⟩, ⟨T, hT₁, hT₂, hT₃⟩⟩ := hs, hsing rw [hadd'] at hadd - have hνinter : ν ((S ∩ T)ᶜ) = 0 := by + have hνinter : ν (S ∩ T)ᶜ = 0 := by rw [compl_inter] refine' nonpos_iff_eq_zero.1 (le_trans (measure_union_le _ _) _) rw [hT₃, hS₃, add_zero] - have heq : s.restrict ((S ∩ T)ᶜ) = (μ.singularPart ν).restrict ((S ∩ T)ᶜ) := by + have heq : s.restrict (S ∩ T)ᶜ = (μ.singularPart ν).restrict (S ∩ T)ᶜ := by ext1 A hA have hf : ν.withDensity f (A ∩ (S ∩ T)ᶜ) = 0 := by refine' withDensity_absolutelyContinuous ν _ _ @@ -250,7 +250,7 @@ theorem eq_singularPart {s : Measure α} {f : α → ℝ≥0∞} (hf : Measurabl exact hνinter ▸ measure_mono (inter_subset_right _ _) rw [restrict_apply hA, restrict_apply hA, ← add_zero (s (A ∩ (S ∩ T)ᶜ)), ← hf, ← add_apply, ← hadd, add_apply, hrn, add_zero] - have heq' : ∀ A : Set α, MeasurableSet A → s A = s.restrict ((S ∩ T)ᶜ) A := by + have heq' : ∀ A : Set α, MeasurableSet A → s A = s.restrict (S ∩ T)ᶜ A := by intro A hA have hsinter : s (A ∩ (S ∩ T)) = 0 := by rw [← nonpos_iff_eq_zero] @@ -318,7 +318,7 @@ theorem eq_withDensity_rnDeriv {s : Measure α} {f : α → ℝ≥0∞} (hf : Me obtain ⟨hmeas, hsing, hadd'⟩ := haveLebesgueDecomposition_spec μ ν obtain ⟨⟨S, hS₁, hS₂, hS₃⟩, ⟨T, hT₁, hT₂, hT₃⟩⟩ := hs, hsing rw [hadd'] at hadd - have hνinter : ν ((S ∩ T)ᶜ) = 0 := by + have hνinter : ν (S ∩ T)ᶜ = 0 := by rw [compl_inter] refine' nonpos_iff_eq_zero.1 (le_trans (measure_union_le _ _) _) rw [hT₃, hS₃, add_zero] @@ -404,7 +404,7 @@ theorem exists_positive_of_not_mutuallySingular (μ ν : Measure α) [IsFiniteMe choose f hf₁ hf₂ hf₃ using this -- set `A` to be the intersection of all the negative parts of obtained Hahn decompositions -- and we show that `μ A = 0` - set A := ⋂ n, f nᶜ with hA₁ + set A := ⋂ n, (f n)ᶜ with hA₁ have hAmeas : MeasurableSet A := MeasurableSet.iInter fun n => (hf₁ n).compl have hA₂ : ∀ n : ℕ, μ.toSignedMeasure - ((1 / (n + 1) : ℝ≥0) • ν).toSignedMeasure ≤[A] 0 := by intro n; exact restrict_le_restrict_subset _ _ (hf₁ n).compl (hf₃ n) (iInter_subset _ _) @@ -722,7 +722,7 @@ instance (priority := 100) haveLebesgueDecomposition_of_sigmaFinite (μ ν : Mea · exact fun n => (S.set_mem n).inter (hA₁ n) -- We will now show `ν Bᶜ = 0`. This follows since `Bᶜ = ⋃ n, S.set n ∩ (A n)ᶜ` and thus, -- `ν Bᶜ = ∑ i, ν (S.set i ∩ (A i)ᶜ) = ∑ i, (νn i) (A i)ᶜ = 0` - · have hcompl : IsCompl (⋃ n, S.set n ∩ A n) (⋃ n, S.set n ∩ A nᶜ) := by + · have hcompl : IsCompl (⋃ n, S.set n ∩ A n) (⋃ n, S.set n ∩ (A n)ᶜ) := by constructor · rw [disjoint_iff_inf_le] rintro x ⟨hx₁, hx₂⟩; rw [mem_iUnion] at hx₁ hx₂ diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index ae2f2ff750a89..7f88cf3b24fe2 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -55,7 +55,7 @@ theorem withDensity_rnDeriv_eq (μ ν : Measure α) [HaveLebesgueDecomposition rw [Measure.coe_zero, Pi.zero_apply, ← this] exact measure_mono (Set.subset_univ _) rw [← measure_add_measure_compl hE₁, hE₂, zero_add] - have : (singularPart μ ν + ν.withDensity (rnDeriv μ ν)) (Eᶜ) = μ (Eᶜ) := by rw [← hadd] + have : (singularPart μ ν + ν.withDensity (rnDeriv μ ν)) Eᶜ = μ Eᶜ := by rw [← hadd] rw [Measure.coe_add, Pi.add_apply, h hE₃] at this exact (add_eq_zero_iff.1 this).1 rw [this, zero_add] at hadd diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index 6ffd9f386358f..b0032c115f4a8 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -220,7 +220,7 @@ theorem ae_le_of_forall_set_lintegral_le_of_sigmaFinite [SigmaFinite μ] {f g : exact ((L1.and L2).and (eventually_mem_spanningSets μ x)).exists refine' le_antisymm _ bot_le calc - μ ({x : α | (fun x : α => f x ≤ g x) x}ᶜ) ≤ μ (⋃ n, s n) := measure_mono B + μ {x : α | (fun x : α => f x ≤ g x) x}ᶜ ≤ μ (⋃ n, s n) := measure_mono B _ ≤ ∑' n, μ (s n) := (measure_iUnion_le _) _ = 0 := by simp only [μs, tsum_zero] #align measure_theory.ae_le_of_forall_set_lintegral_le_of_sigma_finite MeasureTheory.ae_le_of_forall_set_lintegral_le_of_sigmaFinite @@ -496,7 +496,7 @@ theorem ae_eq_zero_of_forall_set_integral_eq_of_finStronglyMeasurable_trim (hm : (hf : FinStronglyMeasurable f (μ.trim hm)) : f =ᵐ[μ] 0 := by obtain ⟨t, ht_meas, htf_zero, htμ⟩ := hf.exists_set_sigmaFinite haveI : SigmaFinite ((μ.restrict t).trim hm) := by rwa [restrict_trim hm μ ht_meas] at htμ - have htf_zero : f =ᵐ[μ.restrict (tᶜ)] 0 := by + have htf_zero : f =ᵐ[μ.restrict tᶜ] 0 := by rw [EventuallyEq, ae_restrict_iff' (MeasurableSet.compl (hm _ ht_meas))] exact eventually_of_forall htf_zero have hf_meas_m : StronglyMeasurable[m] f := hf.stronglyMeasurable diff --git a/Mathlib/MeasureTheory/Function/AEMeasurableSequence.lean b/Mathlib/MeasureTheory/Function/AEMeasurableSequence.lean index 77439138a1867..c3118041536df 100644 --- a/Mathlib/MeasureTheory/Function/AEMeasurableSequence.lean +++ b/Mathlib/MeasureTheory/Function/AEMeasurableSequence.lean @@ -37,7 +37,7 @@ variable {ι : Sort _} {α β γ : Type _} [MeasurableSpace α] [MeasurableSpace whose complement has measure 0 such that for all `x ∈ aeSeqSet`, `f i x` is equal to `(hf i).mk (f i) x` for all `i` and we have the pointwise property `p x (fun n ↦ f n x)`. -/ def aeSeqSet (hf : ∀ i, AEMeasurable (f i) μ) (p : α → (ι → β) → Prop) : Set α := - toMeasurable μ ({ x | (∀ i, f i x = (hf i).mk (f i) x) ∧ p x fun n => f n x }ᶜ)ᶜ + (toMeasurable μ { x | (∀ i, f i x = (hf i).mk (f i) x) ∧ p x fun n => f n x }ᶜ)ᶜ #align ae_seq_set aeSeqSet /-- A sequence of measurable functions that are equal to `f` and verify property `p` on the @@ -101,7 +101,7 @@ theorem measurable (hf : ∀ i, AEMeasurable (f i) μ) (p : α → (ι → β) #align ae_seq.measurable aeSeq.measurable theorem measure_compl_aeSeqSet_eq_zero [Countable ι] (hf : ∀ i, AEMeasurable (f i) μ) - (hp : ∀ᵐ x ∂μ, p x fun n => f n x) : μ (aeSeqSet hf pᶜ) = 0 := by + (hp : ∀ᵐ x ∂μ, p x fun n => f n x) : μ (aeSeqSet hf p)ᶜ = 0 := by rw [aeSeqSet, compl_compl, measure_toMeasurable] have hf_eq := fun i => (hf i).ae_eq_mk simp_rw [Filter.EventuallyEq, ← ae_all_iff] at hf_eq @@ -120,7 +120,7 @@ theorem aeSeq_eq_mk_ae [Countable ι] (hf : ∀ i, AEMeasurable (f i) μ) theorem aeSeq_eq_fun_ae [Countable ι] (hf : ∀ i, AEMeasurable (f i) μ) (hp : ∀ᵐ x ∂μ, p x fun n => f n x) : ∀ᵐ a : α ∂μ, ∀ i : ι, aeSeq hf p i a = f i a := - haveI h_ss : { a : α | ¬∀ i : ι, aeSeq hf p i a = f i a } ⊆ aeSeqSet hf pᶜ := fun _ => + haveI h_ss : { a : α | ¬∀ i : ι, aeSeq hf p i a = f i a } ⊆ (aeSeqSet hf p)ᶜ := fun _ => mt fun hx i => aeSeq_eq_fun_of_mem_aeSeqSet hf hx i measure_mono_null h_ss (measure_compl_aeSeqSet_eq_zero hf hp) #align ae_seq.ae_seq_eq_fun_ae aeSeq.aeSeq_eq_fun_ae diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean index 0c3b8bc4abf76..d4363026ddc34 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean @@ -165,7 +165,7 @@ theorem AEStronglyMeasurable'.aeStronglyMeasurable'_of_measurableSpace_le_on {α {m m₂ m0 : MeasurableSpace α} {μ : Measure α} [TopologicalSpace E] [Zero E] (hm : m ≤ m0) {s : Set α} {f : α → E} (hs_m : MeasurableSet[m] s) (hs : ∀ t, MeasurableSet[m] (s ∩ t) → MeasurableSet[m₂] (s ∩ t)) - (hf : AEStronglyMeasurable' m f μ) (hf_zero : f =ᵐ[μ.restrict (sᶜ)] 0) : + (hf : AEStronglyMeasurable' m f μ) (hf_zero : f =ᵐ[μ.restrict sᶜ] 0) : AEStronglyMeasurable' m₂ f μ := by have h_ind_eq : s.indicator (hf.mk f) =ᵐ[μ] f := by refine' diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Indicator.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Indicator.lean index ccd9babdf6705..11cf4d92802e6 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Indicator.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Indicator.lean @@ -63,11 +63,11 @@ theorem condexp_ae_eq_restrict_zero (hs : MeasurableSet[m] s) (hf : f =ᵐ[μ.re #align measure_theory.condexp_ae_eq_restrict_zero MeasureTheory.condexp_ae_eq_restrict_zero /-- Auxiliary lemma for `condexp_indicator`. -/ -theorem condexp_indicator_aux (hs : MeasurableSet[m] s) (hf : f =ᵐ[μ.restrict (sᶜ)] 0) : +theorem condexp_indicator_aux (hs : MeasurableSet[m] s) (hf : f =ᵐ[μ.restrict sᶜ] 0) : μ[s.indicator f|m] =ᵐ[μ] s.indicator (μ[f|m]) := by by_cases hm : m ≤ m0 swap; · simp_rw [condexp_of_not_le hm, Set.indicator_zero']; rfl - have hsf_zero : ∀ g : α → E, g =ᵐ[μ.restrict (sᶜ)] 0 → s.indicator g =ᵐ[μ] g := fun g => + have hsf_zero : ∀ g : α → E, g =ᵐ[μ.restrict sᶜ] 0 → s.indicator g =ᵐ[μ] g := fun g => indicator_ae_eq_of_restrict_compl_ae_eq_zero (hm _ hs) refine' ((hsf_zero (μ[f|m]) (condexp_ae_eq_restrict_zero hs.compl hf)).trans _).symm exact condexp_congr_ae (hsf_zero f hf).symm @@ -170,11 +170,11 @@ theorem condexp_ae_eq_restrict_of_measurableSpace_eq_on {m m₂ m0 : MeasurableS suffices ∫ x in sᶜ, (μ[s.indicator f|m]) x ∂μ.restrict t = 0 by rw [this, add_zero, Measure.restrict_restrict (hm _ hs_m)] rw [Measure.restrict_restrict (MeasurableSet.compl (hm _ hs_m))] - suffices μ[s.indicator f|m] =ᵐ[μ.restrict (sᶜ)] 0 by + suffices μ[s.indicator f|m] =ᵐ[μ.restrict sᶜ] 0 by rw [Set.inter_comm, ← Measure.restrict_restrict (hm₂ _ ht)] calc - ∫ x : α in t, (μ[s.indicator f|m]) x ∂μ.restrict (sᶜ) = - ∫ x : α in t, 0 ∂μ.restrict (sᶜ) := by + ∫ x : α in t, (μ[s.indicator f|m]) x ∂μ.restrict sᶜ = + ∫ x : α in t, 0 ∂μ.restrict sᶜ := by refine' set_integral_congr_ae (hm₂ _ ht) _ filter_upwards [this] with x hx _ using hx _ = 0 := integral_zero _ _ diff --git a/Mathlib/MeasureTheory/Function/Egorov.lean b/Mathlib/MeasureTheory/Function/Egorov.lean index 06c73ac0c409f..c4982069cd69e 100644 --- a/Mathlib/MeasureTheory/Function/Egorov.lean +++ b/Mathlib/MeasureTheory/Function/Egorov.lean @@ -216,7 +216,7 @@ theorem tendstoUniformlyOn_of_ae_tendsto (hf : ∀ n, StronglyMeasurable (f n)) theorem tendstoUniformlyOn_of_ae_tendsto' [IsFiniteMeasure μ] (hf : ∀ n, StronglyMeasurable (f n)) (hg : StronglyMeasurable g) (hfg : ∀ᵐ x ∂μ, Tendsto (fun n => f n x) atTop (𝓝 (g x))) {ε : ℝ} (hε : 0 < ε) : - ∃ t, MeasurableSet t ∧ μ t ≤ ENNReal.ofReal ε ∧ TendstoUniformlyOn f g atTop (tᶜ) := by + ∃ t, MeasurableSet t ∧ μ t ≤ ENNReal.ofReal ε ∧ TendstoUniformlyOn f g atTop tᶜ := by have ⟨t, _, ht, htendsto⟩ := tendstoUniformlyOn_of_ae_tendsto hf hg MeasurableSet.univ (measure_ne_top μ Set.univ) (by filter_upwards [hfg] with _ htendsto _ using htendsto) hε refine' ⟨_, ht, _⟩ diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 4044643333b15..d5d210c16d962 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -252,8 +252,8 @@ theorem piecewise_apply {s : Set α} (hs : MeasurableSet s) (f g : α →ₛ β) #align measure_theory.simple_func.piecewise_apply MeasureTheory.SimpleFunc.piecewise_apply @[simp] -theorem piecewise_compl {s : Set α} (hs : MeasurableSet (sᶜ)) (f g : α →ₛ β) : - piecewise (sᶜ) hs f g = piecewise s hs.of_compl g f := +theorem piecewise_compl {s : Set α} (hs : MeasurableSet sᶜ) (f g : α →ₛ β) : + piecewise sᶜ hs f g = piecewise s hs.of_compl g f := coe_injective <| by simp [hs]; convert Set.piecewise_compl s f g #align measure_theory.simple_func.piecewise_compl MeasureTheory.SimpleFunc.piecewise_compl diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index 17f93c117819b..0daf8602c5dfa 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -851,8 +851,8 @@ theorem exists_simpleFunc_nonneg_ae_eq {f : Lp.simpleFunc G p μ} (hf : 0 ≤ f) rw [← Lp.simpleFunc.coeFn_nonneg] at hf have hf_ae : 0 ≤ᵐ[μ] simpleFunc.toSimpleFunc f := by filter_upwards [toSimpleFunc_eq_toFun f, hf] with _ h1 _; rwa [h1] - let s := toMeasurable μ { x | ¬0 ≤ simpleFunc.toSimpleFunc f x }ᶜ - have hs_zero : μ (sᶜ) = 0 := by + let s := (toMeasurable μ { x | ¬0 ≤ simpleFunc.toSimpleFunc f x })ᶜ + have hs_zero : μ sᶜ = 0 := by rw [compl_compl, measure_toMeasurable]; rwa [EventuallyLE, ae_iff] at hf_ae have hfs_nonneg : ∀ x ∈ s, 0 ≤ simpleFunc.toSimpleFunc f x := by intro x hxs diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index e239109b0d543..1759868caee97 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -812,7 +812,7 @@ theorem _root_.stronglyMeasurable_of_restrict_of_restrict_compl {_ : MeasurableS [TopologicalSpace β] {f : α → β} {s : Set α} (hs : MeasurableSet s) (h₁ : StronglyMeasurable (s.restrict f)) (h₂ : StronglyMeasurable (sᶜ.restrict f)) : StronglyMeasurable f := - stronglyMeasurable_of_stronglyMeasurable_union_cover s (sᶜ) hs hs.compl (union_compl_self s).ge h₁ + stronglyMeasurable_of_stronglyMeasurable_union_cover s sᶜ hs hs.compl (union_compl_self s).ge h₁ h₂ #align strongly_measurable_of_restrict_of_restrict_compl stronglyMeasurable_of_restrict_of_restrict_compl @@ -1062,7 +1062,7 @@ theorem exists_set_sigmaFinite [Zero β] [TopologicalSpace β] [T2Space β] · rw [Measure.restrict_apply' (MeasurableSet.iUnion hT_meas), Set.union_inter_distrib_right, Set.compl_inter_self t, Set.empty_union] exact (measure_mono (Set.inter_subset_left _ _)).trans_lt (hT_lt_top n) - · rw [← Set.union_iUnion (tᶜ) T] + · rw [← Set.union_iUnion tᶜ T] exact Set.compl_union_self _ #align measure_theory.fin_strongly_measurable.exists_set_sigma_finite MeasureTheory.FinStronglyMeasurable.exists_set_sigmaFinite @@ -1535,7 +1535,7 @@ theorem _root_.aestronglyMeasurable_indicator_iff [Zero β] {s : Set α} (hs : M refine' ⟨indicator s (h.mk f), h.stronglyMeasurable_mk.indicator hs, _⟩ have A : s.indicator f =ᵐ[μ.restrict s] s.indicator (h.mk f) := (indicator_ae_eq_restrict hs).trans (h.ae_eq_mk.trans <| (indicator_ae_eq_restrict hs).symm) - have B : s.indicator f =ᵐ[μ.restrict (sᶜ)] s.indicator (h.mk f) := + have B : s.indicator f =ᵐ[μ.restrict sᶜ] s.indicator (h.mk f) := (indicator_ae_eq_restrict_compl hs).trans (indicator_ae_eq_restrict_compl hs).symm exact ae_of_ae_restrict_of_ae_restrict_compl _ A B #align ae_strongly_measurable_indicator_iff aestronglyMeasurable_indicator_iff @@ -1975,7 +1975,7 @@ end Order variable [Zero β] [T2Space β] theorem exists_set_sigmaFinite (hf : AEFinStronglyMeasurable f μ) : - ∃ t, MeasurableSet t ∧ f =ᵐ[μ.restrict (tᶜ)] 0 ∧ SigmaFinite (μ.restrict t) := by + ∃ t, MeasurableSet t ∧ f =ᵐ[μ.restrict tᶜ] 0 ∧ SigmaFinite (μ.restrict t) := by rcases hf with ⟨g, hg, hfg⟩ obtain ⟨t, ht, hgt_zero, htμ⟩ := hg.exists_set_sigmaFinite refine' ⟨t, ht, _, htμ⟩ @@ -1995,7 +1995,7 @@ protected theorem measurableSet (hf : AEFinStronglyMeasurable f μ) : #align measure_theory.ae_fin_strongly_measurable.measurable_set MeasureTheory.AEFinStronglyMeasurable.measurableSet theorem ae_eq_zero_compl (hf : AEFinStronglyMeasurable f μ) : - f =ᵐ[μ.restrict (hf.sigmaFiniteSetᶜ)] 0 := + f =ᵐ[μ.restrict hf.sigmaFiniteSetᶜ] 0 := hf.exists_set_sigmaFinite.choose_spec.2.1 #align measure_theory.ae_fin_strongly_measurable.ae_eq_zero_compl MeasureTheory.AEFinStronglyMeasurable.ae_eq_zero_compl diff --git a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean index 60df0e6bbd8cc..38f3d5d7c5ab9 100644 --- a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean @@ -535,7 +535,7 @@ theorem tendsto_Lp_of_tendsto_ae_of_meas [IsFiniteMeasure μ] (hp : 1 ≤ p) (hp (dist_comm (g x) (f n x) ▸ (hN x hx).le : dist (f n x) (g x) ≤ ε.toReal / (3 * measureUnivNNReal μ ^ (1 / p.toReal))) refine' le_trans this _ - rw [div_mul_eq_div_mul_one_div, ← ENNReal.ofReal_toReal (measure_lt_top μ (tᶜ)).ne, + rw [div_mul_eq_div_mul_one_div, ← ENNReal.ofReal_toReal (measure_lt_top μ tᶜ).ne, ENNReal.ofReal_rpow_of_nonneg ENNReal.toReal_nonneg hdivp, ← ENNReal.ofReal_mul, mul_assoc] · refine' ENNReal.ofReal_le_ofReal (mul_le_of_le_one_right hε'.le _) rw [mul_comm, mul_one_div, div_le_one] diff --git a/Mathlib/MeasureTheory/Integral/Average.lean b/Mathlib/MeasureTheory/Integral/Average.lean index d368ace0f8216..a79047a94d6c9 100644 --- a/Mathlib/MeasureTheory/Integral/Average.lean +++ b/Mathlib/MeasureTheory/Integral/Average.lean @@ -189,7 +189,7 @@ theorem average_union_mem_segment {f : α → E} {s t : Set α} (hd : AEDisjoint #align measure_theory.average_union_mem_segment MeasureTheory.average_union_mem_segment theorem average_mem_openSegment_compl_self [IsFiniteMeasure μ] {f : α → E} {s : Set α} - (hs : NullMeasurableSet s μ) (hs₀ : μ s ≠ 0) (hsc₀ : μ (sᶜ) ≠ 0) (hfi : Integrable f μ) : + (hs : NullMeasurableSet s μ) (hs₀ : μ s ≠ 0) (hsc₀ : μ sᶜ ≠ 0) (hfi : Integrable f μ) : ⨍ x, f x ∂μ ∈ openSegment ℝ (⨍ x in s, f x ∂μ) (⨍ x in sᶜ, f x ∂μ) := by simpa only [union_compl_self, restrict_univ] using average_union_mem_openSegment aedisjoint_compl_right hs.compl hs₀ hsc₀ (measure_ne_top _ _) diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index f133c22be45c3..75f5e3055ab1e 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -292,7 +292,7 @@ theorem lintegral_mono_ae {f g : α → ℝ≥0∞} (h : ∀ᵐ a ∂μ, f a ≤ rcases exists_measurable_superset_of_null h with ⟨t, hts, ht, ht0⟩ have : ∀ᵐ x ∂μ, x ∉ t := measure_zero_iff_ae_nmem.1 ht0 rw [lintegral, lintegral] - refine' iSup_le fun s => iSup_le fun hfs => le_iSup_of_le (s.restrict (tᶜ)) <| le_iSup_of_le _ _ + refine' iSup_le fun s => iSup_le fun hfs => le_iSup_of_le (s.restrict tᶜ) <| le_iSup_of_le _ _ · intro a by_cases h : a ∈ t <;> simp [h, restrict_apply s ht.compl, ht.compl] exact le_trans (hfs a) (_root_.by_contradiction fun hnfg => h (hts hnfg)) @@ -1812,7 +1812,7 @@ theorem lintegral_withDensity_eq_lintegral_mul₀' {μ : Measure α} {f : α → filter_upwards [Z] intro x hx simp only [hx, Pi.mul_apply] - · have M : MeasurableSet ({ x : α | f' x ≠ 0 }ᶜ) := + · have M : MeasurableSet { x : α | f' x ≠ 0 }ᶜ := (hf.measurable_mk (measurableSet_singleton 0).compl).compl filter_upwards [ae_restrict_mem M] intro x hx diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 924748ea30c12..00ac1ebf0b1c8 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -205,7 +205,7 @@ theorem ofReal_set_integral_one {α : Type _} {_ : MeasurableSpace α} (μ : Mea #align measure_theory.of_real_set_integral_one MeasureTheory.ofReal_set_integral_one theorem integral_piecewise [DecidablePred (· ∈ s)] (hs : MeasurableSet s) (hf : IntegrableOn f s μ) - (hg : IntegrableOn g (sᶜ) μ) : + (hg : IntegrableOn g sᶜ μ) : ∫ x, s.piecewise f g x ∂μ = (∫ x in s, f x ∂μ) + ∫ x in sᶜ, g x ∂μ := by rw [← Set.indicator_add_compl_eq_piecewise, integral_add' (hf.integrable_indicator hs) (hg.integrable_indicator hs.compl), diff --git a/Mathlib/MeasureTheory/MeasurableSpace.lean b/Mathlib/MeasureTheory/MeasurableSpace.lean index ecbf9066b6255..37f6eb0abd620 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace.lean @@ -603,7 +603,7 @@ theorem measurable_of_measurable_union_cover {f : α → β} (s t : Set α) (hs theorem measurable_of_restrict_of_restrict_compl {f : α → β} {s : Set α} (hs : MeasurableSet s) (h₁ : Measurable (s.restrict f)) (h₂ : Measurable (sᶜ.restrict f)) : Measurable f := - measurable_of_measurable_union_cover s (sᶜ) hs hs.compl (union_compl_self s).ge h₁ h₂ + measurable_of_measurable_union_cover s sᶜ hs hs.compl (union_compl_self s).ge h₁ h₂ #align measurable_of_restrict_of_restrict_compl measurable_of_restrict_of_restrict_compl theorem Measurable.dite [∀ x, Decidable (x ∈ s)] {f : s → β} (hf : Measurable f) @@ -1843,7 +1843,7 @@ instance Subtype.instHasCompl : HasCompl (Subtype (MeasurableSet : Set α → Pr #align measurable_set.subtype.has_compl MeasurableSet.Subtype.instHasCompl @[simp] -theorem coe_compl (s : Subtype (MeasurableSet : Set α → Prop)) : ↑(sᶜ) = (sᶜ : Set α) := +theorem coe_compl (s : Subtype (MeasurableSet : Set α → Prop)) : ↑sᶜ = (sᶜ : Set α) := rfl #align measurable_set.coe_compl MeasurableSet.coe_compl diff --git a/Mathlib/MeasureTheory/MeasurableSpaceDef.lean b/Mathlib/MeasureTheory/MeasurableSpaceDef.lean index 1756846d1b263..5cac7369fb253 100644 --- a/Mathlib/MeasureTheory/MeasurableSpaceDef.lean +++ b/Mathlib/MeasureTheory/MeasurableSpaceDef.lean @@ -55,7 +55,7 @@ variable {α β γ δ δ' : Type _} {ι : Sort _} {s t u : Set α} /-- The empty set is a measurable set. Use `MeasurableSet.empty` instead. -/ measurableSet_empty : MeasurableSet' ∅ /-- The complement of a measurable set is a measurable set. Use `MeasurableSet.compl` instead. -/ - measurableSet_compl : ∀ s, MeasurableSet' s → MeasurableSet' (sᶜ) + measurableSet_compl : ∀ s, MeasurableSet' s → MeasurableSet' sᶜ /-- The union of a sequence of measurable sets is a measurable set. Use a more general `MeasurableSet.iUnion` instead. -/ measurableSet_iUnion : ∀ f : ℕ → Set α, (∀ i, MeasurableSet' (f i)) → MeasurableSet' (⋃ i, f i) @@ -87,16 +87,16 @@ theorem MeasurableSet.empty [MeasurableSpace α] : MeasurableSet (∅ : Set α) variable {m : MeasurableSpace α} @[measurability] -protected theorem MeasurableSet.compl : MeasurableSet s → MeasurableSet (sᶜ) := +protected theorem MeasurableSet.compl : MeasurableSet s → MeasurableSet sᶜ := MeasurableSpace.measurableSet_compl _ s #align measurable_set.compl MeasurableSet.compl -protected theorem MeasurableSet.of_compl (h : MeasurableSet (sᶜ)) : MeasurableSet s := +protected theorem MeasurableSet.of_compl (h : MeasurableSet sᶜ) : MeasurableSet s := compl_compl s ▸ h.compl #align measurable_set.of_compl MeasurableSet.of_compl @[simp] -theorem MeasurableSet.compl_iff : MeasurableSet (sᶜ) ↔ MeasurableSet s := +theorem MeasurableSet.compl_iff : MeasurableSet sᶜ ↔ MeasurableSet s := ⟨.of_compl, .compl⟩ #align measurable_set.compl_iff MeasurableSet.compl_iff @@ -358,7 +358,7 @@ instance : PartialOrder (MeasurableSpace α) := inductive GenerateMeasurable (s : Set (Set α)) : Set α → Prop | protected basic : ∀ u ∈ s, GenerateMeasurable s u | protected empty : GenerateMeasurable s ∅ - | protected compl : ∀ t, GenerateMeasurable s t → GenerateMeasurable s (tᶜ) + | protected compl : ∀ t, GenerateMeasurable s t → GenerateMeasurable s tᶜ | protected iUnion : ∀ f : ℕ → Set α, (∀ n, GenerateMeasurable s (f n)) → GenerateMeasurable s (⋃ i, f i) #align measurable_space.generate_measurable MeasurableSpace.GenerateMeasurable @@ -378,7 +378,7 @@ theorem measurableSet_generateFrom {s : Set (Set α)} {t : Set α} (ht : t ∈ s @[elab_as_elim] theorem generateFrom_induction (p : Set α → Prop) (C : Set (Set α)) (hC : ∀ t ∈ C, p t) - (h_empty : p ∅) (h_compl : ∀ t, p t → p (tᶜ)) + (h_empty : p ∅) (h_compl : ∀ t, p t → p tᶜ) (h_Union : ∀ f : ℕ → Set α, (∀ n, p (f n)) → p (⋃ i, f i)) {s : Set α} (hs : MeasurableSet[generateFrom C] s) : p s := by induction hs diff --git a/Mathlib/MeasureTheory/Measure/AEDisjoint.lean b/Mathlib/MeasureTheory/Measure/AEDisjoint.lean index 0b1bd911f1fe3..f566bdd3dcc77 100644 --- a/Mathlib/MeasureTheory/Measure/AEDisjoint.lean +++ b/Mathlib/MeasureTheory/Measure/AEDisjoint.lean @@ -157,11 +157,11 @@ theorem of_null_left (h : μ s = 0) : AEDisjoint μ s t := end AEDisjoint -theorem aedisjoint_compl_left : AEDisjoint μ (sᶜ) s := +theorem aedisjoint_compl_left : AEDisjoint μ sᶜ s := (@disjoint_compl_left _ _ s).aedisjoint #align measure_theory.ae_disjoint_compl_left MeasureTheory.aedisjoint_compl_left -theorem aedisjoint_compl_right : AEDisjoint μ s (sᶜ) := +theorem aedisjoint_compl_right : AEDisjoint μ s sᶜ := (@disjoint_compl_right _ _ s).aedisjoint #align measure_theory.ae_disjoint_compl_right MeasureTheory.aedisjoint_compl_right diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index 4c59c2aa6637a..d2a43b5571b05 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -92,7 +92,7 @@ theorem sum_measure [Countable ι] {μ : ι → Measure α} (h : ∀ i, AEMeasur · rw [restrict_piecewise_compl, compl_iInter] intro t ht refine' - ⟨⋃ i, (h i).mk f ⁻¹' t ∩ s iᶜ, + ⟨⋃ i, (h i).mk f ⁻¹' t ∩ (s i)ᶜ, MeasurableSet.iUnion fun i => (measurable_mk _ ht).inter (measurableSet_toMeasurable _ _).compl, _⟩ @@ -195,7 +195,7 @@ theorem prod_mk {f : α → β} {g : α → γ} (hf : AEMeasurable f μ) (hg : A theorem exists_ae_eq_range_subset (H : AEMeasurable f μ) {t : Set β} (ht : ∀ᵐ x ∂μ, f x ∈ t) (h₀ : t.Nonempty) : ∃ g, Measurable g ∧ range g ⊆ t ∧ f =ᵐ[μ] g := by - let s : Set α := toMeasurable μ ({ x | f x = H.mk f x ∧ f x ∈ t }ᶜ) + let s : Set α := toMeasurable μ { x | f x = H.mk f x ∧ f x ∈ t }ᶜ let g : α → β := piecewise s (fun _ => h₀.some) (H.mk f) refine' ⟨g, _, _, _⟩ · exact Measurable.piecewise (measurableSet_toMeasurable _ _) measurable_const H.measurable_mk @@ -207,7 +207,7 @@ theorem exists_ae_eq_range_subset (H : AEMeasurable f μ) {t : Set β} (ht : ∀ apply subset_toMeasurable simp (config := { contextual := true }) only [hx, mem_compl_iff, mem_setOf_eq, not_and, not_false_iff, imp_true_iff] - · have A : μ (toMeasurable μ ({ x | f x = H.mk f x ∧ f x ∈ t }ᶜ)) = 0 := by + · have A : μ (toMeasurable μ { x | f x = H.mk f x ∧ f x ∈ t }ᶜ) = 0 := by rw [measure_toMeasurable, ← compl_mem_ae_iff, compl_compl] exact H.ae_eq_mk.and ht filter_upwards [compl_mem_ae_iff.2 A]with x hx @@ -342,7 +342,7 @@ theorem aemeasurable_indicator_iff {s} (hs : MeasurableSet s) : refine' ⟨indicator s (h.mk f), h.measurable_mk.indicator hs, _⟩ have A : s.indicator f =ᵐ[μ.restrict s] s.indicator (AEMeasurable.mk f h) := (indicator_ae_eq_restrict hs).trans (h.ae_eq_mk.trans <| (indicator_ae_eq_restrict hs).symm) - have B : s.indicator f =ᵐ[μ.restrict (sᶜ)] s.indicator (AEMeasurable.mk f h) := + have B : s.indicator f =ᵐ[μ.restrict sᶜ] s.indicator (AEMeasurable.mk f h) := (indicator_ae_eq_restrict_compl hs).trans (indicator_ae_eq_restrict_compl hs).symm exact ae_of_ae_restrict_of_ae_restrict_compl _ A B #align ae_measurable_indicator_iff aemeasurable_indicator_iff diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 4fe9df49cd407..46284903c3b2f 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -781,11 +781,11 @@ theorem tendsto_add_haar_inter_smul_one_of_density_one_aux (s : Set E) (hs : Mea apply B.congr' _ filter_upwards [self_mem_nhdsWithin] rintro r (rpos : 0 < r) - convert I (closedBall x r) (sᶜ) (measure_closedBall_pos μ _ rpos).ne' + convert I (closedBall x r) sᶜ (measure_closedBall_pos μ _ rpos).ne' measure_closedBall_lt_top.ne hs.compl rw [compl_compl] have L' : Tendsto (fun r : ℝ => μ (sᶜ ∩ ({x} + r • t)) / μ ({x} + r • t)) (𝓝[>] 0) (𝓝 0) := - tendsto_add_haar_inter_smul_zero_of_density_zero μ (sᶜ) x L t ht h''t + tendsto_add_haar_inter_smul_zero_of_density_zero μ sᶜ x L t ht h''t have L'' : Tendsto (fun r : ℝ => μ ({x} + r • t) / μ ({x} + r • t)) (𝓝[>] 0) (𝓝 1) := by apply tendsto_const_nhds.congr' _ filter_upwards [self_mem_nhdsWithin] diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index d7a66604cd7e0..f3c4a6f32db16 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -153,7 +153,7 @@ theorem measure_union_add_inter' (hs : MeasurableSet s) (t : Set α) : rw [union_comm, inter_comm, measure_union_add_inter t hs, add_comm] #align measure_theory.measure_union_add_inter' MeasureTheory.measure_union_add_inter' -theorem measure_add_measure_compl (h : MeasurableSet s) : μ s + μ (sᶜ) = μ univ := +theorem measure_add_measure_compl (h : MeasurableSet s) : μ s + μ sᶜ = μ univ := measure_add_measure_compl₀ h.nullMeasurableSet #align measure_theory.measure_add_measure_compl MeasureTheory.measure_add_measure_compl @@ -288,7 +288,7 @@ theorem measure_eq_measure_larger_of_between_null_diff {s₁ s₂ s₃ : Set α} (measure_eq_measure_of_between_null_diff h12 h23 h_nulldiff).2 #align measure_theory.measure_eq_measure_larger_of_between_null_diff MeasureTheory.measure_eq_measure_larger_of_between_null_diff -theorem measure_compl (h₁ : MeasurableSet s) (h_fin : μ s ≠ ∞) : μ (sᶜ) = μ univ - μ s := by +theorem measure_compl (h₁ : MeasurableSet s) (h_fin : μ s ≠ ∞) : μ sᶜ = μ univ - μ s := by rw [compl_eq_univ_diff] exact measure_diff (subset_univ s) h₁ h_fin #align measure_theory.measure_compl MeasureTheory.measure_compl @@ -1736,13 +1736,13 @@ theorem restrict_union' (h : Disjoint s t) (hs : MeasurableSet s) : @[simp] theorem restrict_add_restrict_compl (hs : MeasurableSet s) : - μ.restrict s + μ.restrict (sᶜ) = μ := by + μ.restrict s + μ.restrict sᶜ = μ := by rw [← restrict_union (@disjoint_compl_right (Set α) _ _) hs.compl, union_compl_self, restrict_univ] #align measure_theory.measure.restrict_add_restrict_compl MeasureTheory.Measure.restrict_add_restrict_compl @[simp] -theorem restrict_compl_add_restrict (hs : MeasurableSet s) : μ.restrict (sᶜ) + μ.restrict s = μ := +theorem restrict_compl_add_restrict (hs : MeasurableSet s) : μ.restrict sᶜ + μ.restrict s = μ := by rw [add_comm, restrict_add_restrict_compl hs] #align measure_theory.measure.restrict_compl_add_restrict MeasureTheory.Measure.restrict_compl_add_restrict @@ -1996,7 +1996,7 @@ theorem dirac_apply [MeasurableSingletonClass α] (a : α) (s : Set α) : by_cases h : a ∈ s; · rw [dirac_apply_of_mem h, indicator_of_mem h, Pi.one_apply] rw [indicator_of_not_mem h, ← nonpos_iff_eq_zero] calc - dirac a s ≤ dirac a ({a}ᶜ) := measure_mono (subset_compl_comm.1 <| singleton_subset_iff.2 h) + dirac a s ≤ dirac a {a}ᶜ := measure_mono (subset_compl_comm.1 <| singleton_subset_iff.2 h) _ = 0 := by simp [dirac_apply' _ (measurableSet_singleton _).compl] #align measure_theory.measure.dirac_apply MeasureTheory.Measure.dirac_apply @@ -2113,7 +2113,7 @@ theorem sum_of_empty [IsEmpty ι] (μ : ι → Measure α) : sum μ = 0 := by #align measure_theory.measure.sum_of_empty MeasureTheory.Measure.sum_of_empty theorem sum_add_sum_compl (s : Set ι) (μ : ι → Measure α) : - ((sum fun i : s => μ i) + sum fun i : ↥(sᶜ) => μ i) = sum μ := by + ((sum fun i : s => μ i) + sum fun i : ↥sᶜ => μ i) = sum μ := by ext1 t ht simp only [add_apply, sum_apply _ ht] exact @tsum_add_tsum_compl ℝ≥0∞ ι _ _ _ (fun i => μ i t) _ s ENNReal.summable ENNReal.summable @@ -2624,17 +2624,17 @@ end Pointwise /-- The filter of sets `s` such that `sᶜ` has finite measure. -/ def cofinite {m0 : MeasurableSpace α} (μ : Measure α) : Filter α where - sets := { s | μ (sᶜ) < ∞ } + sets := { s | μ sᶜ < ∞ } univ_sets := by simp inter_sets {s t} hs ht := by simp only [compl_inter, mem_setOf_eq] calc - μ (sᶜ ∪ tᶜ) ≤ μ (sᶜ) + μ (tᶜ) := measure_union_le _ _ + μ (sᶜ ∪ tᶜ) ≤ μ sᶜ + μ tᶜ := measure_union_le _ _ _ < ∞ := ENNReal.add_lt_top.2 ⟨hs, ht⟩ sets_of_superset {s t} hs hst := lt_of_le_of_lt (measure_mono <| compl_subset_compl.2 hst) hs #align measure_theory.measure.cofinite MeasureTheory.Measure.cofinite -theorem mem_cofinite : s ∈ μ.cofinite ↔ μ (sᶜ) < ∞ := +theorem mem_cofinite : s ∈ μ.cofinite ↔ μ sᶜ < ∞ := Iff.rfl #align measure_theory.measure.mem_cofinite MeasureTheory.Measure.mem_cofinite @@ -2846,13 +2846,13 @@ theorem ae_restrict_of_ae_restrict_of_subset {s t : Set α} {p : α → Prop} (h #align measure_theory.ae_restrict_of_ae_restrict_of_subset MeasureTheory.ae_restrict_of_ae_restrict_of_subset theorem ae_of_ae_restrict_of_ae_restrict_compl (t : Set α) {p : α → Prop} - (ht : ∀ᵐ x ∂μ.restrict t, p x) (htc : ∀ᵐ x ∂μ.restrict (tᶜ), p x) : ∀ᵐ x ∂μ, p x := + (ht : ∀ᵐ x ∂μ.restrict t, p x) (htc : ∀ᵐ x ∂μ.restrict tᶜ, p x) : ∀ᵐ x ∂μ, p x := nonpos_iff_eq_zero.1 <| calc μ { x | ¬p x } = μ ({ x | ¬p x } ∩ t ∪ { x | ¬p x } ∩ tᶜ) := by rw [← inter_union_distrib_left, union_compl_self, inter_univ] _ ≤ μ ({ x | ¬p x } ∩ t) + μ ({ x | ¬p x } ∩ tᶜ) := (measure_union_le _ _) - _ ≤ μ.restrict t { x | ¬p x } + μ.restrict (tᶜ) { x | ¬p x } := + _ ≤ μ.restrict t { x | ¬p x } + μ.restrict tᶜ { x | ¬p x } := (add_le_add (le_restrict_apply _ _) (le_restrict_apply _ _)) _ = 0 := by rw [ae_iff.1 ht, ae_iff.1 htc, zero_add] @@ -3061,7 +3061,7 @@ theorem measure_ne_top (μ : Measure α) [IsFiniteMeasure μ] (s : Set α) : μ #align measure_theory.measure_ne_top MeasureTheory.measure_ne_top theorem measure_compl_le_add_of_le_add [IsFiniteMeasure μ] (hs : MeasurableSet s) - (ht : MeasurableSet t) {ε : ℝ≥0∞} (h : μ s ≤ μ t + ε) : μ (tᶜ) ≤ μ (sᶜ) + ε := by + (ht : MeasurableSet t) {ε : ℝ≥0∞} (h : μ s ≤ μ t + ε) : μ tᶜ ≤ μ sᶜ + ε := by rw [measure_compl ht (measure_ne_top μ _), measure_compl hs (measure_ne_top μ _), tsub_le_iff_right] calc @@ -3072,7 +3072,7 @@ theorem measure_compl_le_add_of_le_add [IsFiniteMeasure μ] (hs : MeasurableSet #align measure_theory.measure_compl_le_add_of_le_add MeasureTheory.measure_compl_le_add_of_le_add theorem measure_compl_le_add_iff [IsFiniteMeasure μ] (hs : MeasurableSet s) (ht : MeasurableSet t) - {ε : ℝ≥0∞} : μ (sᶜ) ≤ μ (tᶜ) + ε ↔ μ t ≤ μ s + ε := + {ε : ℝ≥0∞} : μ sᶜ ≤ μ tᶜ + ε ↔ μ t ≤ μ s + ε := ⟨fun h => compl_compl s ▸ compl_compl t ▸ measure_compl_le_add_of_le_add hs.compl ht.compl h, measure_compl_le_add_of_le_add ht hs⟩ #align measure_theory.measure_compl_le_add_iff MeasureTheory.measure_compl_le_add_iff @@ -3222,7 +3222,7 @@ instance Measure.dirac.isProbabilityMeasure [MeasurableSpace α] {x : α} : ⟨dirac_apply_of_mem <| mem_univ x⟩ #align measure_theory.measure.dirac.is_probability_measure MeasureTheory.Measure.dirac.isProbabilityMeasure -theorem prob_add_prob_compl [IsProbabilityMeasure μ] (h : MeasurableSet s) : μ s + μ (sᶜ) = 1 := +theorem prob_add_prob_compl [IsProbabilityMeasure μ] (h : MeasurableSet s) : μ s + μ sᶜ = 1 := (measure_add_measure_compl h).trans measure_univ #align measure_theory.prob_add_prob_compl MeasureTheory.prob_add_prob_compl @@ -3251,19 +3251,19 @@ theorem one_le_prob_iff [IsProbabilityMeasure μ] : 1 ≤ μ s ↔ μ s = 1 := /-- Note that this is not quite as useful as it looks because the measure takes values in `ℝ≥0∞`. Thus the subtraction appearing is the truncated subtraction of `ℝ≥0∞`, rather than the better-behaved subtraction of `ℝ`. -/ -theorem prob_compl_eq_one_sub [IsProbabilityMeasure μ] (hs : MeasurableSet s) : μ (sᶜ) = 1 - μ s := +theorem prob_compl_eq_one_sub [IsProbabilityMeasure μ] (hs : MeasurableSet s) : μ sᶜ = 1 - μ s := by simpa only [measure_univ] using measure_compl hs (measure_lt_top μ s).ne #align measure_theory.prob_compl_eq_one_sub MeasureTheory.prob_compl_eq_one_sub @[simp] theorem prob_compl_eq_zero_iff [IsProbabilityMeasure μ] (hs : MeasurableSet s) : - μ (sᶜ) = 0 ↔ μ s = 1 := by + μ sᶜ = 0 ↔ μ s = 1 := by rw [prob_compl_eq_one_sub hs, tsub_eq_zero_iff_le, one_le_prob_iff] #align measure_theory.prob_compl_eq_zero_iff MeasureTheory.prob_compl_eq_zero_iff @[simp] theorem prob_compl_eq_one_iff [IsProbabilityMeasure μ] (hs : MeasurableSet s) : - μ (sᶜ) = 1 ↔ μ s = 0 := by rw [← prob_compl_eq_zero_iff hs.compl, compl_compl] + μ sᶜ = 1 ↔ μ s = 0 := by rw [← prob_compl_eq_zero_iff hs.compl, compl_compl] #align measure_theory.prob_compl_eq_one_iff MeasureTheory.prob_compl_eq_one_iff end IsProbabilityMeasure @@ -3385,7 +3385,7 @@ theorem ite_ae_eq_of_measure_zero {γ} (f : α → γ) (g : α → γ) (s : Set #align measure_theory.ite_ae_eq_of_measure_zero MeasureTheory.ite_ae_eq_of_measure_zero theorem ite_ae_eq_of_measure_compl_zero {γ} (f : α → γ) (g : α → γ) (s : Set α) - (hs_zero : μ (sᶜ) = 0) : (fun x => ite (x ∈ s) (f x) (g x)) =ᵐ[μ] f := by + (hs_zero : μ sᶜ = 0) : (fun x => ite (x ∈ s) (f x) (g x)) =ᵐ[μ] f := by change s ∈ μ.ae at hs_zero filter_upwards [hs_zero] intros @@ -4029,7 +4029,7 @@ theorem exists_ne_forall_mem_nhds_pos_measure_preimage {β} [TopologicalSpace β ∃ a b : β, a ≠ b ∧ (∀ s ∈ 𝓝 a, 0 < μ (f ⁻¹' s)) ∧ ∀ t ∈ 𝓝 b, 0 < μ (f ⁻¹' t) := by -- We use an `OuterMeasure` so that the proof works without `Measurable f` set m : OuterMeasure β := OuterMeasure.map f μ.toOuterMeasure - replace h : ∀ b : β, m ({b}ᶜ) ≠ 0 := fun b => not_eventually.mpr (h b) + replace h : ∀ b : β, m {b}ᶜ ≠ 0 := fun b => not_eventually.mpr (h b) inhabit β have : m univ ≠ 0 := ne_bot_of_le_ne_bot (h default) (m.mono' <| subset_univ _) rcases m.exists_mem_forall_mem_nhds_within_pos this with ⟨b, -, hb⟩ @@ -4179,7 +4179,7 @@ variable {m0 : MeasurableSpace α} {m1 : MeasurableSpace β} {f : α → β} (hf nonrec theorem map_apply (μ : Measure α) (s : Set β) : μ.map f s = μ (f ⁻¹' s) := by refine' le_antisymm _ (le_map_apply hf.measurable.aemeasurable s) - set t := f '' toMeasurable μ (f ⁻¹' s) ∪ range fᶜ + set t := f '' toMeasurable μ (f ⁻¹' s) ∪ (range f)ᶜ have htm : MeasurableSet t := (hf.measurableSet_image.2 <| measurableSet_toMeasurable _ _).union hf.measurableSet_range.compl @@ -4618,7 +4618,7 @@ theorem piecewise_ae_eq_restrict (hs : MeasurableSet s) : piecewise s f g =ᵐ[ #align piecewise_ae_eq_restrict piecewise_ae_eq_restrict theorem piecewise_ae_eq_restrict_compl (hs : MeasurableSet s) : - piecewise s f g =ᵐ[μ.restrict (sᶜ)] g := by + piecewise s f g =ᵐ[μ.restrict sᶜ] g := by rw [ae_restrict_eq hs.compl] exact (piecewise_eqOn_compl s f g).eventuallyEq.filter_mono inf_le_right #align piecewise_ae_eq_restrict_compl piecewise_ae_eq_restrict_compl @@ -4669,12 +4669,12 @@ theorem indicator_ae_eq_restrict (hs : MeasurableSet s) : indicator s f =ᵐ[μ. #align indicator_ae_eq_restrict indicator_ae_eq_restrict theorem indicator_ae_eq_restrict_compl (hs : MeasurableSet s) : - indicator s f =ᵐ[μ.restrict (sᶜ)] 0 := + indicator s f =ᵐ[μ.restrict sᶜ] 0 := piecewise_ae_eq_restrict_compl hs #align indicator_ae_eq_restrict_compl indicator_ae_eq_restrict_compl theorem indicator_ae_eq_of_restrict_compl_ae_eq_zero (hs : MeasurableSet s) - (hf : f =ᵐ[μ.restrict (sᶜ)] 0) : s.indicator f =ᵐ[μ] f := by + (hf : f =ᵐ[μ.restrict sᶜ] 0) : s.indicator f =ᵐ[μ] f := by rw [Filter.EventuallyEq, ae_restrict_iff' hs.compl] at hf filter_upwards [hf]with x hx by_cases hxs : x ∈ s diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean index 383428651bb03..6dbe04d51150e 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean @@ -353,7 +353,7 @@ theorem measure_inter_null_of_null_left {S : Set α} (T : Set α) (h : μ S = 0) /-- The “almost everywhere” filter of co-null sets. -/ def Measure.ae {α} {m : MeasurableSpace α} (μ : Measure α) : Filter α where - sets := { s | μ (sᶜ) = 0 } + sets := { s | μ sᶜ = 0 } univ_sets := by simp inter_sets hs ht := by simp only [compl_inter, mem_setOf_eq]; exact measure_union_null hs ht sets_of_superset hs hst := measure_mono_null (Set.compl_subset_compl.2 hst) hs @@ -371,7 +371,7 @@ notation:50 f " =ᵐ[" μ:50 "] " g:50 => Filter.EventuallyEq (Measure.ae μ) f -- mathport name: «expr ≤ᵐ[ ] » notation:50 f " ≤ᵐ[" μ:50 "] " g:50 => Filter.EventuallyLE (Measure.ae μ) f g -theorem mem_ae_iff {s : Set α} : s ∈ μ.ae ↔ μ (sᶜ) = 0 := +theorem mem_ae_iff {s : Set α} : s ∈ μ.ae ↔ μ sᶜ = 0 := Iff.rfl #align measure_theory.mem_ae_iff MeasureTheory.mem_ae_iff @@ -445,7 +445,7 @@ theorem ae_eq_empty : s =ᵐ[μ] (∅ : Set α) ↔ μ s = 0 := -- Porting note: The priority should be higher than `eventuallyEq_univ`. @[simp high] -theorem ae_eq_univ : s =ᵐ[μ] (univ : Set α) ↔ μ (sᶜ) = 0 := +theorem ae_eq_univ : s =ᵐ[μ] (univ : Set α) ↔ μ sᶜ = 0 := eventuallyEq_univ #align measure_theory.ae_eq_univ MeasureTheory.ae_eq_univ diff --git a/Mathlib/MeasureTheory/Measure/MutuallySingular.lean b/Mathlib/MeasureTheory/Measure/MutuallySingular.lean index d36ca11129e70..152fff84f0014 100644 --- a/Mathlib/MeasureTheory/Measure/MutuallySingular.lean +++ b/Mathlib/MeasureTheory/Measure/MutuallySingular.lean @@ -40,7 +40,7 @@ variable {α : Type _} {m0 : MeasurableSpace α} {μ μ₁ μ₂ ν ν₁ ν₂ /-- Two measures `μ`, `ν` are said to be mutually singular if there exists a measurable set `s` such that `μ s = 0` and `ν sᶜ = 0`. -/ def MutuallySingular {_ : MeasurableSpace α} (μ ν : Measure α) : Prop := - ∃ s : Set α, MeasurableSet s ∧ μ s = 0 ∧ ν (sᶜ) = 0 + ∃ s : Set α, MeasurableSet s ∧ μ s = 0 ∧ ν sᶜ = 0 #align measure_theory.measure.mutually_singular MeasureTheory.Measure.MutuallySingular -- mathport name: measure.mutually_singular diff --git a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean index 89334076deabe..677af520601b8 100644 --- a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean @@ -123,16 +123,16 @@ theorem of_null (h : μ s = 0) : NullMeasurableSet s μ := ⟨∅, MeasurableSet.empty, ae_eq_empty.2 h⟩ #align measure_theory.null_measurable_set.of_null MeasureTheory.NullMeasurableSet.of_null -theorem compl (h : NullMeasurableSet s μ) : NullMeasurableSet (sᶜ) μ := +theorem compl (h : NullMeasurableSet s μ) : NullMeasurableSet sᶜ μ := MeasurableSet.compl h #align measure_theory.null_measurable_set.compl MeasureTheory.NullMeasurableSet.compl -theorem of_compl (h : NullMeasurableSet (sᶜ) μ) : NullMeasurableSet s μ := +theorem of_compl (h : NullMeasurableSet sᶜ μ) : NullMeasurableSet s μ := MeasurableSet.of_compl h #align measure_theory.null_measurable_set.of_compl MeasureTheory.NullMeasurableSet.of_compl @[simp] -theorem compl_iff : NullMeasurableSet (sᶜ) μ ↔ NullMeasurableSet s μ := +theorem compl_iff : NullMeasurableSet sᶜ μ ↔ NullMeasurableSet s μ := MeasurableSet.compl_iff #align measure_theory.null_measurable_set.compl_iff MeasureTheory.NullMeasurableSet.compl_iff @@ -241,13 +241,13 @@ theorem toMeasurable_ae_eq (h : NullMeasurableSet s μ) : toMeasurable μ s =ᵐ exact (exists_measurable_superset_ae_eq h).choose_spec.snd.2 #align measure_theory.null_measurable_set.to_measurable_ae_eq MeasureTheory.NullMeasurableSet.toMeasurable_ae_eq -theorem compl_toMeasurable_compl_ae_eq (h : NullMeasurableSet s μ) : toMeasurable μ (sᶜ)ᶜ =ᵐ[μ] s := +theorem compl_toMeasurable_compl_ae_eq (h : NullMeasurableSet s μ) : (toMeasurable μ sᶜ)ᶜ =ᵐ[μ] s := Iff.mpr ae_eq_set_compl <| toMeasurable_ae_eq h.compl #align measure_theory.null_measurable_set.compl_to_measurable_compl_ae_eq MeasureTheory.NullMeasurableSet.compl_toMeasurable_compl_ae_eq theorem exists_measurable_subset_ae_eq (h : NullMeasurableSet s μ) : ∃ (t : _) (_ : t ⊆ s), MeasurableSet t ∧ t =ᵐ[μ] s := - ⟨toMeasurable μ (sᶜ)ᶜ, compl_subset_comm.2 <| subset_toMeasurable _ _, + ⟨(toMeasurable μ sᶜ)ᶜ, compl_subset_comm.2 <| subset_toMeasurable _ _, (measurableSet_toMeasurable _ _).compl, compl_toMeasurable_compl_ae_eq h⟩ #align measure_theory.null_measurable_set.exists_measurable_subset_ae_eq MeasureTheory.NullMeasurableSet.exists_measurable_subset_ae_eq @@ -341,7 +341,7 @@ theorem measure_union₀' (hs : NullMeasurableSet s μ) (hd : AEDisjoint μ s t) #align measure_theory.measure_union₀' MeasureTheory.measure_union₀' theorem measure_add_measure_compl₀ {s : Set α} (hs : NullMeasurableSet s μ) : - μ s + μ (sᶜ) = μ univ := by rw [← measure_union₀' hs aedisjoint_compl_right, union_compl_self] + μ s + μ sᶜ = μ univ := by rw [← measure_union₀' hs aedisjoint_compl_right, union_compl_self] #align measure_theory.measure_add_measure_compl₀ MeasureTheory.measure_add_measure_compl₀ section MeasurableSingletonClass diff --git a/Mathlib/MeasureTheory/Measure/OpenPos.lean b/Mathlib/MeasureTheory/Measure/OpenPos.lean index 328918b58303c..fd82faa5e0a13 100644 --- a/Mathlib/MeasureTheory/Measure/OpenPos.lean +++ b/Mathlib/MeasureTheory/Measure/OpenPos.lean @@ -96,7 +96,7 @@ theorem eqOn_open_of_ae_eq {f g : X → Y} (h : f =ᵐ[μ.restrict U] g) (hU : I simp only [EventuallyEq, ae_iff, not_imp] at h have : IsOpen (U ∩ { a | f a ≠ g a }) := by refine' isOpen_iff_mem_nhds.mpr fun a ha => inter_mem (hU.mem_nhds ha.1) _ - rcases ha with ⟨ha : a ∈ U, ha' : (f a, g a) ∈ diagonal Yᶜ⟩ + rcases ha with ⟨ha : a ∈ U, ha' : (f a, g a) ∈ (diagonal Y)ᶜ⟩ exact (hf.continuousAt (hU.mem_nhds ha)).prod_mk_nhds (hg.continuousAt (hU.mem_nhds ha)) (isClosed_diagonal.isOpen_compl.mem_nhds ha') diff --git a/Mathlib/MeasureTheory/Measure/OuterMeasure.lean b/Mathlib/MeasureTheory/Measure/OuterMeasure.lean index 5d02d2abf4ffd..1796a58e2e6ef 100644 --- a/Mathlib/MeasureTheory/Measure/OuterMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/OuterMeasure.lean @@ -800,7 +800,7 @@ theorem map_ofFunction {β} {f : α → β} (hf : Injective f) : refine' (map_ofFunction_le _).antisymm fun s => _ simp only [ofFunction_apply, map_apply, le_iInf_iff] intro t ht - refine' iInf_le_of_le (fun n => range fᶜ ∪ f '' t n) (iInf_le_of_le _ _) + refine' iInf_le_of_le (fun n => (range f)ᶜ ∪ f '' t n) (iInf_le_of_le _ _) · rw [← union_iUnion, ← inter_subset, ← image_preimage_eq_inter_range, ← image_iUnion] exact image_subset _ ht · refine' ENNReal.tsum_le_tsum fun n => le_of_eq _ @@ -952,12 +952,12 @@ theorem isCaratheodory_iff_le' {s : Set α} : theorem isCaratheodory_empty : IsCaratheodory m ∅ := by simp [IsCaratheodory, m.empty, diff_empty] #align measure_theory.outer_measure.is_caratheodory_empty MeasureTheory.OuterMeasure.isCaratheodory_empty -theorem isCaratheodory_compl : IsCaratheodory m s₁ → IsCaratheodory m (s₁ᶜ) := by +theorem isCaratheodory_compl : IsCaratheodory m s₁ → IsCaratheodory m s₁ᶜ := by simp [IsCaratheodory, diff_eq, add_comm] #align measure_theory.outer_measure.is_caratheodory_compl MeasureTheory.OuterMeasure.isCaratheodory_compl @[simp] -theorem isCaratheodory_compl_iff : IsCaratheodory m (sᶜ) ↔ IsCaratheodory m s := +theorem isCaratheodory_compl_iff : IsCaratheodory m sᶜ ↔ IsCaratheodory m s := ⟨fun h => by simpa using isCaratheodory_compl m h, isCaratheodory_compl m⟩ #align measure_theory.outer_measure.is_caratheodory_compl_iff MeasureTheory.OuterMeasure.isCaratheodory_compl_iff @@ -1242,7 +1242,7 @@ theorem map_iInf_comap {ι β} [Nonempty ι] {f : α → β} (m : ι → OuterMe map f (⨅ i, comap f (m i)) = ⨅ i, map f (comap f (m i)) := by refine' (map_iInf_le _ _).antisymm fun s => _ simp only [map_apply, comap_apply, iInf_apply, le_iInf_iff] - refine' fun t ht => iInf_le_of_le (fun n => f '' t n ∪ range fᶜ) (iInf_le_of_le _ _) + refine' fun t ht => iInf_le_of_le (fun n => f '' t n ∪ (range f)ᶜ) (iInf_le_of_le _ _) · rw [← iUnion_union, Set.union_comm, ← inter_subset, ← image_iUnion, ← image_preimage_eq_inter_range] exact image_subset _ ht diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 005c5faaba8eb..82dffd2576dce 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -107,12 +107,12 @@ variable {Ω : Type _} [MeasurableSpace Ω] theorem le_measure_compl_liminf_of_limsup_measure_le {ι : Type _} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} [IsProbabilityMeasure μ] [∀ i, IsProbabilityMeasure (μs i)] {E : Set Ω} (E_mble : MeasurableSet E) (h : (L.limsup fun i => μs i E) ≤ μ E) : - μ (Eᶜ) ≤ L.liminf fun i => μs i (Eᶜ) := by + μ Eᶜ ≤ L.liminf fun i => μs i Eᶜ := by rcases L.eq_or_neBot with rfl | hne · simp only [liminf_bot, le_top] - have meas_Ec : μ (Eᶜ) = 1 - μ E := by + have meas_Ec : μ Eᶜ = 1 - μ E := by simpa only [measure_univ] using measure_compl E_mble (measure_lt_top μ E).ne - have meas_i_Ec : ∀ i, μs i (Eᶜ) = 1 - μs i E := by + have meas_i_Ec : ∀ i, μs i Eᶜ = 1 - μs i E := by intro i simpa only [measure_univ] using measure_compl E_mble (measure_lt_top (μs i) E).ne simp_rw [meas_Ec, meas_i_Ec] @@ -127,7 +127,7 @@ theorem le_measure_compl_liminf_of_limsup_measure_le {ι : Type _} {L : Filter theorem le_measure_liminf_of_limsup_measure_compl_le {ι : Type _} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} [IsProbabilityMeasure μ] [∀ i, IsProbabilityMeasure (μs i)] {E : Set Ω} - (E_mble : MeasurableSet E) (h : (L.limsup fun i => μs i (Eᶜ)) ≤ μ (Eᶜ)) : + (E_mble : MeasurableSet E) (h : (L.limsup fun i => μs i Eᶜ) ≤ μ Eᶜ) : μ E ≤ L.liminf fun i => μs i E := compl_compl E ▸ le_measure_compl_liminf_of_limsup_measure_le (MeasurableSet.compl E_mble) h #align measure_theory.le_measure_liminf_of_limsup_measure_compl_le MeasureTheory.le_measure_liminf_of_limsup_measure_compl_le @@ -135,12 +135,12 @@ theorem le_measure_liminf_of_limsup_measure_compl_le {ι : Type _} {L : Filter theorem limsup_measure_compl_le_of_le_liminf_measure {ι : Type _} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} [IsProbabilityMeasure μ] [∀ i, IsProbabilityMeasure (μs i)] {E : Set Ω} (E_mble : MeasurableSet E) (h : μ E ≤ L.liminf fun i => μs i E) : - (L.limsup fun i => μs i (Eᶜ)) ≤ μ (Eᶜ) := by + (L.limsup fun i => μs i Eᶜ) ≤ μ Eᶜ := by rcases L.eq_or_neBot with rfl | hne · simp only [limsup_bot, bot_le] - have meas_Ec : μ (Eᶜ) = 1 - μ E := by + have meas_Ec : μ Eᶜ = 1 - μ E := by simpa only [measure_univ] using measure_compl E_mble (measure_lt_top μ E).ne - have meas_i_Ec : ∀ i, μs i (Eᶜ) = 1 - μs i E := by + have meas_i_Ec : ∀ i, μs i Eᶜ = 1 - μs i E := by intro i simpa only [measure_univ] using measure_compl E_mble (measure_lt_top (μs i) E).ne simp_rw [meas_Ec, meas_i_Ec] @@ -155,7 +155,7 @@ theorem limsup_measure_compl_le_of_le_liminf_measure {ι : Type _} {L : Filter theorem limsup_measure_le_of_le_liminf_measure_compl {ι : Type _} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} [IsProbabilityMeasure μ] [∀ i, IsProbabilityMeasure (μs i)] {E : Set Ω} - (E_mble : MeasurableSet E) (h : μ (Eᶜ) ≤ L.liminf fun i => μs i (Eᶜ)) : + (E_mble : MeasurableSet E) (h : μ Eᶜ ≤ L.liminf fun i => μs i Eᶜ) : (L.limsup fun i => μs i E) ≤ μ E := compl_compl E ▸ limsup_measure_compl_le_of_le_liminf_measure (MeasurableSet.compl E_mble) h #align measure_theory.limsup_measure_le_of_le_liminf_measure_compl MeasureTheory.limsup_measure_le_of_le_liminf_measure_compl @@ -179,10 +179,10 @@ theorem limsup_measure_closed_le_iff_liminf_measure_open_ge {ι : Type _} {L : F constructor · intro h G G_open exact le_measure_liminf_of_limsup_measure_compl_le - G_open.measurableSet (h (Gᶜ) (isClosed_compl_iff.mpr G_open)) + G_open.measurableSet (h Gᶜ (isClosed_compl_iff.mpr G_open)) · intro h F F_closed exact limsup_measure_le_of_le_liminf_measure_compl - F_closed.measurableSet (h (Fᶜ) (isOpen_compl_iff.mpr F_closed)) + F_closed.measurableSet (h Fᶜ (isOpen_compl_iff.mpr F_closed)) #align measure_theory.limsup_measure_closed_le_iff_liminf_measure_open_ge MeasureTheory.limsup_measure_closed_le_iff_liminf_measure_open_ge end LimsupClosedLEAndLELiminfOpen diff --git a/Mathlib/MeasureTheory/Measure/Sub.lean b/Mathlib/MeasureTheory/Measure/Sub.lean index b5e253eb0fd8a..3b044e2467b53 100644 --- a/Mathlib/MeasureTheory/Measure/Sub.lean +++ b/Mathlib/MeasureTheory/Measure/Sub.lean @@ -114,7 +114,7 @@ theorem restrict_sub_eq_restrict_sub_restrict (h_meas_s : MeasurableSet s) : intro ν' h_ν'_in rw [mem_setOf_eq] at h_ν'_in refine' ⟨ν'.restrict s, _, restrict_le_self⟩ - refine' ⟨ν' + (⊤ : Measure α).restrict (sᶜ), _, _⟩ + refine' ⟨ν' + (⊤ : Measure α).restrict sᶜ, _, _⟩ · rw [mem_setOf_eq, add_right_comm, Measure.le_iff] intro t h_meas_t repeat' rw [← measure_inter_add_diff t h_meas_s] diff --git a/Mathlib/MeasureTheory/PiSystem.lean b/Mathlib/MeasureTheory/PiSystem.lean index 379cea08b0f45..1cc2963058611 100644 --- a/Mathlib/MeasureTheory/PiSystem.lean +++ b/Mathlib/MeasureTheory/PiSystem.lean @@ -547,7 +547,7 @@ structure DynkinSystem (α : Type _) where /-- A Dynkin system contains the empty set. -/ has_empty : Has ∅ /-- A Dynkin system is closed under complementation. -/ - has_compl : ∀ {a}, Has a → Has (aᶜ) + has_compl : ∀ {a}, Has a → Has aᶜ /-- A Dynkin system is closed under countable union of pairwise disjoint sets. Use a more general `MeasurableSpace.DynkinSystem.has_iUnion` instead.-/ has_iUnion_nat : ∀ {f : ℕ → Set α}, Pairwise (Disjoint on f) → (∀ i, Has (f i)) → Has (⋃ i, f i) @@ -565,7 +565,7 @@ theorem ext : ∀ {d₁ d₂ : DynkinSystem α}, (∀ s : Set α, d₁.Has s ↔ variable (d : DynkinSystem α) -theorem has_compl_iff {a} : d.Has (aᶜ) ↔ d.Has a := +theorem has_compl_iff {a} : d.Has aᶜ ↔ d.Has a := ⟨fun h => by simpa using d.has_compl h, fun h => d.has_compl h⟩ #align measurable_space.dynkin_system.has_compl_iff MeasurableSpace.DynkinSystem.has_compl_iff @@ -625,12 +625,12 @@ theorem ofMeasurableSpace_le_ofMeasurableSpace_iff {m₁ m₂ : MeasurableSpace inductive GenerateHas (s : Set (Set α)) : Set α → Prop | basic : ∀ t ∈ s, GenerateHas s t | empty : GenerateHas s ∅ - | compl : ∀ {a}, GenerateHas s a → GenerateHas s (aᶜ) + | compl : ∀ {a}, GenerateHas s a → GenerateHas s aᶜ | iUnion : ∀ {f : ℕ → Set α}, Pairwise (Disjoint on f) → (∀ i, GenerateHas s (f i)) → GenerateHas s (⋃ i, f i) #align measurable_space.dynkin_system.generate_has MeasurableSpace.DynkinSystem.GenerateHas -theorem generateHas_compl {C : Set (Set α)} {s : Set α} : GenerateHas C (sᶜ) ↔ GenerateHas C s := by +theorem generateHas_compl {C : Set (Set α)} {s : Set α} : GenerateHas C sᶜ ↔ GenerateHas C s := by refine' ⟨_, GenerateHas.compl⟩ intro h convert GenerateHas.compl h @@ -734,7 +734,7 @@ end DynkinSystem theorem induction_on_inter {C : Set α → Prop} {s : Set (Set α)} [m : MeasurableSpace α] (h_eq : m = generateFrom s) (h_inter : IsPiSystem s) (h_empty : C ∅) (h_basic : ∀ t ∈ s, C t) - (h_compl : ∀ t, MeasurableSet t → C t → C (tᶜ)) + (h_compl : ∀ t, MeasurableSet t → C t → C tᶜ) (h_union : ∀ f : ℕ → Set α, Pairwise (Disjoint on f) → (∀ i, MeasurableSet (f i)) → (∀ i, C (f i)) → C (⋃ i, f i)) : diff --git a/Mathlib/ModelTheory/Definability.lean b/Mathlib/ModelTheory/Definability.lean index b570c2ed0edfa..f5bde2398e01b 100644 --- a/Mathlib/ModelTheory/Definability.lean +++ b/Mathlib/ModelTheory/Definability.lean @@ -138,7 +138,7 @@ theorem definable_finset_biUnion {ι : Type _} {f : ∀ _ : ι, Set (α → M)} #align set.definable_finset_bUnion Set.definable_finset_biUnion @[simp] -theorem Definable.compl {s : Set (α → M)} (hf : A.Definable L s) : A.Definable L (sᶜ) := by +theorem Definable.compl {s : Set (α → M)} (hf : A.Definable L s) : A.Definable L sᶜ := by rcases hf with ⟨φ, hφ⟩ refine' ⟨φ.not, _⟩ ext v @@ -196,7 +196,7 @@ theorem Definable.image_comp_embedding {s : Set (β → M)} (h : A.Definable L s (congr rfl (ext fun x => _)).mp (((h.image_comp_equiv (Equiv.Set.sumCompl (range f))).image_comp_equiv (Equiv.sumCongr (Equiv.ofInjective f f.injective) - (Fintype.equivFin (↥((range f)ᶜ))).symm)).image_comp_sum_inl_fin + (Fintype.equivFin (↥(range f)ᶜ)).symm)).image_comp_sum_inl_fin _) simp only [mem_preimage, mem_image, exists_exists_and_eq_and] refine' exists_congr fun y => and_congr_right fun _ => Eq.congr_left (funext fun a => _) diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index 95f7529f81060..94ea8a0c17b5c 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -1116,7 +1116,8 @@ def nonemptyTheory : L.Theory := /-- A theory indicating that each of a set of constants is distinct. -/ def distinctConstantsTheory (s : Set α) : L[[α]].Theory := - (fun ab : α × α => ((L.con ab.1).term.equal (L.con ab.2).term).not) '' (s ×ˢ s ∩ Set.diagonal αᶜ) + (fun ab : α × α => ((L.con ab.1).term.equal (L.con ab.2).term).not) '' + (s ×ˢ s ∩ (Set.diagonal α)ᶜ) #align first_order.language.distinct_constants_theory FirstOrder.Language.distinctConstantsTheory variable {L} diff --git a/Mathlib/Order/Antichain.lean b/Mathlib/Order/Antichain.lean index 2dd07e6c9c3e7..c9529dbce5b0c 100644 --- a/Mathlib/Order/Antichain.lean +++ b/Mathlib/Order/Antichain.lean @@ -35,13 +35,13 @@ section General variable {α β : Type _} {r r₁ r₂ : α → α → Prop} {r' : β → β → Prop} {s t : Set α} {a b : α} -protected theorem Symmetric.compl (h : Symmetric r) : Symmetric (rᶜ) := fun _ _ hr hr' => +protected theorem Symmetric.compl (h : Symmetric r) : Symmetric rᶜ := fun _ _ hr hr' => hr <| h hr' #align symmetric.compl Symmetric.compl /-- An antichain is a set such that no two distinct elements are related. -/ def IsAntichain (r : α → α → Prop) (s : Set α) : Prop := - s.Pairwise (rᶜ) + s.Pairwise rᶜ #align is_antichain IsAntichain namespace IsAntichain diff --git a/Mathlib/Order/Atoms.lean b/Mathlib/Order/Atoms.lean index ebc49819c2afc..8544a76ae28c8 100644 --- a/Mathlib/Order/Atoms.lean +++ b/Mathlib/Order/Atoms.lean @@ -907,7 +907,7 @@ theorem isCoatom_iff (s : Set α) : IsCoatom s ↔ ∃ x, s = {x}ᶜ := by #align set.is_coatom_iff Set.isCoatom_iff theorem isCoatom_singleton_compl (x : α) : IsCoatom ({x}ᶜ : Set α) := - (isCoatom_iff ({x}ᶜ)).mpr ⟨x, rfl⟩ + (isCoatom_iff {x}ᶜ).mpr ⟨x, rfl⟩ #align set.is_coatom_singleton_compl Set.isCoatom_singleton_compl instance : IsAtomistic (Set α) where diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 22edc9b081e3b..66649d7ba67f0 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -757,32 +757,32 @@ class HasCompl (α : Type _) where export HasCompl (compl) @[inherit_doc] -postfix:999 "ᶜ" => compl +postfix:1024 "ᶜ" => compl instance Prop.hasCompl : HasCompl Prop := ⟨Not⟩ #align Prop.has_compl Prop.hasCompl instance Pi.hasCompl {ι : Type u} {α : ι → Type v} [∀ i, HasCompl (α i)] : HasCompl (∀ i, α i) := - ⟨fun x i ↦ x iᶜ⟩ + ⟨fun x i ↦ (x i)ᶜ⟩ #align pi.has_compl Pi.hasCompl theorem Pi.compl_def {ι : Type u} {α : ι → Type v} [∀ i, HasCompl (α i)] (x : ∀ i, α i) : - xᶜ = fun i ↦ x iᶜ := + xᶜ = fun i ↦ (x i)ᶜ := rfl #align pi.compl_def Pi.compl_def @[simp] theorem Pi.compl_apply {ι : Type u} {α : ι → Type v} [∀ i, HasCompl (α i)] (x : ∀ i, α i) (i : ι) : - (xᶜ) i = x iᶜ := + xᶜ i = (x i)ᶜ := rfl #align pi.compl_apply Pi.compl_apply -instance IsIrrefl.compl (r) [IsIrrefl α r] : IsRefl α (rᶜ) := +instance IsIrrefl.compl (r) [IsIrrefl α r] : IsRefl α rᶜ := ⟨@irrefl α r _⟩ #align is_irrefl.compl IsIrrefl.compl -instance IsRefl.compl (r) [IsRefl α r] : IsIrrefl α (rᶜ) := +instance IsRefl.compl (r) [IsRefl α r] : IsIrrefl α rᶜ := ⟨fun a ↦ not_not_intro (refl a)⟩ #align is_refl.compl IsRefl.compl diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index 62cc63f4e9f35..223a6646deb2e 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -569,7 +569,7 @@ theorem compl_sup_eq_top : xᶜ ⊔ x = ⊤ := sup_comm.trans sup_compl_eq_top #align compl_sup_eq_top compl_sup_eq_top -theorem isCompl_compl : IsCompl x (xᶜ) := +theorem isCompl_compl : IsCompl x xᶜ := IsCompl.of_eq inf_compl_eq_bot' sup_compl_eq_top #align is_compl_compl isCompl_compl @@ -732,11 +732,11 @@ theorem compl_himp_compl : xᶜ ⇨ yᶜ = y ⇨ x := @compl_sdiff_compl αᵒᵈ _ _ _ #align compl_himp_compl compl_himp_compl -theorem disjoint_compl_left_iff : Disjoint (xᶜ) y ↔ y ≤ x := by +theorem disjoint_compl_left_iff : Disjoint xᶜ y ↔ y ≤ x := by rw [← le_compl_iff_disjoint_left, compl_compl] #align disjoint_compl_left_iff disjoint_compl_left_iff -theorem disjoint_compl_right_iff : Disjoint x (yᶜ) ↔ x ≤ y := by +theorem disjoint_compl_right_iff : Disjoint x yᶜ ↔ x ≤ y := by rw [← le_compl_iff_disjoint_right, compl_compl] #align disjoint_compl_right_iff disjoint_compl_right_iff @@ -823,7 +823,7 @@ protected def Function.Injective.generalizedBooleanAlgebra [Sup α] [Inf α] [Bo protected def Function.Injective.booleanAlgebra [Sup α] [Inf α] [Top α] [Bot α] [HasCompl α] [SDiff α] [BooleanAlgebra β] (f : α → β) (hf : Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) - (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) (map_compl : ∀ a, f (aᶜ) = f aᶜ) + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : BooleanAlgebra α := { hf.generalizedBooleanAlgebra f map_sup map_inf map_bot map_sdiff with compl := compl, diff --git a/Mathlib/Order/Circular.lean b/Mathlib/Order/Circular.lean index 99f338fb2bc08..6ed9658c64e9f 100644 --- a/Mathlib/Order/Circular.lean +++ b/Mathlib/Order/Circular.lean @@ -369,13 +369,13 @@ theorem right_mem_cIcc (a b : α) : b ∈ cIcc a b := btw_rfl_right #align set.right_mem_cIcc Set.right_mem_cIcc -theorem compl_cIcc {a b : α} : cIcc a bᶜ = cIoo b a := by +theorem compl_cIcc {a b : α} : (cIcc a b)ᶜ = cIoo b a := by ext rw [Set.mem_cIoo, sbtw_iff_not_btw] rfl #align set.compl_cIcc Set.compl_cIcc -theorem compl_cIoo {a b : α} : cIoo a bᶜ = cIcc b a := by +theorem compl_cIoo {a b : α} : (cIoo a b)ᶜ = cIcc b a := by ext rw [Set.mem_cIcc, btw_iff_not_sbtw] rfl diff --git a/Mathlib/Order/CompleteBooleanAlgebra.lean b/Mathlib/Order/CompleteBooleanAlgebra.lean index 65b53a3e4eae7..075c17522b8ba 100644 --- a/Mathlib/Order/CompleteBooleanAlgebra.lean +++ b/Mathlib/Order/CompleteBooleanAlgebra.lean @@ -323,28 +323,28 @@ section CompleteBooleanAlgebra variable [CompleteBooleanAlgebra α] {a b : α} {s : Set α} {f : ι → α} -theorem compl_iInf : iInf fᶜ = ⨆ i, f iᶜ := +theorem compl_iInf : (iInf f)ᶜ = ⨆ i, (f i)ᶜ := le_antisymm (compl_le_of_compl_le <| le_iInf fun i => compl_le_of_compl_le <| le_iSup (HasCompl.compl ∘ f) i) (iSup_le fun _ => compl_le_compl <| iInf_le _ _) #align compl_infi compl_iInf -theorem compl_iSup : iSup fᶜ = ⨅ i, f iᶜ := +theorem compl_iSup : (iSup f)ᶜ = ⨅ i, (f i)ᶜ := compl_injective (by simp [compl_iInf]) #align compl_supr compl_iSup -theorem compl_sInf : sInf sᶜ = ⨆ i ∈ s, iᶜ := by simp only [sInf_eq_iInf, compl_iInf] +theorem compl_sInf : (sInf s)ᶜ = ⨆ i ∈ s, iᶜ := by simp only [sInf_eq_iInf, compl_iInf] #align compl_Inf compl_sInf -theorem compl_sSup : sSup sᶜ = ⨅ i ∈ s, iᶜ := by simp only [sSup_eq_iSup, compl_iSup] +theorem compl_sSup : (sSup s)ᶜ = ⨅ i ∈ s, iᶜ := by simp only [sSup_eq_iSup, compl_iSup] #align compl_Sup compl_sSup -theorem compl_sInf' : sInf sᶜ = sSup (HasCompl.compl '' s) := +theorem compl_sInf' : (sInf s)ᶜ = sSup (HasCompl.compl '' s) := compl_sInf.trans sSup_image.symm #align compl_Inf' compl_sInf' -theorem compl_sSup' : sSup sᶜ = sInf (HasCompl.compl '' s) := +theorem compl_sSup' : (sSup s)ᶜ = sInf (HasCompl.compl '' s) := compl_sSup.trans sInf_image.symm #align compl_Sup' compl_sSup' @@ -404,7 +404,7 @@ protected def Function.Injective.completeBooleanAlgebra [Sup α] [Inf α] [SupSe (hf : Function.Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_sSup : ∀ s, f (sSup s) = ⨆ a ∈ s, f a) (map_sInf : ∀ s, f (sInf s) = ⨅ a ∈ s, f a) (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) - (map_compl : ∀ a, f (aᶜ) = f aᶜ) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : + (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : CompleteBooleanAlgebra α := { hf.completeDistribLattice f map_sup map_inf map_sSup map_sInf map_top map_bot, hf.booleanAlgebra f map_sup map_inf map_top map_bot map_compl map_sdiff with } diff --git a/Mathlib/Order/Disjointed.lean b/Mathlib/Order/Disjointed.lean index 41d6dfb31ae92..baf5403130785 100644 --- a/Mathlib/Order/Disjointed.lean +++ b/Mathlib/Order/Disjointed.lean @@ -147,7 +147,7 @@ theorem iSup_disjointed (f : ℕ → α) : (⨆ n, disjointed f n) = ⨆ n, f n iSup_eq_iSup_of_partialSups_eq_partialSups (partialSups_disjointed f) #align supr_disjointed iSup_disjointed -theorem disjointed_eq_inf_compl (f : ℕ → α) (n : ℕ) : disjointed f n = f n ⊓ ⨅ i < n, f iᶜ := by +theorem disjointed_eq_inf_compl (f : ℕ → α) (n : ℕ) : disjointed f n = f n ⊓ ⨅ i < n, (f i)ᶜ := by cases n · rw [disjointed_zero, eq_comm, inf_eq_left] simp_rw [le_iInf_iff] @@ -171,7 +171,8 @@ theorem iUnion_disjointed {f : ℕ → Set α} : (⋃ n, disjointed f n) = ⋃ n iSup_disjointed f #align Union_disjointed iUnion_disjointed -theorem disjointed_eq_inter_compl (f : ℕ → Set α) (n : ℕ) : disjointed f n = f n ∩ ⋂ i < n, f iᶜ := +theorem disjointed_eq_inter_compl (f : ℕ → Set α) (n : ℕ) : + disjointed f n = f n ∩ ⋂ i < n, (f i)ᶜ := disjointed_eq_inf_compl f n #align disjointed_eq_inter_compl disjointed_eq_inter_compl diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index b743b4c7373ce..58d0d5692435a 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -683,14 +683,14 @@ theorem inf_principal_neBot_iff {s : Set α} : NeBot (l ⊓ 𝓟 s) ↔ ∀ U l.basis_sets.inf_principal_neBot_iff #align filter.inf_principal_ne_bot_iff Filter.inf_principal_neBot_iff -theorem mem_iff_inf_principal_compl {f : Filter α} {s : Set α} : s ∈ f ↔ f ⊓ 𝓟 (sᶜ) = ⊥ := by +theorem mem_iff_inf_principal_compl {f : Filter α} {s : Set α} : s ∈ f ↔ f ⊓ 𝓟 sᶜ = ⊥ := by refine' not_iff_not.1 ((inf_principal_neBot_iff.trans _).symm.trans neBot_iff) exact ⟨fun h hs => by simpa [Set.not_nonempty_empty] using h s hs, fun hs t ht => inter_compl_nonempty_iff.2 fun hts => hs <| mem_of_superset ht hts⟩ #align filter.mem_iff_inf_principal_compl Filter.mem_iff_inf_principal_compl -theorem not_mem_iff_inf_principal_compl {f : Filter α} {s : Set α} : s ∉ f ↔ NeBot (f ⊓ 𝓟 (sᶜ)) := +theorem not_mem_iff_inf_principal_compl {f : Filter α} {s : Set α} : s ∉ f ↔ NeBot (f ⊓ 𝓟 sᶜ) := (not_congr mem_iff_inf_principal_compl).trans neBot_iff.symm #align filter.not_mem_iff_inf_principal_compl Filter.not_mem_iff_inf_principal_compl @@ -718,24 +718,24 @@ theorem disjoint_pure_pure {x y : α} : Disjoint (pure x : Filter α) (pure y) #align filter.disjoint_pure_pure Filter.disjoint_pure_pure @[simp] -theorem compl_diagonal_mem_prod {l₁ l₂ : Filter α} : diagonal αᶜ ∈ l₁ ×ˢ l₂ ↔ Disjoint l₁ l₂ := by +theorem compl_diagonal_mem_prod {l₁ l₂ : Filter α} : (diagonal α)ᶜ ∈ l₁ ×ˢ l₂ ↔ Disjoint l₁ l₂ := by simp only [mem_prod_iff, Filter.disjoint_iff, prod_subset_compl_diagonal_iff_disjoint] #align filter.compl_diagonal_mem_prod Filter.compl_diagonal_mem_prod -- porting note: use `∃ i, p i ∧ _` instead of `∃ i (hi : p i), _`. theorem HasBasis.disjoint_iff_left (h : l.HasBasis p s) : - Disjoint l l' ↔ ∃ i, p i ∧ s iᶜ ∈ l' := by + Disjoint l l' ↔ ∃ i, p i ∧ (s i)ᶜ ∈ l' := by simp only [h.disjoint_iff l'.basis_sets, id, ← disjoint_principal_left, (hasBasis_principal _).disjoint_iff l'.basis_sets, true_and, Unique.exists_iff] #align filter.has_basis.disjoint_iff_left Filter.HasBasis.disjoint_iff_leftₓ -- porting note: use `∃ i, p i ∧ _` instead of `∃ i (hi : p i), _`. theorem HasBasis.disjoint_iff_right (h : l.HasBasis p s) : - Disjoint l' l ↔ ∃ i, p i ∧ s iᶜ ∈ l' := + Disjoint l' l ↔ ∃ i, p i ∧ (s i)ᶜ ∈ l' := disjoint_comm.trans h.disjoint_iff_left #align filter.has_basis.disjoint_iff_right Filter.HasBasis.disjoint_iff_rightₓ -theorem le_iff_forall_inf_principal_compl {f g : Filter α} : f ≤ g ↔ ∀ V ∈ g, f ⊓ 𝓟 (Vᶜ) = ⊥ := +theorem le_iff_forall_inf_principal_compl {f g : Filter α} : f ≤ g ↔ ∀ V ∈ g, f ⊓ 𝓟 Vᶜ = ⊥ := forall₂_congr fun _ _ => mem_iff_inf_principal_compl #align filter.le_iff_forall_inf_principal_compl Filter.le_iff_forall_inf_principal_compl diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index ff5a59ad20b65..dde80f9602ad3 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -994,7 +994,7 @@ theorem principal_neBot_iff {s : Set α} : NeBot (𝓟 s) ↔ s.Nonempty := alias principal_neBot_iff ↔ _ _root_.Set.Nonempty.principal_neBot #align set.nonempty.principal_ne_bot Set.Nonempty.principal_neBot -theorem isCompl_principal (s : Set α) : IsCompl (𝓟 s) (𝓟 (sᶜ)) := +theorem isCompl_principal (s : Set α) : IsCompl (𝓟 s) (𝓟 sᶜ) := IsCompl.of_eq (by rw [inf_principal, inter_compl_self, principal_empty]) <| by rw [sup_principal, union_compl_self, principal_univ] #align filter.is_compl_principal Filter.isCompl_principal @@ -1019,13 +1019,13 @@ theorem inf_principal_eq_bot {f : Filter α} {s : Set α} : f ⊓ 𝓟 s = ⊥ rfl #align filter.inf_principal_eq_bot Filter.inf_principal_eq_bot -theorem mem_of_eq_bot {f : Filter α} {s : Set α} (h : f ⊓ 𝓟 (sᶜ) = ⊥) : s ∈ f := by +theorem mem_of_eq_bot {f : Filter α} {s : Set α} (h : f ⊓ 𝓟 sᶜ = ⊥) : s ∈ f := by rwa [inf_principal_eq_bot, compl_compl] at h #align filter.mem_of_eq_bot Filter.mem_of_eq_bot theorem diff_mem_inf_principal_compl {f : Filter α} {s : Set α} (hs : s ∈ f) (t : Set α) : - s \ t ∈ f ⊓ 𝓟 (tᶜ) := - inter_mem_inf hs <| mem_principal_self (tᶜ) + s \ t ∈ f ⊓ 𝓟 tᶜ := + inter_mem_inf hs <| mem_principal_self tᶜ #align filter.diff_mem_inf_principal_compl Filter.diff_mem_inf_principal_compl theorem principal_le_iff {s : Set α} {f : Filter α} : 𝓟 s ≤ f ↔ ∀ V ∈ f, s ⊆ V := by @@ -2368,11 +2368,12 @@ theorem comap_neBot_iff_frequently {f : Filter β} {m : α → β} : simp only [comap_neBot_iff, frequently_iff, mem_range, @and_comm (_ ∈ _), exists_exists_eq_and] #align filter.comap_ne_bot_iff_frequently Filter.comap_neBot_iff_frequently -theorem comap_neBot_iff_compl_range {f : Filter β} {m : α → β} : NeBot (comap m f) ↔ range mᶜ ∉ f := +theorem comap_neBot_iff_compl_range {f : Filter β} {m : α → β} : + NeBot (comap m f) ↔ (range m)ᶜ ∉ f := comap_neBot_iff_frequently #align filter.comap_ne_bot_iff_compl_range Filter.comap_neBot_iff_compl_range -theorem comap_eq_bot_iff_compl_range {f : Filter β} {m : α → β} : comap m f = ⊥ ↔ range mᶜ ∈ f := +theorem comap_eq_bot_iff_compl_range {f : Filter β} {m : α → β} : comap m f = ⊥ ↔ (range m)ᶜ ∈ f := not_iff_not.mp <| neBot_iff.symm.trans comap_neBot_iff_compl_range #align filter.comap_eq_bot_iff_compl_range Filter.comap_eq_bot_iff_compl_range @@ -3122,7 +3123,7 @@ protected theorem Tendsto.if' {α β : Type _} {l₁ : Filter α} {l₂ : Filter #align filter.tendsto.if' Filter.Tendsto.if' protected theorem Tendsto.piecewise {l₁ : Filter α} {l₂ : Filter β} {f g : α → β} {s : Set α} - [∀ x, Decidable (x ∈ s)] (h₀ : Tendsto f (l₁ ⊓ 𝓟 s) l₂) (h₁ : Tendsto g (l₁ ⊓ 𝓟 (sᶜ)) l₂) : + [∀ x, Decidable (x ∈ s)] (h₀ : Tendsto f (l₁ ⊓ 𝓟 s) l₂) (h₁ : Tendsto g (l₁ ⊓ 𝓟 sᶜ) l₂) : Tendsto (piecewise s f g) l₁ l₂ := Tendsto.if h₀ h₁ #align filter.tendsto.piecewise Filter.Tendsto.piecewise diff --git a/Mathlib/Order/Filter/Cofinite.lean b/Mathlib/Order/Filter/Cofinite.lean index c16f1ddbd40d9..c48908115c1f9 100644 --- a/Mathlib/Order/Filter/Cofinite.lean +++ b/Mathlib/Order/Filter/Cofinite.lean @@ -96,7 +96,7 @@ theorem eventually_cofinite_ne (x : α) : ∀ᶠ a in cofinite, a ≠ x := theorem le_cofinite_iff_compl_singleton_mem : l ≤ cofinite ↔ ∀ x, {x}ᶜ ∈ l := by refine' ⟨fun h x => h (finite_singleton x).compl_mem_cofinite, fun h s (hs : sᶜ.Finite) => _⟩ - rw [← compl_compl s, ← biUnion_of_singleton (sᶜ), compl_iUnion₂, Filter.biInter_mem hs] + rw [← compl_compl s, ← biUnion_of_singleton sᶜ, compl_iUnion₂, Filter.biInter_mem hs] exact fun x _ => h x #align filter.le_cofinite_iff_compl_singleton_mem Filter.le_cofinite_iff_compl_singleton_mem diff --git a/Mathlib/Order/Filter/Prod.lean b/Mathlib/Order/Filter/Prod.lean index 65b2c7a32923b..ac99e69f28c9b 100644 --- a/Mathlib/Order/Filter/Prod.lean +++ b/Mathlib/Order/Filter/Prod.lean @@ -515,7 +515,7 @@ theorem coprod_neBot_right [NeBot g] [Nonempty α] : (f.coprod g).NeBot := #align filter.coprod_ne_bot_right Filter.coprod_neBot_right theorem principal_coprod_principal (s : Set α) (t : Set β) : - (𝓟 s).coprod (𝓟 t) = 𝓟 ((sᶜ ×ˢ tᶜ)ᶜ) := by + (𝓟 s).coprod (𝓟 t) = 𝓟 (sᶜ ×ˢ tᶜ)ᶜ := by rw [Filter.coprod, comap_principal, comap_principal, sup_principal, Set.prod_eq, compl_inter, preimage_compl, preimage_compl, compl_compl, compl_compl] #align filter.principal_coprod_principal Filter.principal_coprod_principal diff --git a/Mathlib/Order/Filter/Ultrafilter.lean b/Mathlib/Order/Filter/Ultrafilter.lean index c719c70f9559d..a868b0727bd4d 100644 --- a/Mathlib/Order/Filter/Ultrafilter.lean +++ b/Mathlib/Order/Filter/Ultrafilter.lean @@ -502,7 +502,7 @@ theorem compl_mem_hyperfilter_of_finite {s : Set α} (hf : Set.Finite s) : sᶜ alias compl_mem_hyperfilter_of_finite ← _root_.Set.Finite.compl_mem_hyperfilter #align set.finite.compl_mem_hyperfilter Set.Finite.compl_mem_hyperfilter -theorem mem_hyperfilter_of_finite_compl {s : Set α} (hf : Set.Finite (sᶜ)) : s ∈ hyperfilter α := +theorem mem_hyperfilter_of_finite_compl {s : Set α} (hf : Set.Finite sᶜ) : s ∈ hyperfilter α := compl_compl s ▸ hf.compl_mem_hyperfilter #align filter.mem_hyperfilter_of_finite_compl Filter.mem_hyperfilter_of_finite_compl diff --git a/Mathlib/Order/Heyting/Basic.lean b/Mathlib/Order/Heyting/Basic.lean index 495fa936d0676..4ab298de0bb81 100644 --- a/Mathlib/Order/Heyting/Basic.lean +++ b/Mathlib/Order/Heyting/Basic.lean @@ -840,19 +840,19 @@ alias le_compl_comm ← le_compl_iff_le_compl alias le_compl_comm ↔ le_compl_of_le_compl _ #align le_compl_of_le_compl le_compl_of_le_compl -theorem disjoint_compl_left : Disjoint (aᶜ) a := +theorem disjoint_compl_left : Disjoint aᶜ a := disjoint_iff_inf_le.mpr <| le_himp_iff.1 (himp_bot _).ge #align disjoint_compl_left disjoint_compl_left -theorem disjoint_compl_right : Disjoint a (aᶜ) := +theorem disjoint_compl_right : Disjoint a aᶜ := disjoint_compl_left.symm #align disjoint_compl_right disjoint_compl_right -theorem LE.le.disjoint_compl_left (h : b ≤ a) : Disjoint (aᶜ) b := +theorem LE.le.disjoint_compl_left (h : b ≤ a) : Disjoint aᶜ b := disjoint_compl_left.mono_right h #align has_le.le.disjoint_compl_left LE.le.disjoint_compl_left -theorem LE.le.disjoint_compl_right (h : a ≤ b) : Disjoint a (bᶜ) := +theorem LE.le.disjoint_compl_right (h : a ≤ b) : Disjoint a bᶜ := disjoint_compl_right.mono_left h #align has_le.le.disjoint_compl_right LE.le.disjoint_compl_right @@ -913,12 +913,12 @@ theorem compl_compl_compl (a : α) : aᶜᶜᶜ = aᶜ := #align compl_compl_compl compl_compl_compl @[simp] -theorem disjoint_compl_compl_left_iff : Disjoint (aᶜᶜ) b ↔ Disjoint a b := by +theorem disjoint_compl_compl_left_iff : Disjoint aᶜᶜ b ↔ Disjoint a b := by simp_rw [← le_compl_iff_disjoint_left, compl_compl_compl] #align disjoint_compl_compl_left_iff disjoint_compl_compl_left_iff @[simp] -theorem disjoint_compl_compl_right_iff : Disjoint a (bᶜᶜ) ↔ Disjoint a b := by +theorem disjoint_compl_compl_right_iff : Disjoint a bᶜᶜ ↔ Disjoint a b := by simp_rw [← le_compl_iff_disjoint_right, compl_compl_compl] #align disjoint_compl_compl_right_iff disjoint_compl_compl_right_iff @@ -953,12 +953,12 @@ instance : CoheytingAlgebra αᵒᵈ := top_sdiff := @himp_bot α _ } @[simp] -theorem ofDual_hnot (a : αᵒᵈ) : ofDual (¬a) = ofDual aᶜ := +theorem ofDual_hnot (a : αᵒᵈ) : ofDual (¬a) = (ofDual a)ᶜ := rfl #align of_dual_hnot ofDual_hnot @[simp] -theorem toDual_compl (a : α) : toDual (aᶜ) = ¬toDual a := +theorem toDual_compl (a : α) : toDual aᶜ = ¬toDual a := rfl #align to_dual_compl toDual_compl @@ -1133,7 +1133,7 @@ instance : HeytingAlgebra αᵒᵈ := himp_bot := @top_sdiff' α _ } @[simp] -theorem ofDual_compl (a : αᵒᵈ) : ofDual (aᶜ) = ¬ofDual a := +theorem ofDual_compl (a : αᵒᵈ) : ofDual aᶜ = ¬ofDual a := rfl #align of_dual_compl ofDual_compl @@ -1143,7 +1143,7 @@ theorem ofDual_himp (a b : αᵒᵈ) : ofDual (a ⇨ b) = ofDual b \ ofDual a := #align of_dual_himp ofDual_himp @[simp] -theorem toDual_hnot (a : α) : toDual (¬a) = toDual aᶜ := +theorem toDual_hnot (a : α) : toDual (¬a) = (toDual a)ᶜ := rfl #align to_dual_hnot toDual_hnot @@ -1261,7 +1261,7 @@ protected def Function.Injective.generalizedCoheytingAlgebra [Sup α] [Inf α] [ protected def Function.Injective.heytingAlgebra [Sup α] [Inf α] [Top α] [Bot α] [HasCompl α] [HImp α] [HeytingAlgebra β] (f : α → β) (hf : Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) - (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) (map_compl : ∀ a, f (aᶜ) = f aᶜ) + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) : HeytingAlgebra α := { hf.generalizedHeytingAlgebra f map_sup map_inf map_top map_himp, ‹Bot α›, ‹HasCompl α› with bot_le := fun a => by @@ -1294,7 +1294,7 @@ protected def Function.Injective.biheytingAlgebra [Sup α] [Inf α] [Top α] [Bo [HasCompl α] [HNot α] [HImp α] [SDiff α] [BiheytingAlgebra β] (f : α → β) (hf : Injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) - (map_compl : ∀ a, f (aᶜ) = f aᶜ) (map_hnot : ∀ a, f (¬a) = ¬f a) + (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_hnot : ∀ a, f (¬a) = ¬f a) (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : BiheytingAlgebra α := { hf.heytingAlgebra f map_sup map_inf map_top map_bot map_compl map_himp, diff --git a/Mathlib/Order/Heyting/Hom.lean b/Mathlib/Order/Heyting/Hom.lean index 92e1ff2dc3675..8ac96115e277a 100644 --- a/Mathlib/Order/Heyting/Hom.lean +++ b/Mathlib/Order/Heyting/Hom.lean @@ -190,7 +190,7 @@ section HeytingAlgebra variable [HeytingAlgebra α] [HeytingAlgebra β] [HeytingHomClass F α β] (f : F) @[simp] -theorem map_compl (a : α) : f (aᶜ) = f aᶜ := by rw [← himp_bot, ← himp_bot, map_himp, map_bot] +theorem map_compl (a : α) : f aᶜ = (f a)ᶜ := by rw [← himp_bot, ← himp_bot, map_himp, map_bot] #align map_compl map_compl @[simp] diff --git a/Mathlib/Order/Heyting/Regular.lean b/Mathlib/Order/Heyting/Regular.lean index 236f5be0e7d67..eba8b9e839bb5 100644 --- a/Mathlib/Order/Heyting/Regular.lean +++ b/Mathlib/Order/Heyting/Regular.lean @@ -74,15 +74,15 @@ theorem IsRegular.himp (ha : IsRegular a) (hb : IsRegular b) : IsRegular (a ⇨ rw [IsRegular, compl_compl_himp_distrib, ha.eq, hb.eq] #align heyting.is_regular.himp Heyting.IsRegular.himp -theorem isRegular_compl (a : α) : IsRegular (aᶜ) := +theorem isRegular_compl (a : α) : IsRegular aᶜ := compl_compl_compl _ #align heyting.is_regular_compl Heyting.isRegular_compl -protected theorem IsRegular.disjoint_compl_left_iff (ha : IsRegular a) : Disjoint (aᶜ) b ↔ b ≤ a := +protected theorem IsRegular.disjoint_compl_left_iff (ha : IsRegular a) : Disjoint aᶜ b ↔ b ≤ a := by rw [← le_compl_iff_disjoint_left, ha.eq] #align heyting.is_regular.disjoint_compl_left_iff Heyting.IsRegular.disjoint_compl_left_iff -protected theorem IsRegular.disjoint_compl_right_iff (hb : IsRegular b) : Disjoint a (bᶜ) ↔ a ≤ b := +protected theorem IsRegular.disjoint_compl_right_iff (hb : IsRegular b) : Disjoint a bᶜ ↔ a ≤ b := by rw [← le_compl_iff_disjoint_right, hb.eq] #align heyting.is_regular.disjoint_compl_right_iff Heyting.IsRegular.disjoint_compl_right_iff @@ -90,7 +90,7 @@ protected theorem IsRegular.disjoint_compl_right_iff (hb : IsRegular b) : Disjoi /-- A Heyting algebra with regular excluded middle is a boolean algebra. -/ @[reducible] def _root_.BooleanAlgebra.ofRegular (h : ∀ a : α, IsRegular (a ⊔ aᶜ)) : BooleanAlgebra α := - have : ∀ a : α, IsCompl a (aᶜ) := fun a => + have : ∀ a : α, IsCompl a aᶜ := fun a => ⟨disjoint_compl_right, codisjoint_iff.2 <| by erw [← (h a), compl_sup, inf_compl_eq_bot, compl_bot]⟩ { ‹HeytingAlgebra α›, @@ -166,7 +166,7 @@ theorem coe_himp (a b : Regular α) : (↑(a ⇨ b) : α) = (a : α) ⇨ b := #align heyting.regular.coe_himp Heyting.Regular.coe_himp @[simp, norm_cast] -theorem coe_compl (a : Regular α) : (↑(aᶜ) : α) = (a : α)ᶜ := +theorem coe_compl (a : Regular α) : (↑aᶜ : α) = (a : α)ᶜ := rfl #align heyting.regular.coe_compl Heyting.Regular.coe_compl diff --git a/Mathlib/Order/Hom/Lattice.lean b/Mathlib/Order/Hom/Lattice.lean index 6db13bbd14ab7..de871009e29bf 100644 --- a/Mathlib/Order/Hom/Lattice.lean +++ b/Mathlib/Order/Hom/Lattice.lean @@ -285,7 +285,7 @@ section BooleanAlgebra variable [BooleanAlgebra α] [BooleanAlgebra β] [BoundedLatticeHomClass F α β] (f : F) /-- Special case of `map_compl` for boolean algebras. -/ -theorem map_compl' (a : α) : f (aᶜ) = f aᶜ := +theorem map_compl' (a : α) : f aᶜ = (f a)ᶜ := (isCompl_compl.map _).compl_eq.symm #align map_compl' map_compl' diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index b311008bcd165..8df8812588e9f 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -1020,11 +1020,11 @@ section CompleteBooleanAlgebra variable [CompleteBooleanAlgebra α] (f : Filter β) (u : β → α) -theorem limsup_compl : limsup u fᶜ = liminf (compl ∘ u) f := by +theorem limsup_compl : (limsup u f)ᶜ = liminf (compl ∘ u) f := by simp only [limsup_eq_iInf_iSup, compl_iInf, compl_iSup, liminf_eq_iSup_iInf, Function.comp_apply] #align filter.limsup_compl Filter.limsup_compl -theorem liminf_compl : liminf u fᶜ = limsup (compl ∘ u) f := by +theorem liminf_compl : (liminf u f)ᶜ = limsup (compl ∘ u) f := by simp only [limsup_eq_iInf_iSup, compl_iInf, compl_iSup, liminf_eq_iSup_iInf, Function.comp_apply] #align filter.liminf_compl Filter.liminf_compl diff --git a/Mathlib/Order/PrimeIdeal.lean b/Mathlib/Order/PrimeIdeal.lean index fe6a4d042a51d..14cff8d1b4410 100644 --- a/Mathlib/Order/PrimeIdeal.lean +++ b/Mathlib/Order/PrimeIdeal.lean @@ -94,7 +94,7 @@ end PrimePair -/ @[mk_iff] class IsPrime [Preorder P] (I : Ideal P) extends IsProper I : Prop where - compl_filter : IsPFilter ((I : Set P)ᶜ) + compl_filter : IsPFilter (I : Set P)ᶜ #align order.ideal.is_prime Order.Ideal.IsPrime section Preorder @@ -220,7 +220,7 @@ variable [Preorder P] -/ @[mk_iff] class IsPrime (F : PFilter P) : Prop where - compl_ideal : IsIdeal ((F : Set P)ᶜ) + compl_ideal : IsIdeal (F : Set P)ᶜ #align order.pfilter.is_prime Order.PFilter.IsPrime /-- Create an element of type `Order.Ideal.PrimePair` from a filter satisfying the predicate diff --git a/Mathlib/Order/UpperLower/Basic.lean b/Mathlib/Order/UpperLower/Basic.lean index 5147e21dea78e..968fe3b834232 100644 --- a/Mathlib/Order/UpperLower/Basic.lean +++ b/Mathlib/Order/UpperLower/Basic.lean @@ -82,21 +82,21 @@ theorem isUpperSet_univ : IsUpperSet (univ : Set α) := fun _ _ _ => id theorem isLowerSet_univ : IsLowerSet (univ : Set α) := fun _ _ _ => id #align is_lower_set_univ isLowerSet_univ -theorem IsUpperSet.compl (hs : IsUpperSet s) : IsLowerSet (sᶜ) := fun _a _b h hb ha => hb <| hs h ha +theorem IsUpperSet.compl (hs : IsUpperSet s) : IsLowerSet sᶜ := fun _a _b h hb ha => hb <| hs h ha #align is_upper_set.compl IsUpperSet.compl -theorem IsLowerSet.compl (hs : IsLowerSet s) : IsUpperSet (sᶜ) := fun _a _b h hb ha => hb <| hs h ha +theorem IsLowerSet.compl (hs : IsLowerSet s) : IsUpperSet sᶜ := fun _a _b h hb ha => hb <| hs h ha #align is_lower_set.compl IsLowerSet.compl @[simp] -theorem isUpperSet_compl : IsUpperSet (sᶜ) ↔ IsLowerSet s := +theorem isUpperSet_compl : IsUpperSet sᶜ ↔ IsLowerSet s := ⟨fun h => by convert h.compl rw [compl_compl], IsLowerSet.compl⟩ #align is_upper_set_compl isUpperSet_compl @[simp] -theorem isLowerSet_compl : IsLowerSet (sᶜ) ↔ IsUpperSet s := +theorem isLowerSet_compl : IsLowerSet sᶜ ↔ IsUpperSet s := ⟨fun h => by convert h.compl rw [compl_compl], IsUpperSet.compl⟩ @@ -794,7 +794,7 @@ namespace UpperSet variable {s t : UpperSet α} {a : α} @[simp] -theorem coe_compl (s : UpperSet α) : (s.compl : Set α) = ↑sᶜ := +theorem coe_compl (s : UpperSet α) : (s.compl : Set α) = (↑s)ᶜ := rfl #align upper_set.coe_compl UpperSet.coe_compl @@ -870,7 +870,7 @@ namespace LowerSet variable {s t : LowerSet α} {a : α} @[simp] -theorem coe_compl (s : LowerSet α) : (s.compl : Set α) = ↑sᶜ := +theorem coe_compl (s : LowerSet α) : (s.compl : Set α) = (↑s)ᶜ := rfl #align lower_set.coe_compl LowerSet.coe_compl diff --git a/Mathlib/Probability/CondCount.lean b/Mathlib/Probability/CondCount.lean index 94b85cdbb28bd..11ab1c21b1881 100644 --- a/Mathlib/Probability/CondCount.lean +++ b/Mathlib/Probability/CondCount.lean @@ -164,7 +164,7 @@ theorem condCount_union (hs : s.Finite) (htu : Disjoint t u) : #align probability_theory.cond_count_union ProbabilityTheory.condCount_union theorem condCount_compl (t : Set Ω) (hs : s.Finite) (hs' : s.Nonempty) : - condCount s t + condCount s (tᶜ) = 1 := by + condCount s t + condCount s tᶜ = 1 := by rw [← condCount_union hs disjoint_compl_right, Set.union_compl_self, (condCount_isProbabilityMeasure hs hs').measure_univ] #align probability_theory.cond_count_compl ProbabilityTheory.condCount_compl @@ -193,7 +193,7 @@ theorem condCount_disjoint_union (hs : s.Finite) (ht : t.Finite) (hst : Disjoint /-- A version of the law of total probability for counting probabilities. -/ theorem condCount_add_compl_eq (u t : Set Ω) (hs : s.Finite) : - condCount (s ∩ u) t * condCount s u + condCount (s ∩ uᶜ) t * condCount s (uᶜ) = + condCount (s ∩ u) t * condCount s u + condCount (s ∩ uᶜ) t * condCount s uᶜ = condCount s t := by -- Porting note: The original proof used `conv_rhs`. However, that tactic timed out. have : condCount s t = (condCount (s ∩ u) t * condCount (s ∩ u ∪ s ∩ uᶜ) (s ∩ u) + diff --git a/Mathlib/Probability/ConditionalProbability.lean b/Mathlib/Probability/ConditionalProbability.lean index c905d9dbb6918..f46a87dada0e9 100644 --- a/Mathlib/Probability/ConditionalProbability.lean +++ b/Mathlib/Probability/ConditionalProbability.lean @@ -152,7 +152,7 @@ theorem cond_mul_eq_inter [IsFiniteMeasure μ] (hms : MeasurableSet s) (hcs : μ /-- A version of the law of total probability. -/ theorem cond_add_cond_compl_eq [IsFiniteMeasure μ] (hms : MeasurableSet s) (hcs : μ s ≠ 0) - (hcs' : μ (sᶜ) ≠ 0) : μ[t|s] * μ s + μ[t|sᶜ] * μ (sᶜ) = μ t := by + (hcs' : μ sᶜ ≠ 0) : μ[t|s] * μ s + μ[t|sᶜ] * μ sᶜ = μ t := by rw [cond_mul_eq_inter μ hms hcs, cond_mul_eq_inter μ hms.compl hcs', Set.inter_comm _ t, Set.inter_comm _ t] exact measure_inter_add_diff t hms diff --git a/Mathlib/Probability/Independence/Basic.lean b/Mathlib/Probability/Independence/Basic.lean index 726304c44ed9a..3409e3975a762 100644 --- a/Mathlib/Probability/Independence/Basic.lean +++ b/Mathlib/Probability/Independence/Basic.lean @@ -907,7 +907,7 @@ theorem iIndepSet.iIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β rw [iIndepFun_iff_measure_inter_preimage_eq_mul] rintro S π _hπ simp_rw [Set.indicator_const_preimage_eq_union] - refine' @hs S (fun i => ite (1 ∈ π i) (s i) ∅ ∪ ite ((0 : β) ∈ π i) (s iᶜ) ∅) fun i _hi => _ + refine' @hs S (fun i => ite (1 ∈ π i) (s i) ∅ ∪ ite ((0 : β) ∈ π i) (s i)ᶜ ∅) fun i _hi => _ have hsi : MeasurableSet[generateFrom {s i}] (s i) := measurableSet_generateFrom (Set.mem_singleton _) refine' diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index 6af28afa8d93e..c71a815401448 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -553,7 +553,7 @@ theorem hasCondCdf_ae (ρ : Measure (α × ℝ)) [IsFiniteMeasure ρ] : ∀ᵐ a /-- A measurable set of elements of `α` such that `ρ` has a conditional cdf at all `a ∈ condCdfSet`. -/ def condCdfSet (ρ : Measure (α × ℝ)) : Set α := - toMeasurable ρ.fst {b | ¬HasCondCdf ρ b}ᶜ + (toMeasurable ρ.fst {b | ¬HasCondCdf ρ b})ᶜ #align probability_theory.cond_cdf_set ProbabilityTheory.condCdfSet theorem measurableSet_condCdfSet (ρ : Measure (α × ℝ)) : MeasurableSet (condCdfSet ρ) := @@ -1010,7 +1010,7 @@ theorem measurable_measure_condCdf (ρ : Measure (α × ℝ)) : exact (measurable_condCdf ρ u).ennreal_ofReal · intro t ht ht_cd_meas have : - (fun a => (condCdf ρ a).measure (tᶜ)) = + (fun a => (condCdf ρ a).measure tᶜ) = (fun a => (condCdf ρ a).measure univ) - fun a => (condCdf ρ a).measure t := by ext1 a rw [measure_compl ht (measure_ne_top (condCdf ρ a).measure _), Pi.sub_apply] diff --git a/Mathlib/Probability/Kernel/Disintegration.lean b/Mathlib/Probability/Kernel/Disintegration.lean index 749dcc80fa7d5..892f1dfa5ce49 100644 --- a/Mathlib/Probability/Kernel/Disintegration.lean +++ b/Mathlib/Probability/Kernel/Disintegration.lean @@ -108,7 +108,7 @@ theorem set_lintegral_condKernelReal_prod {s : Set α} (hs : MeasurableSet s) {t exact set_lintegral_condKernelReal_Iic ρ q hs · intro t ht ht_lintegral calc - ∫⁻ a in s, condKernelReal ρ a (tᶜ) ∂ρ.fst = + ∫⁻ a in s, condKernelReal ρ a tᶜ ∂ρ.fst = ∫⁻ a in s, condKernelReal ρ a univ - condKernelReal ρ a t ∂ρ.fst := by congr with a; rw [measure_compl ht (measure_ne_top (condKernelReal ρ a) _)] _ = ∫⁻ a in s, condKernelReal ρ a univ ∂ρ.fst - ∫⁻ a in s, condKernelReal ρ a t ∂ρ.fst := by @@ -169,7 +169,7 @@ theorem lintegral_condKernelReal_mem {s : Set (α × ℝ)} (hs : MeasurableSet s · intro t ht ht_eq calc ∫⁻ a, condKernelReal ρ a {x : ℝ | (a, x) ∈ tᶜ} ∂ρ.fst = - ∫⁻ a, condKernelReal ρ a ({x : ℝ | (a, x) ∈ t}ᶜ) ∂ρ.fst := rfl + ∫⁻ a, condKernelReal ρ a {x : ℝ | (a, x) ∈ t}ᶜ ∂ρ.fst := rfl _ = ∫⁻ a, condKernelReal ρ a univ - condKernelReal ρ a {x : ℝ | (a, x) ∈ t} ∂ρ.fst := by congr with a : 1 exact measure_compl (measurable_prod_mk_left ht) (measure_ne_top (condKernelReal ρ a) _) @@ -183,7 +183,7 @@ theorem lintegral_condKernelReal_mem {s : Set (α × ℝ)} (hs : MeasurableSet s rw [lintegral_condKernelReal_univ] exact measure_lt_top ρ univ _ = ρ univ - ρ t := by rw [ht_eq, lintegral_condKernelReal_univ] - _ = ρ (tᶜ) := (measure_compl ht (measure_ne_top _ _)).symm + _ = ρ tᶜ := (measure_compl ht (measure_ne_top _ _)).symm · intro f hf_disj hf_meas hf_eq have h_eq : ∀ a, {x | (a, x) ∈ ⋃ i, f i} = ⋃ i, {x | (a, x) ∈ f i} := by intro a @@ -246,7 +246,7 @@ theorem ae_condKernelReal_eq_one {s : Set ℝ} (hs : MeasurableSet s) (hρ : ρ rw [kernel.const_apply] at h simp only [mem_compl_iff, mem_setOf_eq, kernel.prodMkLeft_apply'] at h filter_upwards [h] with a ha - change condKernelReal ρ a (sᶜ) = 0 at ha + change condKernelReal ρ a sᶜ = 0 at ha rwa [prob_compl_eq_zero_iff hs] at ha #align probability_theory.ae_cond_kernel_real_eq_one ProbabilityTheory.ae_condKernelReal_eq_one @@ -274,12 +274,12 @@ theorem exists_cond_kernel (γ : Type _) [MeasurableSpace γ] : -- however an issue: that `η` may not be a Markov kernel since its value is only a -- probability distribution almost everywhere with respect to `ρ.fst`, not everywhere. -- We modify it to obtain a Markov kernel which is almost everywhere equal. - let ρ_set := toMeasurable ρ.fst ({a | condKernelReal ρ' a (range f) = 1}ᶜ)ᶜ + let ρ_set := (toMeasurable ρ.fst {a | condKernelReal ρ' a (range f) = 1}ᶜ)ᶜ have hm : MeasurableSet ρ_set := (measurableSet_toMeasurable _ _).compl have h_eq_one_of_mem : ∀ a ∈ ρ_set, condKernelReal ρ' a (range f) = 1 := by intro a ha rw [mem_compl_iff] at ha - have h_ss := subset_toMeasurable ρ.fst ({a : α | condKernelReal ρ' a (range f) = 1}ᶜ) + have h_ss := subset_toMeasurable ρ.fst {a : α | condKernelReal ρ' a (range f) = 1}ᶜ suffices ha' : a ∉ {a : α | condKernelReal ρ' a (range f) = 1}ᶜ · rwa [not_mem_compl_iff] at ha' exact not_mem_subset h_ss ha diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 48b9fdb57bcfb..3a106a7228a4e 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -822,8 +822,8 @@ theorem prime_rename_iff (s : Set σ) {p : MvPolynomial s R} : classical symm let eqv := - (sumAlgEquiv R (↥(sᶜ)) s).symm.trans - (renameEquiv R <| (Equiv.sumComm (↥(sᶜ)) s).trans <| Equiv.Set.sumCompl s) + (sumAlgEquiv R (↥sᶜ) s).symm.trans + (renameEquiv R <| (Equiv.sumComm (↥sᶜ) s).trans <| Equiv.Set.sumCompl s) have : (rename (↑)).toRingHom = eqv.toAlgHom.toRingHom.comp C := by apply ringHom_ext · intro @@ -833,8 +833,8 @@ theorem prime_rename_iff (s : Set σ) {p : MvPolynomial s R} : dsimp erw [iterToSum_C_X, rename_X, rename_X] rfl - rw [← @prime_C_iff (MvPolynomial s R) (↥(sᶜ)) instCommRingMvPolynomialToCommSemiring p] - rw [@MulEquiv.prime_iff (MvPolynomial (↑(sᶜ)) (MvPolynomial (↑s) R)) (MvPolynomial σ R) (_) (_)] + rw [← @prime_C_iff (MvPolynomial s R) (↥sᶜ) instCommRingMvPolynomialToCommSemiring p] + rw [@MulEquiv.prime_iff (MvPolynomial ↑sᶜ (MvPolynomial ↑s R)) (MvPolynomial σ R) (_) (_)] rotate_left exact eqv.toMulEquiv convert Iff.rfl diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 3fce3863100e1..737811af0cb48 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -679,10 +679,10 @@ instance canonicallyOrderedCommSemiring : CanonicallyOrderedCommSemiring Cardina add_le_add_left := fun a b => add_le_add_left exists_add_of_le := fun {a b} => inductionOn₂ a b fun α β ⟨⟨f, hf⟩⟩ => - have : Sum α (range fᶜ : Set β) ≃ β := + have : Sum α ((range f)ᶜ : Set β) ≃ β := (Equiv.sumCongr (Equiv.ofInjective f hf) (Equiv.refl _)).trans <| Equiv.Set.sumCompl (range f) - ⟨#↥(range fᶜ), mk_congr this.symm⟩ + ⟨#↥(range f)ᶜ, mk_congr this.symm⟩ le_self_add := fun a b => (add_zero a).ge.trans <| add_le_add_left (Cardinal.zero_le _) _ eq_zero_or_eq_zero_of_mul_eq_zero := fun {a b} => inductionOn₂ a b fun α β => by diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 30697af15710f..2e4b0e84e1d00 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -128,7 +128,7 @@ def StrictOrder.cof (r : α → α → Prop) : Cardinal := /-- The set in the definition of `Order.StrictOrder.cof` is nonempty. -/ theorem StrictOrder.cof_nonempty (r : α → α → Prop) [IsIrrefl α r] : { c | ∃ S : Set α, Unbounded r S ∧ (#S) = c }.Nonempty := - @Order.cof_nonempty α _ (IsRefl.swap (rᶜ)) + @Order.cof_nonempty α _ (IsRefl.swap rᶜ) #align strict_order.cof_nonempty StrictOrder.cof_nonempty /-! ### Cofinality of ordinals -/ diff --git a/Mathlib/SetTheory/Cardinal/Ordinal.lean b/Mathlib/SetTheory/Cardinal/Ordinal.lean index c77b9d33a3884..4ed826a723bbe 100644 --- a/Mathlib/SetTheory/Cardinal/Ordinal.lean +++ b/Mathlib/SetTheory/Cardinal/Ordinal.lean @@ -1153,7 +1153,7 @@ theorem mk_compl_of_infinite {α : Type _} [Infinite α] (s : Set α) (h2 : (#s) #align cardinal.mk_compl_of_infinite Cardinal.mk_compl_of_infinite theorem mk_compl_finset_of_infinite {α : Type _} [Infinite α] (s : Finset α) : - (#(↑sᶜ : Set α)) = (#α) := by + (#((↑s)ᶜ : Set α)) = (#α) := by apply mk_compl_of_infinite exact (finset_card_lt_aleph0 s).trans_le (aleph0_le_mk α) #align cardinal.mk_compl_finset_of_infinite Cardinal.mk_compl_finset_of_infinite @@ -1194,7 +1194,7 @@ theorem mk_compl_eq_mk_compl_finite_same {α : Type u} [Finite α] {s t : Set α theorem extend_function {α β : Type _} {s : Set α} (f : s ↪ β) - (h : Nonempty ((sᶜ : Set α) ≃ (range fᶜ : Set β))) : ∃ g : α ≃ β, ∀ x : s, g x = f x := by + (h : Nonempty ((sᶜ : Set α) ≃ ((range f)ᶜ : Set β))) : ∃ g : α ≃ β, ∀ x : s, g x = f x := by intros ; have := h; cases' this with g let h : α ≃ β := (Set.sumCompl (s : Set α)).symm.trans diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 2ee626c805eef..36a231cc57a84 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -1710,7 +1710,7 @@ theorem lsub_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : h.not_lt (lt_lsub f i) #align ordinal.lsub_not_mem_range Ordinal.lsub_not_mem_range -theorem nonempty_compl_range {ι : Type u} (f : ι → Ordinal.{max u v}) : Set.range fᶜ.Nonempty := +theorem nonempty_compl_range {ι : Type u} (f : ι → Ordinal.{max u v}) : (Set.range f)ᶜ.Nonempty := ⟨_, lsub_not_mem_range.{_, v} f⟩ #align ordinal.nonempty_compl_range Ordinal.nonempty_compl_range @@ -2005,7 +2005,7 @@ theorem IsNormal.eq_iff_zero_and_succ {f g : Ordinal.{u} → Ordinal.{u}} (hf : /-- The minimum excluded ordinal in a family of ordinals. -/ def mex {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal := - sInf (Set.range fᶜ) + sInf (Set.range f)ᶜ #align ordinal.mex Ordinal.mex theorem mex_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ∉ Set.range f := diff --git a/Mathlib/Topology/Algebra/GroupWithZero.lean b/Mathlib/Topology/Algebra/GroupWithZero.lean index 6fe7454906bae..12f5f0e381ef8 100644 --- a/Mathlib/Topology/Algebra/GroupWithZero.lean +++ b/Mathlib/Topology/Algebra/GroupWithZero.lean @@ -107,7 +107,7 @@ theorem tendsto_inv₀ {x : G₀} (hx : x ≠ 0) : Tendsto Inv.inv (𝓝 x) ( continuousAt_inv₀ hx #align tendsto_inv₀ tendsto_inv₀ -theorem continuousOn_inv₀ : ContinuousOn (Inv.inv : G₀ → G₀) ({0}ᶜ) := fun _x hx => +theorem continuousOn_inv₀ : ContinuousOn (Inv.inv : G₀ → G₀) {0}ᶜ := fun _x hx => (continuousAt_inv₀ hx).continuousWithinAt #align continuous_on_inv₀ continuousOn_inv₀ @@ -322,7 +322,7 @@ theorem continuousAt_zpow₀ (x : G₀) (m : ℤ) (h : x ≠ 0 ∨ 0 ≤ m) : exact (continuousAt_pow x (m + 1)).inv₀ (pow_ne_zero _ hx) #align continuous_at_zpow₀ continuousAt_zpow₀ -theorem continuousOn_zpow₀ (m : ℤ) : ContinuousOn (fun x : G₀ => x ^ m) ({0}ᶜ) := fun _x hx => +theorem continuousOn_zpow₀ (m : ℤ) : ContinuousOn (fun x : G₀ => x ^ m) {0}ᶜ := fun _x hx => (continuousAt_zpow₀ _ _ (Or.inl hx)).continuousWithinAt #align continuous_on_zpow₀ continuousOn_zpow₀ diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean index 80a3e3b2289e0..80b17504b49e3 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean @@ -770,7 +770,7 @@ end Countable variable [ContinuousAdd α] theorem tsum_add_tsum_compl {s : Set β} (hs : Summable (f ∘ (↑) : s → α)) - (hsc : Summable (f ∘ (↑) : ↑(sᶜ) → α)) : ((∑' x : s, f x) + ∑' x : ↑(sᶜ), f x) = ∑' x, f x := + (hsc : Summable (f ∘ (↑) : ↑sᶜ → α)) : ((∑' x : s, f x) + ∑' x : ↑sᶜ, f x) = ∑' x, f x := (hs.hasSum.add_compl hsc.hasSum).tsum_eq.symm #align tsum_add_tsum_compl tsum_add_tsum_compl @@ -851,7 +851,7 @@ theorem Summable.update (hf : Summable f) (b : β) [DecidableEq β] (a : α) : #align summable.update Summable.update theorem HasSum.hasSum_compl_iff {s : Set β} (hf : HasSum (f ∘ (↑) : s → α) a₁) : - HasSum (f ∘ (↑) : ↑(sᶜ) → α) a₂ ↔ HasSum f (a₁ + a₂) := by + HasSum (f ∘ (↑) : ↑sᶜ → α) a₂ ↔ HasSum f (a₁ + a₂) := by refine' ⟨fun h => hf.add_compl h, fun h => _⟩ rw [hasSum_subtype_iff_indicator] at hf ⊢ rw [Set.indicator_compl] @@ -859,12 +859,12 @@ theorem HasSum.hasSum_compl_iff {s : Set β} (hf : HasSum (f ∘ (↑) : s → #align has_sum.has_sum_compl_iff HasSum.hasSum_compl_iff theorem HasSum.hasSum_iff_compl {s : Set β} (hf : HasSum (f ∘ (↑) : s → α) a₁) : - HasSum f a₂ ↔ HasSum (f ∘ (↑) : ↑(sᶜ) → α) (a₂ - a₁) := + HasSum f a₂ ↔ HasSum (f ∘ (↑) : ↑sᶜ → α) (a₂ - a₁) := Iff.symm <| hf.hasSum_compl_iff.trans <| by rw [add_sub_cancel'_right] #align has_sum.has_sum_iff_compl HasSum.hasSum_iff_compl theorem Summable.summable_compl_iff {s : Set β} (hf : Summable (f ∘ (↑) : s → α)) : - Summable (f ∘ (↑) : ↑(sᶜ) → α) ↔ Summable f := + Summable (f ∘ (↑) : ↑sᶜ → α) ↔ Summable f := ⟨fun ⟨_, ha⟩ => (hf.hasSum.hasSum_compl_iff.1 ha).summable, fun ⟨_, ha⟩ => (hf.hasSum.hasSum_iff_compl.1 ha).summable⟩ #align summable.summable_compl_iff Summable.summable_compl_iff @@ -885,7 +885,7 @@ protected theorem Finset.summable_compl_iff (s : Finset β) : #align finset.summable_compl_iff Finset.summable_compl_iff theorem Set.Finite.summable_compl_iff {s : Set β} (hs : s.Finite) : - Summable (f ∘ (↑) : ↑(sᶜ) → α) ↔ Summable f := + Summable (f ∘ (↑) : ↑sᶜ → α) ↔ Summable f := (hs.summable f).summable_compl_iff #align set.finite.summable_compl_iff Set.Finite.summable_compl_iff @@ -913,7 +913,7 @@ theorem tsum_sub (hf : Summable f) (hg : Summable g) : #align tsum_sub tsum_sub theorem sum_add_tsum_compl {s : Finset β} (hf : Summable f) : - ((∑ x in s, f x) + ∑' x : ↑((s : Set β)ᶜ), f x) = ∑' x, f x := + ((∑ x in s, f x) + ∑' x : ↑(s : Set β)ᶜ, f x) = ∑' x, f x := ((s.hasSum f).add_compl (s.summable_compl_iff.2 hf).hasSum).tsum_eq.symm #align sum_add_tsum_compl sum_add_tsum_compl @@ -1185,8 +1185,8 @@ theorem Summable.subtype (hf : Summable f) (s : Set β) : Summable (f ∘ (↑) #align summable.subtype Summable.subtype theorem summable_subtype_and_compl {s : Set β} : - ((Summable fun x : s => f x) ∧ Summable fun x : ↑(sᶜ) => f x) ↔ Summable f := - ⟨and_imp.2 Summable.add_compl, fun h => ⟨h.subtype s, h.subtype (sᶜ)⟩⟩ + ((Summable fun x : s => f x) ∧ Summable fun x : ↑sᶜ => f x) ↔ Summable f := + ⟨and_imp.2 Summable.add_compl, fun h => ⟨h.subtype s, h.subtype sᶜ⟩⟩ #align summable_subtype_and_compl summable_subtype_and_compl theorem Summable.sigma_factor {γ : β → Type _} {f : (Σb : β, γ b) → α} (ha : Summable f) (b : β) : @@ -1230,7 +1230,7 @@ theorem tsum_comm [T0Space α] {f : β → γ → α} (h : Summable (Function.un end LocInstances theorem tsum_subtype_add_tsum_subtype_compl [T2Space α] {f : β → α} (hf : Summable f) (s : Set β) : - ((∑' x : s, f x) + ∑' x : ↑(sᶜ), f x) = ∑' x, f x := + ((∑' x : s, f x) + ∑' x : ↑sᶜ, f x) = ∑' x, f x := ((hf.subtype s).hasSum.add_compl (hf.subtype { x | x ∉ s }).hasSum).unique hf.hasSum #align tsum_subtype_add_tsum_subtype_compl tsum_subtype_add_tsum_subtype_compl diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean index e66db511e4334..7dbb6589b967b 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean @@ -250,11 +250,11 @@ theorem summable_abs_iff [LinearOrderedAddCommGroup α] [UniformSpace α] [Unifo [CompleteSpace α] {f : ι → α} : (Summable fun x => |f x|) ↔ Summable f := let s := { x | 0 ≤ f x } have h1 : ∀ x : s, |f x| = f x := fun x => abs_of_nonneg x.2 - have h2 : ∀ x : ↑(sᶜ), |f x| = -f x := fun x => abs_of_neg (not_le.1 x.2) + have h2 : ∀ x : ↑sᶜ, |f x| = -f x := fun x => abs_of_neg (not_le.1 x.2) calc (Summable fun x => |f x|) ↔ - (Summable fun x : s => |f x|) ∧ Summable fun x : ↑(sᶜ) => |f x| := + (Summable fun x : s => |f x|) ∧ Summable fun x : ↑sᶜ => |f x| := summable_subtype_and_compl.symm - _ ↔ (Summable fun x : s => f x) ∧ Summable fun x : ↑(sᶜ) => -f x := by simp only [h1, h2] + _ ↔ (Summable fun x : s => f x) ∧ Summable fun x : ↑sᶜ => -f x := by simp only [h1, h2] _ ↔ Summable f := by simp only [summable_neg_iff, summable_subtype_and_compl] #align summable_abs_iff summable_abs_iff diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index d93b202b2365f..fc637321efeac 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -94,7 +94,7 @@ theorem unique_topology_of_t2 {t : TopologicalSpace 𝕜} (h₁ : @TopologicalAd mem_compl_singleton_iff.mpr <| Ne.symm <| norm_ne_zero_iff.mp hξ₀.ne.symm -- Thus, its balanced core `𝓑` is too. Let's show that the closed ball of radius `ε` contains -- `𝓑`, which will imply that the closed ball is indeed a `𝓣`-neighborhood of 0. - have : balancedCore 𝕜 ({ξ₀}ᶜ) ∈ @nhds 𝕜 t 0 := balancedCore_mem_nhds_zero this + have : balancedCore 𝕜 {ξ₀}ᶜ ∈ @nhds 𝕜 t 0 := balancedCore_mem_nhds_zero this refine' mem_of_superset this fun ξ hξ => _ -- Let `ξ ∈ 𝓑`. We want to show `‖ξ‖ < ε`. If `ξ = 0`, this is trivial. by_cases hξ0 : ξ = 0 @@ -104,7 +104,7 @@ theorem unique_topology_of_t2 {t : TopologicalSpace 𝕜} (h₁ : @TopologicalAd -- Now suppose `ξ ≠ 0`. By contradiction, let's assume `ε < ‖ξ‖`, and show that -- `ξ₀ ∈ 𝓑 ⊆ {ξ₀}ᶜ`, which is a contradiction. by_contra' h - suffices (ξ₀ * ξ⁻¹) • ξ ∈ balancedCore 𝕜 ({ξ₀}ᶜ) by + suffices (ξ₀ * ξ⁻¹) • ξ ∈ balancedCore 𝕜 {ξ₀}ᶜ by rw [smul_eq_mul 𝕜, mul_assoc, inv_mul_cancel hξ0, mul_one] at this exact not_mem_compl_iff.mpr (mem_singleton ξ₀) ((balancedCore_subset _) this) -- For that, we use that `𝓑` is balanced : since `‖ξ₀‖ < ε < ‖ξ‖`, we have `‖ξ₀ / ξ‖ ≤ 1`, @@ -189,7 +189,7 @@ theorem LinearMap.continuous_of_nonzero_on_open (l : E →ₗ[𝕜] 𝕜) (s : S (hs₂ : s.Nonempty) (hs₃ : ∀ x ∈ s, l x ≠ 0) : Continuous l := by refine' l.continuous_of_isClosed_ker (l.isClosed_or_dense_ker.resolve_right fun hl => _) rcases hs₂ with ⟨x, hx⟩ - have : x ∈ interior ((LinearMap.ker l : Set E)ᶜ) := by + have : x ∈ interior (LinearMap.ker l : Set E)ᶜ := by rw [mem_interior_iff_mem_nhds] exact mem_of_superset (hs₁.mem_nhds hx) hs₃ rwa [hl.interior_compl] at this diff --git a/Mathlib/Topology/Algebra/Order/T5.lean b/Mathlib/Topology/Algebra/Order/T5.lean index 38fe92c70bc12..3d763474714eb 100644 --- a/Mathlib/Topology/Algebra/Order/T5.lean +++ b/Mathlib/Topology/Algebra/Order/T5.lean @@ -44,7 +44,7 @@ theorem compl_section_ordSeparatingSet_mem_nhdsWithin_Ici (hd : Disjoint s (clos · exact mem_of_superset hmem' (disjoint_left.1 H) · simp only [Set.disjoint_left, not_forall, Classical.not_not] at H rcases H with ⟨c, ⟨hac, hcb⟩, hc⟩ - have hsub' : Icc a b ⊆ ordConnectedComponent (tᶜ) a := + have hsub' : Icc a b ⊆ ordConnectedComponent tᶜ a := subset_ordConnectedComponent (left_mem_Icc.2 hab) hsub have hd : Disjoint s (ordConnectedSection (ordSeparatingSet s t)) := disjoint_left_ordSeparatingSet.mono_right ordConnectedSection_subset diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index ce5242892b3ce..d0c92185e1291 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -87,7 +87,7 @@ def TopologicalSpace.ofClosed {α : Type u} (T : Set (Set α)) (empty_mem : ∅ (union_mem : ∀ A, A ∈ T → ∀ B, B ∈ T → A ∪ B ∈ T) : TopologicalSpace α where IsOpen X := Xᶜ ∈ T isOpen_univ := by simp [empty_mem] - isOpen_inter s t hs ht := by simpa only [compl_inter] using union_mem (sᶜ) hs (tᶜ) ht + isOpen_inter s t hs ht := by simpa only [compl_inter] using union_mem sᶜ hs tᶜ ht isOpen_sUnion s hs := by simp only [Set.compl_sUnion] exact sInter_mem (compl '' s) fun z ⟨y, hy, hz⟩ => hz ▸ hs y hy @@ -191,14 +191,14 @@ theorem IsOpen.and : IsOpen { a | p₁ a } → IsOpen { a | p₂ a } → IsOpen /-- A set is closed if its complement is open -/ class IsClosed (s : Set α) : Prop where /-- The complement of a closed set is an open set. -/ - isOpen_compl : IsOpen (sᶜ) + isOpen_compl : IsOpen sᶜ #align is_closed IsClosed set_option quotPrecheck false in /-- Notation for `IsClosed` with respect to a non-standard topology. -/ scoped[Topology] notation (name := IsClosed_of) "IsClosed[" t "]" => @IsClosed _ t -@[simp] theorem isOpen_compl_iff {s : Set α} : IsOpen (sᶜ) ↔ IsClosed s := +@[simp] theorem isOpen_compl_iff {s : Set α} : IsOpen sᶜ ↔ IsClosed s := ⟨fun h => ⟨h⟩, fun h => h.isOpen_compl⟩ #align is_open_compl_iff isOpen_compl_iff @@ -229,7 +229,7 @@ theorem isClosed_biInter {s : Set β} {f : β → Set α} (h : ∀ i ∈ s, IsCl #align is_closed_bInter isClosed_biInter @[simp] -theorem isClosed_compl_iff {s : Set α} : IsClosed (sᶜ) ↔ IsOpen s := by +theorem isClosed_compl_iff {s : Set α} : IsClosed sᶜ ↔ IsOpen s := by rw [← isOpen_compl_iff, compl_compl] #align is_closed_compl_iff isClosed_compl_iff @@ -542,18 +542,18 @@ theorem interior_subset_closure {s : Set α} : interior s ⊆ closure s := Subset.trans interior_subset subset_closure #align interior_subset_closure interior_subset_closure -theorem closure_eq_compl_interior_compl {s : Set α} : closure s = interior (sᶜ)ᶜ := by +theorem closure_eq_compl_interior_compl {s : Set α} : closure s = (interior sᶜ)ᶜ := by rw [interior, closure, compl_sUnion, compl_image_set_of] simp only [compl_subset_compl, isOpen_compl_iff] #align closure_eq_compl_interior_compl closure_eq_compl_interior_compl @[simp] -theorem interior_compl {s : Set α} : interior (sᶜ) = closure sᶜ := by +theorem interior_compl {s : Set α} : interior sᶜ = (closure s)ᶜ := by simp [closure_eq_compl_interior_compl] #align interior_compl interior_compl @[simp] -theorem closure_compl {s : Set α} : closure (sᶜ) = interior sᶜ := by +theorem closure_compl {s : Set α} : closure sᶜ = (interior s)ᶜ := by simp [closure_eq_compl_interior_compl] #align closure_compl closure_compl @@ -608,11 +608,11 @@ theorem dense_iff_closure_eq {s : Set α} : Dense s ↔ closure s = univ := alias dense_iff_closure_eq ↔ Dense.closure_eq _ #align dense.closure_eq Dense.closure_eq -theorem interior_eq_empty_iff_dense_compl {s : Set α} : interior s = ∅ ↔ Dense (sᶜ) := by +theorem interior_eq_empty_iff_dense_compl {s : Set α} : interior s = ∅ ↔ Dense sᶜ := by rw [dense_iff_closure_eq, closure_compl, compl_univ_iff] #align interior_eq_empty_iff_dense_compl interior_eq_empty_iff_dense_compl -theorem Dense.interior_compl {s : Set α} (h : Dense s) : interior (sᶜ) = ∅ := +theorem Dense.interior_compl {s : Set α} (h : Dense s) : interior sᶜ = ∅ := interior_eq_empty_iff_dense_compl.2 <| by rwa [compl_compl] #align dense.interior_compl Dense.interior_compl @@ -707,7 +707,7 @@ theorem self_diff_frontier (s : Set α) : s \ frontier s = interior s := by inter_eq_self_of_subset_right interior_subset, empty_union] #align self_diff_frontier self_diff_frontier -theorem frontier_eq_closure_inter_closure {s : Set α} : frontier s = closure s ∩ closure (sᶜ) := by +theorem frontier_eq_closure_inter_closure {s : Set α} : frontier s = closure s ∩ closure sᶜ := by rw [closure_compl, frontier, diff_eq] #align frontier_eq_closure_inter_closure frontier_eq_closure_inter_closure @@ -729,7 +729,7 @@ theorem frontier_interior_subset {s : Set α} : frontier (interior s) ⊆ fronti /-- The complement of a set has the same frontier as the original set. -/ @[simp] -theorem frontier_compl (s : Set α) : frontier (sᶜ) = frontier s := by +theorem frontier_compl (s : Set α) : frontier sᶜ = frontier s := by simp only [frontier_eq_closure_inter_closure, compl_compl, inter_comm] #align frontier_compl frontier_compl @@ -749,8 +749,8 @@ theorem frontier_inter_subset (s t : Set α) : #align frontier_inter_subset frontier_inter_subset theorem frontier_union_subset (s t : Set α) : - frontier (s ∪ t) ⊆ frontier s ∩ closure (tᶜ) ∪ closure (sᶜ) ∩ frontier t := by - simpa only [frontier_compl, ← compl_union] using frontier_inter_subset (sᶜ) (tᶜ) + frontier (s ∪ t) ⊆ frontier s ∩ closure tᶜ ∪ closure sᶜ ∩ frontier t := by + simpa only [frontier_compl, ← compl_union] using frontier_inter_subset sᶜ tᶜ #align frontier_union_subset frontier_union_subset theorem IsClosed.frontier_eq {s : Set α} (hs : IsClosed s) : frontier s = s \ interior s := by @@ -798,12 +798,12 @@ theorem Disjoint.frontier_right (hs : IsOpen s) (hd : Disjoint s t) : Disjoint s #align disjoint.frontier_right Disjoint.frontier_right theorem frontier_eq_inter_compl_interior {s : Set α} : - frontier s = interior sᶜ ∩ interior (sᶜ)ᶜ := by + frontier s = (interior s)ᶜ ∩ (interior sᶜ)ᶜ := by rw [← frontier_compl, ← closure_compl]; rfl #align frontier_eq_inter_compl_interior frontier_eq_inter_compl_interior theorem compl_frontier_eq_union_interior {s : Set α} : - frontier sᶜ = interior s ∪ interior (sᶜ) := by + (frontier s)ᶜ = interior s ∪ interior sᶜ := by rw [frontier_eq_inter_compl_interior] simp only [compl_inter, compl_compl] #align compl_frontier_eq_union_interior compl_frontier_eq_union_interior @@ -835,7 +835,7 @@ scoped[Topology] notation "𝓝" => nhds scoped[Topology] notation "𝓝[" s "] " x:100 => nhdsWithin x s /-- Notation for the filter of punctured neighborhoods of a point. -/ -scoped[Topology] notation "𝓝[≠] " x:100 => nhdsWithin x ({x}ᶜ) +scoped[Topology] notation "𝓝[≠] " x:100 => nhdsWithin x {x}ᶜ /-- Notation for the filter of right neighborhoods of a point. -/ scoped[Topology] notation "𝓝[≥] " x:100 => nhdsWithin x (Set.Ici x) @@ -1169,7 +1169,7 @@ def AccPt (x : α) (F : Filter α) : Prop := NeBot (𝓝[≠] x ⊓ F) #align acc_pt AccPt -theorem acc_iff_cluster (x : α) (F : Filter α) : AccPt x F ↔ ClusterPt x (𝓟 ({x}ᶜ) ⊓ F) := by +theorem acc_iff_cluster (x : α) (F : Filter α) : AccPt x F ↔ ClusterPt x (𝓟 {x}ᶜ ⊓ F) := by rw [AccPt, nhdsWithin, ClusterPt, inf_assoc] #align acc_iff_cluster acc_iff_cluster @@ -1311,7 +1311,7 @@ theorem dense_compl_singleton (x : α) [NeBot (𝓝[≠] x)] : Dense ({x}ᶜ : S /-- If `x` is not an isolated point of a topological space, then the closure of `{x}ᶜ` is the whole space. -/ -- porting note: was a `@[simp]` lemma but `simp` can prove it -theorem closure_compl_singleton (x : α) [NeBot (𝓝[≠] x)] : closure ({x}ᶜ) = (univ : Set α) := +theorem closure_compl_singleton (x : α) [NeBot (𝓝[≠] x)] : closure {x}ᶜ = (univ : Set α) := (dense_compl_singleton x).closure_eq #align closure_compl_singleton closure_compl_singleton @@ -1433,8 +1433,8 @@ theorem Dense.inter_nhds_nonempty {s t : Set α} (hs : Dense s) {x : α} (ht : t theorem closure_diff {s t : Set α} : closure s \ closure t ⊆ closure (s \ t) := calc - closure s \ closure t = closure tᶜ ∩ closure s := by simp only [diff_eq, inter_comm] - _ ⊆ closure (closure tᶜ ∩ s) := (isOpen_compl_iff.mpr <| isClosed_closure).inter_closure + closure s \ closure t = (closure t)ᶜ ∩ closure s := by simp only [diff_eq, inter_comm] + _ ⊆ closure ((closure t)ᶜ ∩ s) := (isOpen_compl_iff.mpr <| isClosed_closure).inter_closure _ = closure (s \ closure t) := by simp only [diff_eq, inter_comm] _ ⊆ closure (s \ t) := closure_mono <| diff_subset_diff (Subset.refl s) subset_closure #align closure_diff closure_diff @@ -1469,7 +1469,7 @@ Then `f` tends to `a` along `l` restricted to `s` if and only if it tends to `a` theorem tendsto_inf_principal_nhds_iff_of_forall_eq {f : β → α} {l : Filter β} {s : Set β} {a : α} (h : ∀ (x) (_ : x ∉ s), f x = a) : Tendsto f (l ⊓ 𝓟 s) (𝓝 a) ↔ Tendsto f l (𝓝 a) := by rw [tendsto_iff_comap, tendsto_iff_comap] - replace h : 𝓟 (sᶜ) ≤ comap f (𝓝 a) + replace h : 𝓟 sᶜ ≤ comap f (𝓝 a) · rintro U ⟨t, ht, htU⟩ x hx have : f x ∈ t := (h x hx).symm ▸ mem_of_mem_nhds ht exact htU this diff --git a/Mathlib/Topology/Bornology/Basic.lean b/Mathlib/Topology/Bornology/Basic.lean index 377a95707f902..554c61628625c 100644 --- a/Mathlib/Topology/Bornology/Basic.lean +++ b/Mathlib/Topology/Bornology/Basic.lean @@ -141,7 +141,7 @@ def IsCobounded (s : Set α) : Prop := /-- `IsBounded` is the predicate that `s` is bounded relative to the ambient bornology on `α`. -/ def IsBounded (s : Set α) : Prop := - IsCobounded (sᶜ) + IsCobounded sᶜ #align bornology.is_bounded Bornology.IsBounded theorem isCobounded_def {s : Set α} : IsCobounded s ↔ s ∈ cobounded α := @@ -153,12 +153,12 @@ theorem isBounded_def {s : Set α} : IsBounded s ↔ sᶜ ∈ cobounded α := #align bornology.is_bounded_def Bornology.isBounded_def @[simp] -theorem isBounded_compl_iff : IsBounded (sᶜ) ↔ IsCobounded s := by +theorem isBounded_compl_iff : IsBounded sᶜ ↔ IsCobounded s := by rw [isBounded_def, isCobounded_def, compl_compl] #align bornology.is_bounded_compl_iff Bornology.isBounded_compl_iff @[simp] -theorem isCobounded_compl_iff : IsCobounded (sᶜ) ↔ IsBounded s := +theorem isCobounded_compl_iff : IsCobounded sᶜ ↔ IsBounded s := Iff.rfl #align bornology.is_cobounded_compl_iff Bornology.isCobounded_compl_iff @@ -239,7 +239,7 @@ theorem ext_iff_isBounded {t t' : Bornology α} : t = t' ↔ ∀ s, @IsBounded α t s ↔ @IsBounded α t' s := ⟨fun h s => h ▸ Iff.rfl, fun h => by ext s - simpa [@isBounded_def _ t, isBounded_def, compl_compl] using h (sᶜ)⟩ + simpa [@isBounded_def _ t, isBounded_def, compl_compl] using h sᶜ⟩ -- porting note: Lean 3 could do this without `@isBounded_def _ t` #align bornology.ext_iff_is_bounded Bornology.ext_iff_isBounded @@ -344,7 +344,7 @@ theorem IsBounded.all (s : Set α) : IsBounded s := #align bornology.is_bounded.all Bornology.IsBounded.all theorem IsCobounded.all (s : Set α) : IsCobounded s := - compl_compl s ▸ IsBounded.all (sᶜ) + compl_compl s ▸ IsBounded.all sᶜ #align bornology.is_cobounded.all Bornology.IsCobounded.all variable (α) diff --git a/Mathlib/Topology/Category/Compactum.lean b/Mathlib/Topology/Category/Compactum.lean index 9beb4cf3804d7..7a64c28bdeacd 100644 --- a/Mathlib/Topology/Category/Compactum.lean +++ b/Mathlib/Topology/Category/Compactum.lean @@ -296,7 +296,7 @@ theorem str_eq_of_le_nhds {X : Compactum} (F : Ultrafilter X) (x : X) : ↑F ≤ intro A hA h by_contra H rw [le_nhds_iff] at cond - specialize cond (Aᶜ) H hA.isOpen_compl + specialize cond Aᶜ H hA.isOpen_compl rw [Ultrafilter.mem_coe, Ultrafilter.compl_mem_iff_not_mem] at cond contradiction -- If A ∈ F, then x ∈ cl A. diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index 10a33075185bd..307492698b7d3 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -361,7 +361,7 @@ theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : apply (IsOpen.continuousOn_iff _).mp · rw [continuousOn_iff_continuous_restrict] have : ∀ a, a ∉ Set.range c.inl → a ∈ Set.range c.inr := by - rintro a (h : a ∈ Set.range c.inlᶜ) + rintro a (h : a ∈ (Set.range c.inl)ᶜ) rwa [eq_compl_iff_isCompl.mpr h₃.symm] convert_to Continuous (g ∘ (Homeomorph.ofEmbedding _ h₂.toEmbedding).symm ∘ Subtype.map _ this) @@ -373,7 +373,7 @@ theorem binaryCofan_isColimit_iff {X Y : TopCat} (c : BinaryCofan X Y) : · continuity · rw [embedding_subtype_val.toInducing.continuous_iff] exact continuous_subtype_val - · change IsOpen (Set.range c.inlᶜ) + · change IsOpen (Set.range c.inl)ᶜ rw [← eq_compl_iff_isCompl.mpr h₃.symm] exact h₂.open_range · intro T f g diff --git a/Mathlib/Topology/Compactification/OnePoint.lean b/Mathlib/Topology/Compactification/OnePoint.lean index cf5c8a19a6969..ee20691066c3d 100644 --- a/Mathlib/Topology/Compactification/OnePoint.lean +++ b/Mathlib/Topology/Compactification/OnePoint.lean @@ -132,7 +132,7 @@ theorem range_coe_inter_infty : range ((↑) : X → OnePoint X) ∩ {∞} = ∅ #align alexandroff.range_coe_inter_infty OnePoint.range_coe_inter_infty @[simp] -theorem compl_range_coe : range ((↑) : X → OnePoint X)ᶜ = {∞} := +theorem compl_range_coe : (range ((↑) : X → OnePoint X))ᶜ = {∞} := compl_range_some X #align alexandroff.compl_range_coe OnePoint.compl_range_coe @@ -187,7 +187,7 @@ that `(↑)` has dense range, so it is a dense embedding. variable [TopologicalSpace X] instance : TopologicalSpace (OnePoint X) where - IsOpen s := (∞ ∈ s → IsCompact ((((↑) : X → OnePoint X) ⁻¹' s)ᶜ)) ∧ + IsOpen s := (∞ ∈ s → IsCompact (((↑) : X → OnePoint X) ⁻¹' s)ᶜ) ∧ IsOpen (((↑) : X → OnePoint X) ⁻¹' s) isOpen_univ := by simp isOpen_inter s t := by @@ -207,17 +207,17 @@ instance : TopologicalSpace (OnePoint X) where variable {s : Set (OnePoint X)} {t : Set X} theorem isOpen_def : - IsOpen s ↔ (∞ ∈ s → IsCompact (((↑) ⁻¹' s : Set X)ᶜ)) ∧ IsOpen ((↑) ⁻¹' s : Set X) := + IsOpen s ↔ (∞ ∈ s → IsCompact ((↑) ⁻¹' s : Set X)ᶜ) ∧ IsOpen ((↑) ⁻¹' s : Set X) := Iff.rfl #align alexandroff.is_open_def OnePoint.isOpen_def theorem isOpen_iff_of_mem' (h : ∞ ∈ s) : - IsOpen s ↔ IsCompact (((↑) ⁻¹' s : Set X)ᶜ) ∧ IsOpen ((↑) ⁻¹' s : Set X) := by + IsOpen s ↔ IsCompact ((↑) ⁻¹' s : Set X)ᶜ ∧ IsOpen ((↑) ⁻¹' s : Set X) := by simp [isOpen_def, h] #align alexandroff.is_open_iff_of_mem' OnePoint.isOpen_iff_of_mem' theorem isOpen_iff_of_mem (h : ∞ ∈ s) : - IsOpen s ↔ IsClosed (((↑) ⁻¹' s : Set X)ᶜ) ∧ IsCompact (((↑) ⁻¹' s : Set X)ᶜ) := by + IsOpen s ↔ IsClosed ((↑) ⁻¹' s : Set X)ᶜ ∧ IsCompact ((↑) ⁻¹' s : Set X)ᶜ := by simp only [isOpen_iff_of_mem' h, isClosed_compl_iff, and_comm] #align alexandroff.is_open_iff_of_mem OnePoint.isOpen_iff_of_mem @@ -241,7 +241,7 @@ theorem isOpen_image_coe {s : Set X} : IsOpen ((↑) '' s : Set (OnePoint X)) #align alexandroff.is_open_image_coe OnePoint.isOpen_image_coe theorem isOpen_compl_image_coe {s : Set X} : - IsOpen (((↑) '' s : Set (OnePoint X))ᶜ) ↔ IsClosed s ∧ IsCompact s := by + IsOpen ((↑) '' s : Set (OnePoint X))ᶜ ↔ IsClosed s ∧ IsCompact s := by rw [isOpen_iff_of_mem, ← preimage_compl, compl_compl, preimage_image_eq _ coe_injective] exact infty_not_mem_image_coe #align alexandroff.is_open_compl_image_coe OnePoint.isOpen_compl_image_coe @@ -364,7 +364,7 @@ theorem tendsto_nhds_infty' {α : Type _} {f : OnePoint X → α} {l : Filter α theorem tendsto_nhds_infty {α : Type _} {f : OnePoint X → α} {l : Filter α} : Tendsto f (𝓝 ∞) l ↔ - ∀ s ∈ l, f ∞ ∈ s ∧ ∃ t : Set X, IsClosed t ∧ IsCompact t ∧ MapsTo (f ∘ (↑)) (tᶜ) s := + ∀ s ∈ l, f ∞ ∈ s ∧ ∃ t : Set X, IsClosed t ∧ IsCompact t ∧ MapsTo (f ∘ (↑)) tᶜ s := tendsto_nhds_infty'.trans <| by simp only [tendsto_pure_left, hasBasis_coclosedCompact.tendsto_left_iff, forall_and, and_assoc, exists_prop] @@ -377,7 +377,7 @@ theorem continuousAt_infty' {Y : Type _} [TopologicalSpace Y] {f : OnePoint X theorem continuousAt_infty {Y : Type _} [TopologicalSpace Y] {f : OnePoint X → Y} : ContinuousAt f ∞ ↔ - ∀ s ∈ 𝓝 (f ∞), ∃ t : Set X, IsClosed t ∧ IsCompact t ∧ MapsTo (f ∘ (↑)) (tᶜ) s := + ∀ s ∈ 𝓝 (f ∞), ∃ t : Set X, IsClosed t ∧ IsCompact t ∧ MapsTo (f ∘ (↑)) tᶜ s := continuousAt_infty'.trans <| by simp only [hasBasis_coclosedCompact.tendsto_left_iff, and_assoc] #align alexandroff.continuous_at_infty OnePoint.continuousAt_infty diff --git a/Mathlib/Topology/Connected.lean b/Mathlib/Topology/Connected.lean index 388c418e625fa..3344a03760bc3 100644 --- a/Mathlib/Topology/Connected.lean +++ b/Mathlib/Topology/Connected.lean @@ -454,7 +454,7 @@ theorem IsPreconnected.subset_clopen {s t : Set α} (hs : IsPreconnected s) (ht contained in `u`, then the whole set `s` is contained in `u`. -/ theorem IsPreconnected.subset_of_closure_inter_subset (hs : IsPreconnected s) (hu : IsOpen u) (h'u : (s ∩ u).Nonempty) (h : closure u ∩ s ⊆ u) : s ⊆ u := by - have A : s ⊆ u ∪ closure uᶜ := by + have A : s ⊆ u ∪ (closure u)ᶜ := by intro x hx by_cases xu : x ∈ u · exact Or.inl xu @@ -1292,7 +1292,7 @@ theorem isTotallyDisconnected_of_clopen_set {X : Type _} [TopologicalSpace X] rcases h_contra with ⟨x, hx, y, hy, hxy⟩ obtain ⟨U, h_clopen, hxU, hyU⟩ := hX hxy specialize - hS U (Uᶜ) h_clopen.1 h_clopen.compl.1 (fun a _ => em (a ∈ U)) ⟨x, hx, hxU⟩ ⟨y, hy, hyU⟩ + hS U Uᶜ h_clopen.1 h_clopen.compl.1 (fun a _ => em (a ∈ U)) ⟨x, hx, hxU⟩ ⟨y, hy, hyU⟩ rw [inter_compl_self, Set.inter_empty] at hS exact Set.not_nonempty_empty hS #align is_totally_disconnected_of_clopen_set isTotallyDisconnected_of_clopen_set diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index c6b98d67f0bd2..0efa88ae030f6 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -268,7 +268,7 @@ def of : α ≃ CofiniteTopology α := instance [Inhabited α] : Inhabited (CofiniteTopology α) where default := of default instance : TopologicalSpace (CofiniteTopology α) where - IsOpen s := s.Nonempty → Set.Finite (sᶜ) + IsOpen s := s.Nonempty → Set.Finite sᶜ isOpen_univ := by simp isOpen_inter s t := by rintro hs ht ⟨x, hxs, hxt⟩ @@ -1401,7 +1401,7 @@ theorem pi_generateFrom_eq_finite {π : ι → Type _} {g : ∀ a, Set (Set (π letI := generateFrom { t | ∃ s : ∀ a, Set (π a), (∀ a, s a ∈ g a) ∧ t = pi univ s } refine isOpen_iff_forall_mem_open.2 fun f hf => ?_ choose c hcg hfc using fun a => sUnion_eq_univ_iff.1 (hg a) (f a) - refine ⟨pi i t ∩ pi (↑iᶜ : Set ι) c, inter_subset_left _ _, ?_, ⟨hf, fun a _ => hfc a⟩⟩ + refine ⟨pi i t ∩ pi ((↑i)ᶜ : Set ι) c, inter_subset_left _ _, ?_, ⟨hf, fun a _ => hfc a⟩⟩ rw [← univ_pi_piecewise] refine GenerateOpen.basic _ ⟨_, fun a => ?_, rfl⟩ by_cases a ∈ i <;> simp [*] diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousFunction/Bounded.lean index c546f84ab2b46..e19bdda2bcfaf 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousFunction/Bounded.lean @@ -483,16 +483,16 @@ theorem extend_of_empty [IsEmpty α] (f : α ↪ δ) (g : α →ᵇ β) (h : δ @[simp] theorem dist_extend_extend (f : α ↪ δ) (g₁ g₂ : α →ᵇ β) (h₁ h₂ : δ →ᵇ β) : dist (g₁.extend f h₁) (g₂.extend f h₂) = - max (dist g₁ g₂) (dist (h₁.restrict (range fᶜ)) (h₂.restrict (range fᶜ))) := by + max (dist g₁ g₂) (dist (h₁.restrict (range f)ᶜ) (h₂.restrict (range f)ᶜ)) := by refine' le_antisymm ((dist_le <| le_max_iff.2 <| Or.inl dist_nonneg).2 fun x => _) (max_le _ _) · rcases _root_.em (∃ y, f y = x) with (⟨x, rfl⟩ | hx) · simp only [extend_apply] exact (dist_coe_le_dist x).trans (le_max_left _ _) · simp only [extend_apply' hx] - lift x to (range fᶜ : Set δ) using hx + lift x to ((range f)ᶜ : Set δ) using hx calc - dist (h₁ x) (h₂ x) = dist (h₁.restrict (range fᶜ) x) (h₂.restrict (range fᶜ) x) := rfl - _ ≤ dist (h₁.restrict (range fᶜ)) (h₂.restrict (range fᶜ)) := (dist_coe_le_dist x) + dist (h₁ x) (h₂ x) = dist (h₁.restrict (range f)ᶜ x) (h₂.restrict (range f)ᶜ x) := rfl + _ ≤ dist (h₁.restrict (range f)ᶜ) (h₂.restrict (range f)ᶜ) := (dist_coe_le_dist x) _ ≤ _ := le_max_right _ _ · refine' (dist_le dist_nonneg).2 fun x => _ rw [← extend_apply f g₁ h₁, ← extend_apply f g₂ h₂] diff --git a/Mathlib/Topology/ContinuousFunction/Ideals.lean b/Mathlib/Topology/ContinuousFunction/Ideals.lean index 7fefe681bf786..f3f6d0eff6b68 100644 --- a/Mathlib/Topology/ContinuousFunction/Ideals.lean +++ b/Mathlib/Topology/ContinuousFunction/Ideals.lean @@ -215,7 +215,7 @@ theorem idealOfSet_ofIdeal_eq_closure (I : Ideal C(X, 𝕜)) : -- Let `t := {x : X | ε / 2 ≤ ‖f x‖₊}}` which is closed and disjoint from `set_of_ideal I`. set t := {x : X | ε / 2 ≤ ‖f x‖₊} have ht : IsClosed t := isClosed_le continuous_const (map_continuous f).nnnorm - have htI : Disjoint t (setOfIdeal Iᶜ) := by + have htI : Disjoint t (setOfIdeal I)ᶜ := by refine' Set.subset_compl_iff_disjoint_left.mp fun x hx => _ simpa only [Set.mem_setOf, Set.mem_compl_iff, not_le] using (nnnorm_eq_zero.mpr (mem_idealOfSet.mp hf hx)).trans_lt (half_pos hε) @@ -384,7 +384,7 @@ theorem setOfIdeal_eq_compl_singleton (I : Ideal C(X, 𝕜)) [hI : I.IsMaximal] #align continuous_map.set_of_ideal_eq_compl_singleton ContinuousMap.setOfIdeal_eq_compl_singleton theorem ideal_isMaximal_iff (I : Ideal C(X, 𝕜)) [hI : IsClosed (I : Set C(X, 𝕜))] : - I.IsMaximal ↔ ∃ x : X, idealOfSet 𝕜 ({x}ᶜ) = I := by + I.IsMaximal ↔ ∃ x : X, idealOfSet 𝕜 {x}ᶜ = I := by refine' ⟨_, fun h => let ⟨x, hx⟩ := h diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 3c497a5549f02..a8ff534caee99 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -787,7 +787,7 @@ theorem continuousWithinAt_diff_self {f : α → β} {s : Set α} {x : α} : @[simp] theorem continuousWithinAt_compl_self {f : α → β} {a : α} : - ContinuousWithinAt f ({a}ᶜ) a ↔ ContinuousAt f a := by + ContinuousWithinAt f {a}ᶜ a ↔ ContinuousAt f a := by rw [compl_eq_univ_diff, continuousWithinAt_diff_self, continuousWithinAt_univ] #align continuous_within_at_compl_self continuousWithinAt_compl_self @@ -1122,7 +1122,7 @@ theorem ContinuousOn.if' {s : Set α} {p : α → Prop} {f g : α → β} [∀ a cases' hx with hx hx · apply ContinuousWithinAt.union · exact (hf x hx).congr (fun y hy => if_pos hy.2) (if_pos hx.2) - · have : x ∉ closure ({ a | p a }ᶜ) := fun h => hx' ⟨subset_closure hx.2, by + · have : x ∉ closure { a | p a }ᶜ := fun h => hx' ⟨subset_closure hx.2, by rwa [closure_compl] at h⟩ exact continuousWithinAt_of_not_mem_closure fun h => this (closure_inter_subset_inter_closure _ _ h).2 @@ -1165,7 +1165,7 @@ theorem ContinuousOn.if {α β : Type _} [TopologicalSpace α] [TopologicalSpace theorem ContinuousOn.piecewise {s t : Set α} {f g : α → β} [∀ a, Decidable (a ∈ t)] (ht : ∀ a ∈ s ∩ frontier t, f a = g a) (hf : ContinuousOn f <| s ∩ closure t) - (hg : ContinuousOn g <| s ∩ closure (tᶜ)) : ContinuousOn (piecewise t f g) s := + (hg : ContinuousOn g <| s ∩ closure tᶜ) : ContinuousOn (piecewise t f g) s := hf.if ht hg #align continuous_on.piecewise ContinuousOn.piecewise @@ -1205,7 +1205,7 @@ theorem Continuous.if_const (p : Prop) {f g : α → β} [Decidable p] (hf : Con theorem continuous_piecewise {s : Set α} {f g : α → β} [∀ a, Decidable (a ∈ s)] (hs : ∀ a ∈ frontier s, f a = g a) (hf : ContinuousOn f (closure s)) - (hg : ContinuousOn g (closure (sᶜ))) : Continuous (piecewise s f g) := + (hg : ContinuousOn g (closure sᶜ)) : Continuous (piecewise s f g) := continuous_if hs hf hg #align continuous_piecewise continuous_piecewise @@ -1236,13 +1236,13 @@ theorem ite_inter_closure_eq_of_inter_frontier_eq {s s' t : Set α} #align ite_inter_closure_eq_of_inter_frontier_eq ite_inter_closure_eq_of_inter_frontier_eq theorem ite_inter_closure_compl_eq_of_inter_frontier_eq {s s' t : Set α} - (ht : s ∩ frontier t = s' ∩ frontier t) : t.ite s s' ∩ closure (tᶜ) = s' ∩ closure (tᶜ) := by + (ht : s ∩ frontier t = s' ∩ frontier t) : t.ite s s' ∩ closure tᶜ = s' ∩ closure tᶜ := by rw [← ite_compl, ite_inter_closure_eq_of_inter_frontier_eq] rwa [frontier_compl, eq_comm] #align ite_inter_closure_compl_eq_of_inter_frontier_eq ite_inter_closure_compl_eq_of_inter_frontier_eq theorem continuousOn_piecewise_ite' {s s' t : Set α} {f f' : α → β} [∀ x, Decidable (x ∈ t)] - (h : ContinuousOn f (s ∩ closure t)) (h' : ContinuousOn f' (s' ∩ closure (tᶜ))) + (h : ContinuousOn f (s ∩ closure t)) (h' : ContinuousOn f' (s' ∩ closure tᶜ)) (H : s ∩ frontier t = s' ∩ frontier t) (Heq : EqOn f f' (s ∩ frontier t)) : ContinuousOn (t.piecewise f f') (t.ite s s') := by apply ContinuousOn.piecewise diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index debee63bf1ba3..6d53951b5170d 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -85,7 +85,7 @@ theorem dense_image (di : DenseInducing i) {s : Set α} : Dense (i '' s) ↔ Den /-- If `i : α → β` is a dense embedding with dense complement of the range, then any compact set in `α` has empty interior. -/ -theorem interior_compact_eq_empty [T2Space β] (di : DenseInducing i) (hd : Dense (range iᶜ)) +theorem interior_compact_eq_empty [T2Space β] (di : DenseInducing i) (hd : Dense (range i)ᶜ) {s : Set α} (hs : IsCompact s) : interior s = ∅ := by refine' eq_empty_iff_forall_not_mem.2 fun x hx => _ rw [mem_interior_iff_mem_nhds] at hx diff --git a/Mathlib/Topology/GDelta.lean b/Mathlib/Topology/GDelta.lean index a9e72cfcc1700..f260543be63b8 100644 --- a/Mathlib/Topology/GDelta.lean +++ b/Mathlib/Topology/GDelta.lean @@ -141,16 +141,16 @@ theorem isGδ_compl_singleton (a : α) : IsGδ ({a}ᶜ : Set α) := isOpen_compl_singleton.isGδ #align is_Gδ_compl_singleton isGδ_compl_singleton -theorem Set.Countable.isGδ_compl {s : Set α} (hs : s.Countable) : IsGδ (sᶜ) := by +theorem Set.Countable.isGδ_compl {s : Set α} (hs : s.Countable) : IsGδ sᶜ := by rw [← biUnion_of_singleton s, compl_iUnion₂] exact isGδ_biInter hs fun x _ => isGδ_compl_singleton x #align set.countable.is_Gδ_compl Set.Countable.isGδ_compl -theorem Set.Finite.isGδ_compl {s : Set α} (hs : s.Finite) : IsGδ (sᶜ) := +theorem Set.Finite.isGδ_compl {s : Set α} (hs : s.Finite) : IsGδ sᶜ := hs.countable.isGδ_compl #align set.finite.is_Gδ_compl Set.Finite.isGδ_compl -theorem Set.Subsingleton.isGδ_compl {s : Set α} (hs : s.Subsingleton) : IsGδ (sᶜ) := +theorem Set.Subsingleton.isGδ_compl {s : Set α} (hs : s.Subsingleton) : IsGδ sᶜ := hs.finite.isGδ_compl #align set.subsingleton.is_Gδ_compl Set.Subsingleton.isGδ_compl diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index bef0722f52f38..cf11817f90d76 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -83,7 +83,7 @@ theorem specializes_TFAE (x y : X) : tfae_have 2 → 3 · exact fun h s hso hy => h (hso.mem_nhds hy) tfae_have 3 → 4 - · exact fun h s hsc hx => of_not_not fun hy => h (sᶜ) hsc.isOpen_compl hy hx + · exact fun h s hsc hx => of_not_not fun hy => h sᶜ hsc.isOpen_compl hy hx tfae_have 4 → 5 · exact fun h => h _ isClosed_closure (subset_closure <| mem_singleton _) tfae_have 6 ↔ 5 diff --git a/Mathlib/Topology/Instances/RatLemmas.lean b/Mathlib/Topology/Instances/RatLemmas.lean index f99b3037990d6..7a27bf59f5976 100644 --- a/Mathlib/Topology/Instances/RatLemmas.lean +++ b/Mathlib/Topology/Instances/RatLemmas.lean @@ -46,7 +46,7 @@ theorem interior_compact_eq_empty (hs : IsCompact s) : interior s = ∅ := denseEmbedding_coe_real.toDenseInducing.interior_compact_eq_empty dense_irrational hs #align rat.interior_compact_eq_empty Rat.interior_compact_eq_empty -theorem dense_compl_compact (hs : IsCompact s) : Dense (sᶜ) := +theorem dense_compl_compact (hs : IsCompact s) : Dense sᶜ := interior_eq_empty_iff_dense_compl.1 (interior_compact_eq_empty hs) #align rat.dense_compl_compact Rat.dense_compl_compact diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index f0497d6f45233..5451caf8b5ada 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -102,7 +102,7 @@ theorem isOpen_iff_coe_preimage_of_iSup_eq_top (s : Set β) : theorem isClosed_iff_coe_preimage_of_iSup_eq_top (s : Set β) : IsClosed s ↔ ∀ i, IsClosed ((↑) ⁻¹' s : Set (U i)) := by - simpa using isOpen_iff_coe_preimage_of_iSup_eq_top hU (sᶜ) + simpa using isOpen_iff_coe_preimage_of_iSup_eq_top hU sᶜ #align is_closed_iff_coe_preimage_of_supr_eq_top isClosed_iff_coe_preimage_of_iSup_eq_top theorem isClosedMap_iff_isClosedMap_of_iSup_eq_top : diff --git a/Mathlib/Topology/LocalHomeomorph.lean b/Mathlib/Topology/LocalHomeomorph.lean index 92c71281d5e36..e262f3cfe993f 100644 --- a/Mathlib/Topology/LocalHomeomorph.lean +++ b/Mathlib/Topology/LocalHomeomorph.lean @@ -548,7 +548,7 @@ theorem of_symm_image_eq (h : e.symm '' (e.target ∩ t) = e.source ∩ s) : e.I LocalEquiv.IsImage.of_symm_image_eq h #align local_homeomorph.is_image.of_symm_image_eq LocalHomeomorph.IsImage.of_symm_image_eq -protected theorem compl (h : e.IsImage s t) : e.IsImage (sᶜ) (tᶜ) := fun _ hx => (h hx).not +protected theorem compl (h : e.IsImage s t) : e.IsImage sᶜ tᶜ := fun _ hx => (h hx).not #align local_homeomorph.is_image.compl LocalHomeomorph.IsImage.compl protected theorem inter {s' t'} (h : e.IsImage s t) (h' : e.IsImage s' t') : diff --git a/Mathlib/Topology/LocallyConstant/Basic.lean b/Mathlib/Topology/LocallyConstant/Basic.lean index 6fb763bd9152f..2522bc2d98d81 100644 --- a/Mathlib/Topology/LocallyConstant/Basic.lean +++ b/Mathlib/Topology/LocallyConstant/Basic.lean @@ -76,7 +76,7 @@ theorem isOpen_fiber {f : X → Y} (hf : IsLocallyConstant f) (y : Y) : IsOpen { #align is_locally_constant.is_open_fiber IsLocallyConstant.isOpen_fiber theorem isClosed_fiber {f : X → Y} (hf : IsLocallyConstant f) (y : Y) : IsClosed { x | f x = y } := - ⟨hf ({y}ᶜ)⟩ + ⟨hf {y}ᶜ⟩ #align is_locally_constant.is_closed_fiber IsLocallyConstant.isClosed_fiber theorem isClopen_fiber {f : X → Y} (hf : IsLocallyConstant f) (y : Y) : IsClopen { x | f x = y } := @@ -157,7 +157,7 @@ theorem apply_eq_of_isPreconnected {f : X → Y} (hf : IsLocallyConstant f) {s : let U := f ⁻¹' {f y} suffices : x ∉ Uᶜ; exact Classical.not_not.1 this intro hxV - specialize hs U (Uᶜ) (hf {f y}) (hf ({f y}ᶜ)) _ ⟨y, ⟨hy, rfl⟩⟩ ⟨x, ⟨hx, hxV⟩⟩ + specialize hs U Uᶜ (hf {f y}) (hf {f y}ᶜ) _ ⟨y, ⟨hy, rfl⟩⟩ ⟨x, ⟨hx, hxV⟩⟩ · simp only [union_compl_self, subset_univ] · simp only [inter_empty, Set.not_nonempty_empty, inter_compl_self] at hs #align is_locally_constant.apply_eq_of_is_preconnected IsLocallyConstant.apply_eq_of_isPreconnected diff --git a/Mathlib/Topology/LocallyFinite.lean b/Mathlib/Topology/LocallyFinite.lean index efb31c7a2fcaf..830c195f30247 100644 --- a/Mathlib/Topology/LocallyFinite.lean +++ b/Mathlib/Topology/LocallyFinite.lean @@ -142,7 +142,7 @@ theorem isClosed_iUnion (hf : LocallyFinite f) (hc : ∀ i, IsClosed (f i)) : /-- If `f : β → Set α` is a locally finite family of closed sets, then for any `x : α`, the intersection of the complements to `f i`, `x ∉ f i`, is a neighbourhood of `x`. -/ theorem iInter_compl_mem_nhds (hf : LocallyFinite f) (hc : ∀ i, IsClosed (f i)) (x : X) : - (⋂ (i) (_ : x ∉ f i), f iᶜ) ∈ 𝓝 x := by + (⋂ (i) (_ : x ∉ f i), (f i)ᶜ) ∈ 𝓝 x := by refine' IsOpen.mem_nhds _ (mem_iInter₂.2 fun i => id) suffices IsClosed (⋃ i : { i // x ∉ f i }, f i) by rwa [← isOpen_compl_iff, compl_iUnion, iInter_subtype] at this diff --git a/Mathlib/Topology/MetricSpace/Baire.lean b/Mathlib/Topology/MetricSpace/Baire.lean index 894edd99516bb..3278fe6f58a13 100644 --- a/Mathlib/Topology/MetricSpace/Baire.lean +++ b/Mathlib/Topology/MetricSpace/Baire.lean @@ -294,7 +294,7 @@ is dense. Formulated here with `⋃`. -/ theorem IsGδ.dense_iUnion_interior_of_closed [Encodable ι] {s : Set α} (hs : IsGδ s) (hd : Dense s) {f : ι → Set α} (hc : ∀ i, IsClosed (f i)) (hU : s ⊆ ⋃ i, f i) : Dense (⋃ i, interior (f i)) := by - let g i := frontier (f i)ᶜ + let g i := (frontier (f i))ᶜ have hgo : ∀ i, IsOpen (g i) := fun i => isClosed_frontier.isOpen_compl have hgd : Dense (⋂ i, g i) := by refine' dense_iInter_of_open hgo fun i x => _ diff --git a/Mathlib/Topology/MetricSpace/HausdorffDimension.lean b/Mathlib/Topology/MetricSpace/HausdorffDimension.lean index 110605249f78d..90312b87df086 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDimension.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDimension.lean @@ -580,7 +580,7 @@ end Real variable {E F : Type _} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] [NormedAddCommGroup F] [NormedSpace ℝ F] -theorem dense_compl_of_dimH_lt_finrank {s : Set E} (hs : dimH s < finrank ℝ E) : Dense (sᶜ) := by +theorem dense_compl_of_dimH_lt_finrank {s : Set E} (hs : dimH s < finrank ℝ E) : Dense sᶜ := by refine fun x => mem_closure_iff_nhds.2 fun t ht => nonempty_iff_ne_empty.2 fun he => hs.not_le ?_ rw [← diff_eq, diff_eq_empty] at he rw [← Real.dimH_of_mem_nhds ht] @@ -624,7 +624,7 @@ vector spaces. Suppose that `f` is `C¹` smooth on a convex set `s` of Hausdorff less than the dimension of `F`. Then the complement of the image `f '' s` is dense in `F`. -/ theorem ContDiffOn.dense_compl_image_of_dimH_lt_finrank [FiniteDimensional ℝ F] {f : E → F} {s t : Set E} (h : ContDiffOn ℝ 1 f s) (hc : Convex ℝ s) (ht : t ⊆ s) - (htF : dimH t < finrank ℝ F) : Dense ((f '' t)ᶜ) := + (htF : dimH t < finrank ℝ F) : Dense (f '' t)ᶜ := dense_compl_of_dimH_lt_finrank <| (h.dimH_image_le hc ht).trans_lt htF set_option linter.uppercaseLean3 false in #align cont_diff_on.dense_compl_image_of_dimH_lt_finrank ContDiffOn.dense_compl_image_of_dimH_lt_finrank @@ -633,6 +633,6 @@ set_option linter.uppercaseLean3 false in real vector space `F` of strictly larger dimension, then the complement of the range of `f` is dense in `F`. -/ theorem ContDiff.dense_compl_range_of_finrank_lt_finrank [FiniteDimensional ℝ F] {f : E → F} - (h : ContDiff ℝ 1 f) (hEF : finrank ℝ E < finrank ℝ F) : Dense (range fᶜ) := + (h : ContDiff ℝ 1 f) (hEF : finrank ℝ E < finrank ℝ F) : Dense (range f)ᶜ := dense_compl_of_dimH_lt_finrank <| h.dimH_range_le.trans_lt <| Nat.cast_lt.2 hEF #align cont_diff.dense_compl_range_of_finrank_lt_finrank ContDiff.dense_compl_range_of_finrank_lt_finrank diff --git a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean index 108797f06a823..1626784e9a922 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean @@ -207,17 +207,17 @@ theorem infEdist_smul {M} [SMul M α] [IsometricSMul M α] (c : M) (x : α) (s : theorem _root_.IsOpen.exists_iUnion_isClosed {U : Set α} (hU : IsOpen U) : ∃ F : ℕ → Set α, (∀ n, IsClosed (F n)) ∧ (∀ n, F n ⊆ U) ∧ (⋃ n, F n) = U ∧ Monotone F := by obtain ⟨a, a_pos, a_lt_one⟩ : ∃ a : ℝ≥0∞, 0 < a ∧ a < 1 := exists_between zero_lt_one - let F := fun n : ℕ => (fun x => infEdist x (Uᶜ)) ⁻¹' Ici (a ^ n) + let F := fun n : ℕ => (fun x => infEdist x Uᶜ) ⁻¹' Ici (a ^ n) have F_subset : ∀ n, F n ⊆ U := fun n x hx ↦ by by_contra h - have : infEdist x (Uᶜ) ≠ 0 := ((ENNReal.pow_pos a_pos _).trans_le hx).ne' + have : infEdist x Uᶜ ≠ 0 := ((ENNReal.pow_pos a_pos _).trans_le hx).ne' exact this (infEdist_zero_of_mem h) refine ⟨F, fun n => IsClosed.preimage continuous_infEdist isClosed_Ici, F_subset, ?_, ?_⟩ show (⋃ n, F n) = U · refine' Subset.antisymm (by simp only [iUnion_subset_iff, F_subset, forall_const]) fun x hx => _ have : ¬x ∈ Uᶜ := by simpa using hx rw [mem_iff_infEdist_zero_of_closed hU.isClosed_compl] at this - have B : 0 < infEdist x (Uᶜ) := by simpa [pos_iff_ne_zero] using this + have B : 0 < infEdist x Uᶜ := by simpa [pos_iff_ne_zero] using this have : Filter.Tendsto (fun n => a ^ n) atTop (𝓝 0) := ENNReal.tendsto_pow_atTop_nhds_0_of_lt_1 a_lt_one rcases((tendsto_order.1 this).2 _ B).exists with ⟨n, hn⟩ @@ -542,7 +542,7 @@ theorem ball_infDist_subset_compl : ball x (infDist x s) ⊆ sᶜ := (disjoint_ball_infDist (s := s)).subset_compl_right #align metric.ball_inf_dist_subset_compl Metric.ball_infDist_subset_compl -theorem ball_infDist_compl_subset : ball x (infDist x (sᶜ)) ⊆ s := +theorem ball_infDist_compl_subset : ball x (infDist x sᶜ) ⊆ s := ball_infDist_subset_compl.trans_eq (compl_compl s) #align metric.ball_inf_dist_compl_subset Metric.ball_infDist_compl_subset diff --git a/Mathlib/Topology/MetricSpace/Metrizable.lean b/Mathlib/Topology/MetricSpace/Metrizable.lean index 564e438d6fba7..c2b98add1b017 100644 --- a/Mathlib/Topology/MetricSpace/Metrizable.lean +++ b/Mathlib/Topology/MetricSpace/Metrizable.lean @@ -162,7 +162,7 @@ theorem exists_embedding_l_infty : ∃ f : X → ℕ →ᵇ ℝ, Embedding f := · exact ⟨fun x => (f x).extend (Encodable.encode' s) 0, (BoundedContinuousFunction.isometry_extend (Encodable.encode' s) (0 : ℕ →ᵇ ℝ)).embedding.comp hf⟩ - have hd : ∀ UV : s, Disjoint (closure UV.1.1) (UV.1.2ᶜ) := + have hd : ∀ UV : s, Disjoint (closure UV.1.1) UV.1.2ᶜ := fun UV => disjoint_compl_right.mono_right (compl_subset_compl.2 UV.2.2) -- Choose a sequence of `εₙ > 0`, `n : s`, that is bounded above by `1` and tends to zero -- along the `cofinite` filter. @@ -173,7 +173,7 @@ theorem exists_embedding_l_infty : ∃ f : X → ℕ →ᵇ ℝ, Embedding f := /- For each `UV = (U, V) ∈ s` we use Urysohn's lemma to choose a function `f UV` that is equal to zero on `U` and is equal to `ε UV` on the complement to `V`. -/ have : ∀ UV : s, ∃ f : C(X, ℝ), - EqOn f 0 UV.1.1 ∧ EqOn f (fun _ => ε UV) (UV.1.2ᶜ) ∧ ∀ x, f x ∈ Icc 0 (ε UV) := by + EqOn f 0 UV.1.1 ∧ EqOn f (fun _ => ε UV) UV.1.2ᶜ ∧ ∀ x, f x ∈ Icc 0 (ε UV) := by intro UV rcases exists_continuous_zero_one_of_closed isClosed_closure (hB.isOpen UV.2.1.2).isClosed_compl (hd UV) with diff --git a/Mathlib/Topology/MetricSpace/PiNat.lean b/Mathlib/Topology/MetricSpace/PiNat.lean index 3115f3288b2c0..f70e236b96371 100644 --- a/Mathlib/Topology/MetricSpace/PiNat.lean +++ b/Mathlib/Topology/MetricSpace/PiNat.lean @@ -911,14 +911,14 @@ protected def metricSpace : MetricSpace (∀ i, F i) where calc dist x y = ∑' i : ι, min ((1 / 2) ^ encode i : ℝ) (dist (x i) (y i)) := rfl _ = (∑ i in K, min ((1 / 2) ^ encode i : ℝ) (dist (x i) (y i))) + - ∑' i : ↑((K : Set ι)ᶜ), min ((1 / 2) ^ encode (i : ι) : ℝ) (dist (x i) (y i)) := + ∑' i : ↑(K : Set ι)ᶜ, min ((1 / 2) ^ encode (i : ι) : ℝ) (dist (x i) (y i)) := (sum_add_tsum_compl (dist_summable _ _)).symm _ ≤ (∑ i in K, dist (x i) (y i)) + - ∑' i : ↑((K : Set ι)ᶜ), ((1 / 2) ^ encode (i : ι) : ℝ) := by + ∑' i : ↑(K : Set ι)ᶜ, ((1 / 2) ^ encode (i : ι) : ℝ) := by refine' add_le_add (Finset.sum_le_sum fun i _ => min_le_right _ _) _ refine' tsum_le_tsum (fun i => min_le_left _ _) _ _ - · apply Summable.subtype (dist_summable x y) ((↑K : Set ι)ᶜ) - · apply Summable.subtype summable_geometric_two_encode ((↑K : Set ι)ᶜ) + · apply Summable.subtype (dist_summable x y) (↑K : Set ι)ᶜ + · apply Summable.subtype summable_geometric_two_encode (↑K : Set ι)ᶜ _ < (∑ _i in K, δ) + ε / 2 := by apply add_lt_add_of_le_of_lt _ hK refine Finset.sum_le_sum fun i hi => (hxy i ?_).le diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index 7db1f51d48cfd..8036f8c84ee4c 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -245,7 +245,7 @@ boundary. Porting note: definitions and lemmas in this section now take `(s : Opens α)` instead of `{s : Set α} (hs : IsOpen s)` so that we can turn various definitions and lemmas into instances. -Also, some lemmas used to assume `Set.Nonempty (sᶜ)` in Lean 3. In fact, this assumption is not +Also, some lemmas used to assume `Set.Nonempty sᶜ` in Lean 3. In fact, this assumption is not needed, so it was dropped. -/ @@ -266,11 +266,11 @@ by `dist' x y = dist x y + |1 / dist x sᶜ - 1 / dist y sᶜ|`, where the secon the boundary to ensure that Cauchy sequences for `dist'` remain well inside `s`. -/ -- Porting note: in mathlib3 this was only a local instance. instance instDist : Dist (CompleteCopy s) where - dist x y := dist x.1 y.1 + abs (1 / infDist x.1 (sᶜ) - 1 / infDist y.1 (sᶜ)) + dist x y := dist x.1 y.1 + abs (1 / infDist x.1 sᶜ - 1 / infDist y.1 sᶜ) #align polish_space.has_dist_complete_copy TopologicalSpace.Opens.CompleteCopy.instDistₓ theorem dist_eq (x y : CompleteCopy s) : - dist x y = dist x.1 y.1 + abs (1 / infDist x.1 (sᶜ) - 1 / infDist y.1 (sᶜ)) := + dist x y = dist x.1 y.1 + abs (1 / infDist x.1 sᶜ - 1 / infDist y.1 sᶜ) := rfl #align polish_space.dist_complete_copy_eq TopologicalSpace.Opens.CompleteCopy.dist_eqₓ @@ -295,9 +295,9 @@ instance instMetricSpace : MetricSpace (CompleteCopy s) := by · simp only [dist_eq, dist_self, one_div, sub_self, abs_zero, add_zero] · simp only [dist_eq, dist_comm, abs_sub_comm] · calc - dist x z = dist x.1 z.1 + |1 / infDist x.1 (sᶜ) - 1 / infDist z.1 (sᶜ)| := rfl - _ ≤ dist x.1 y.1 + dist y.1 z.1 + (|1 / infDist x.1 (sᶜ) - 1 / infDist y.1 (sᶜ)| + - |1 / infDist y.1 (sᶜ) - 1 / infDist z.1 (sᶜ)|) := + dist x z = dist x.1 z.1 + |1 / infDist x.1 sᶜ - 1 / infDist z.1 sᶜ| := rfl + _ ≤ dist x.1 y.1 + dist y.1 z.1 + (|1 / infDist x.1 sᶜ - 1 / infDist y.1 sᶜ| + + |1 / infDist y.1 sᶜ - 1 / infDist z.1 sᶜ|) := add_le_add (dist_triangle _ _ _) (dist_triangle (1 / infDist _ _) _ _) _ = dist x y + dist y z := add_add_add_comm .. · refine ⟨fun h x hx ↦ ?_, fun h ↦ isOpen_iff_mem_nhds.2 fun x hx ↦ ?_⟩ @@ -305,8 +305,8 @@ instance instMetricSpace : MetricSpace (CompleteCopy s) := by exact ⟨ε, ε0, fun y hy ↦ hε <| (dist_comm _ _).trans_lt <| (dist_val_le_dist _ _).trans_lt hy⟩ · rcases h x hx with ⟨ε, ε0, hε⟩ simp only [dist_eq, one_div] at hε - have : Tendsto (fun y : s ↦ dist x.1 y.1 + |(infDist x.1 (sᶜ))⁻¹ - (infDist y.1 (sᶜ))⁻¹|) - (𝓝 x) (𝓝 (dist x.1 x.1 + |(infDist x.1 (sᶜ))⁻¹ - (infDist x.1 (sᶜ))⁻¹|)) + have : Tendsto (fun y : s ↦ dist x.1 y.1 + |(infDist x.1 sᶜ)⁻¹ - (infDist y.1 sᶜ)⁻¹|) + (𝓝 x) (𝓝 (dist x.1 x.1 + |(infDist x.1 sᶜ)⁻¹ - (infDist x.1 sᶜ)⁻¹|)) · refine (tendsto_const_nhds.dist continuous_subtype_val.continuousAt).add (tendsto_const_nhds.sub <| ?_).abs refine (continuousAt_inv_infDist_pt ?_).comp continuous_subtype_val.continuousAt @@ -328,23 +328,23 @@ instance instCompleteSpace [CompleteSpace α] : CompleteSpace (CompleteCopy s) : obtain ⟨x, xlim⟩ : ∃ x, Tendsto (fun n => (u n).1) atTop (𝓝 x) := cauchySeq_tendsto_of_complete A by_cases xs : x ∈ s · exact ⟨⟨x, xs⟩, tendsto_subtype_rng.2 xlim⟩ - obtain ⟨C, hC⟩ : ∃ C, ∀ n, 1 / infDist (u n).1 (sᶜ) < C - · refine ⟨(1 / 2) ^ 0 + 1 / infDist (u 0).1 (sᶜ), fun n ↦ ?_⟩ + obtain ⟨C, hC⟩ : ∃ C, ∀ n, 1 / infDist (u n).1 sᶜ < C + · refine ⟨(1 / 2) ^ 0 + 1 / infDist (u 0).1 sᶜ, fun n ↦ ?_⟩ rw [← sub_lt_iff_lt_add] calc - _ ≤ |1 / infDist (u n).1 (sᶜ) - 1 / infDist (u 0).1 (sᶜ)| := le_abs_self _ - _ = |1 / infDist (u 0).1 (sᶜ) - 1 / infDist (u n).1 (sᶜ)| := abs_sub_comm _ _ + _ ≤ |1 / infDist (u n).1 sᶜ - 1 / infDist (u 0).1 sᶜ| := le_abs_self _ + _ = |1 / infDist (u 0).1 sᶜ - 1 / infDist (u n).1 sᶜ| := abs_sub_comm _ _ _ ≤ dist (u 0) (u n) := le_add_of_nonneg_left dist_nonneg _ < (1 / 2) ^ 0 := hu 0 0 n le_rfl n.zero_le have Cpos : 0 < C := lt_of_le_of_lt (div_nonneg zero_le_one infDist_nonneg) (hC 0) - have Hmem : ∀ {y}, y ∈ s ↔ 0 < infDist y (sᶜ) := fun {y} ↦ by + have Hmem : ∀ {y}, y ∈ s ↔ 0 < infDist y sᶜ := fun {y} ↦ by rw [← s.isOpen.isClosed_compl.not_mem_iff_infDist_pos ⟨x, xs⟩]; exact not_not.symm - have I : ∀ n, 1 / C ≤ infDist (u n).1 (sᶜ) := fun n ↦ by - have : 0 < infDist (u n).1 (sᶜ) := Hmem.1 (u n).2 + have I : ∀ n, 1 / C ≤ infDist (u n).1 sᶜ := fun n ↦ by + have : 0 < infDist (u n).1 sᶜ := Hmem.1 (u n).2 rw [div_le_iff' Cpos] exact (div_le_iff this).1 (hC n).le - have I' : 1 / C ≤ infDist x (sᶜ) := - have : Tendsto (fun n => infDist (u n).1 (sᶜ)) atTop (𝓝 (infDist x (sᶜ))) := + have I' : 1 / C ≤ infDist x sᶜ := + have : Tendsto (fun n => infDist (u n).1 sᶜ) atTop (𝓝 (infDist x sᶜ)) := ((continuous_infDist_pt (sᶜ : Set α)).tendsto x).comp xlim ge_of_tendsto' this I exact absurd (Hmem.2 <| lt_of_lt_of_le (div_pos one_pos Cpos) I') xs @@ -398,7 +398,7 @@ theorem _root_.IsClosed.isClopenable [TopologicalSpace α] [PolishSpace α] {s : #align is_closed.is_clopenable IsClosed.isClopenable theorem IsClopenable.compl [TopologicalSpace α] {s : Set α} (hs : IsClopenable s) : - IsClopenable (sᶜ) := by + IsClopenable sᶜ := by rcases hs with ⟨t, t_le, t_polish, h, h'⟩ exact ⟨t, t_le, t_polish, @IsOpen.isClosed_compl α t s h', @IsClosed.isOpen_compl α t s h⟩ #align polish_space.is_clopenable.compl PolishSpace.IsClopenable.compl diff --git a/Mathlib/Topology/Order/LowerTopology.lean b/Mathlib/Topology/Order/LowerTopology.lean index eb3a6e990d4f2..05e82e462e2f9 100644 --- a/Mathlib/Topology/Order/LowerTopology.lean +++ b/Mathlib/Topology/Order/LowerTopology.lean @@ -119,17 +119,17 @@ instance : Preorder (WithLowerTopology α) := ‹Preorder α› instance : TopologicalSpace (WithLowerTopology α) := - generateFrom { s | ∃ a, Ici aᶜ = s } + generateFrom { s | ∃ a, (Ici a)ᶜ = s } theorem isOpen_preimage_ofLower (S : Set α) : IsOpen (WithLowerTopology.ofLower ⁻¹' S) ↔ - (generateFrom { s : Set α | ∃ a : α, Ici aᶜ = s }).IsOpen S := + (generateFrom { s : Set α | ∃ a : α, (Ici a)ᶜ = s }).IsOpen S := Iff.rfl #align with_lower_topology.is_open_preimage_of_lower WithLowerTopology.isOpen_preimage_ofLower theorem isOpen_def (T : Set (WithLowerTopology α)) : - IsOpen T ↔ - (generateFrom { s : Set α | ∃ a : α, Ici aᶜ = s }).IsOpen (WithLowerTopology.toLower ⁻¹' T) := + IsOpen T ↔ (generateFrom { s : Set α | ∃ a : α, (Ici a)ᶜ = s }).IsOpen + (WithLowerTopology.toLower ⁻¹' T) := Iff.rfl #align with_lower_topology.is_open_def WithLowerTopology.isOpen_def @@ -139,7 +139,7 @@ end WithLowerTopology The lower topology is the topology generated by the complements of the closed intervals to infinity. -/ class LowerTopology (α : Type _) [t : TopologicalSpace α] [Preorder α] : Prop where - topology_eq_lowerTopology : t = generateFrom { s | ∃ a, Ici aᶜ = s } + topology_eq_lowerTopology : t = generateFrom { s | ∃ a, (Ici a)ᶜ = s } #align lower_topology LowerTopology instance [Preorder α] : LowerTopology (WithLowerTopology α) := @@ -168,7 +168,7 @@ def withLowerTopologyHomeomorph : WithLowerTopology α ≃ₜ α := WithLowerTopology.ofLower.toHomeomorphOfInducing ⟨by erw [topology_eq α, induced_id]; rfl⟩ #align lower_topology.with_lower_topology_homeomorph LowerTopology.withLowerTopologyHomeomorph -theorem isOpen_iff_generate_Ici_compl : IsOpen s ↔ GenerateOpen { t | ∃ a, Ici aᶜ = t } s := by +theorem isOpen_iff_generate_Ici_compl : IsOpen s ↔ GenerateOpen { t | ∃ a, (Ici a)ᶜ = t } s := by rw [topology_eq α]; rfl #align lower_topology.is_open_iff_generate_Ici_compl LowerTopology.isOpen_iff_generate_Ici_compl @@ -214,7 +214,7 @@ protected theorem isTopologicalBasis : IsTopologicalBasis (lowerBasis α) := by ext s constructor · rintro ⟨F, hF, rfl⟩ - refine' ⟨(fun a => Ici aᶜ) '' F, ⟨hF.image _, image_subset_iff.2 fun _ _ => ⟨_, rfl⟩⟩, _⟩ + refine' ⟨(fun a => (Ici a)ᶜ) '' F, ⟨hF.image _, image_subset_iff.2 fun _ _ => ⟨_, rfl⟩⟩, _⟩ simp only [sInter_image] · rintro ⟨F, ⟨hF, hs⟩, rfl⟩ haveI := hF.to_subtype @@ -262,7 +262,7 @@ instance [Preorder α] [TopologicalSpace α] [LowerTopology α] [OrderBot α] [P dsimp simp_rw [coe_upperClosure, compl_iUnion, prod_eq, preimage_iInter, preimage_compl] -- without `let`, `refine` tries to use the product topology and fails - let _ : TopologicalSpace (α × β) := generateFrom { s | ∃ a, Ici aᶜ = s } + let _ : TopologicalSpace (α × β) := generateFrom { s | ∃ a, (Ici a)ᶜ = s } refine (isOpen_biInter hs fun a _ => ?_).inter (isOpen_biInter ht fun b _ => ?_) · exact GenerateOpen.basic _ ⟨(a, ⊥), by simp [Ici_prod_eq, prod_univ]⟩ · exact GenerateOpen.basic _ ⟨(⊥, b), by simp [Ici_prod_eq, univ_prod]⟩ diff --git a/Mathlib/Topology/Paracompact.lean b/Mathlib/Topology/Paracompact.lean index ce218946fbe5b..6e1309fbeacdb 100644 --- a/Mathlib/Topology/Paracompact.lean +++ b/Mathlib/Topology/Paracompact.lean @@ -103,11 +103,11 @@ theorem precise_refinement_set [ParacompactSpace X] {s : Set X} (hs : IsClosed s (uo : ∀ i, IsOpen (u i)) (us : s ⊆ ⋃ i, u i) : ∃ v : ι → Set X, (∀ i, IsOpen (v i)) ∧ (s ⊆ ⋃ i, v i) ∧ LocallyFinite v ∧ ∀ i, v i ⊆ u i := by -- Porting note: Added proof of uc - have uc : (iUnion fun i => Option.elim' (sᶜ) u i) = univ := by + have uc : (iUnion fun i => Option.elim' sᶜ u i) = univ := by apply Subset.antisymm (subset_univ _) · simp_rw [← compl_union_self s, Option.elim', iUnion_option] - apply union_subset_union_right (sᶜ) us - rcases precise_refinement (Option.elim' (sᶜ) u) (Option.forall.2 ⟨isOpen_compl_iff.2 hs, uo⟩) + apply union_subset_union_right sᶜ us + rcases precise_refinement (Option.elim' sᶜ u) (Option.forall.2 ⟨isOpen_compl_iff.2 hs, uo⟩) uc with ⟨v, vo, vc, vf, vu⟩ refine' ⟨v ∘ some, fun i ↦ vo _, _, vf.comp_injective (Option.some_injective _), fun i ↦ vu _⟩ @@ -216,7 +216,7 @@ theorem refinement_of_locallyCompact_sigmaCompact_of_nhds_basis_set [LocallyComp fun n ↦ ((K.isCompact _).diff isOpen_interior).inter_right hs -- Next we choose a finite covering `B (c n i) (r n i)` of each -- `Kdiff (n + 1) ∩ s` such that `B (c n i) (r n i) ∩ s` is disjoint with `K n` - have : ∀ (n) (x : ↑(Kdiff (n + 1) ∩ s)), K nᶜ ∈ 𝓝 (x : X) := + have : ∀ (n) (x : ↑(Kdiff (n + 1) ∩ s)), (K n)ᶜ ∈ 𝓝 (x : X) := fun n x ↦ IsOpen.mem_nhds (K.isClosed n).isOpen_compl fun hx' ↦ x.2.1.2 <| K.subset_interior_succ _ hx' -- Porting note: Commented out `haveI` for now. @@ -306,7 +306,7 @@ theorem normal_of_paracompact_t2 [T2Space X] [ParacompactSpace X] : NormalSpace choose u v hu hv hxu htv huv using SetCoe.forall'.1 H rcases precise_refinement_set hs u hu fun x hx ↦ mem_iUnion.2 ⟨⟨x, hx⟩, hxu _⟩ with ⟨u', hu'o, hcov', hu'fin, hsub⟩ - refine' ⟨⋃ i, u' i, closure (⋃ i, u' i)ᶜ, isOpen_iUnion hu'o, isClosed_closure.isOpen_compl, + refine' ⟨⋃ i, u' i, (closure (⋃ i, u' i))ᶜ, isOpen_iUnion hu'o, isClosed_closure.isOpen_compl, hcov', _, disjoint_compl_right.mono le_rfl (compl_le_compl subset_closure)⟩ rw [hu'fin.closure_iUnion, compl_iUnion, subset_iInter_iff] refine' fun i x hxt hxu ↦ diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index d829a85c05465..edcbcaa2071e8 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -681,7 +681,7 @@ theorem compl_singleton_mem_nhdsSet_iff [T1Space α] {x : α} {s : Set α} : {x} theorem nhdsSet_le_iff [T1Space α] {s t : Set α} : 𝓝ˢ s ≤ 𝓝ˢ t ↔ s ⊆ t := by refine' ⟨_, fun h => monotone_nhdsSet h⟩ simp_rw [Filter.le_def]; intro h x hx - specialize h ({x}ᶜ) + specialize h {x}ᶜ simp_rw [compl_singleton_mem_nhdsSet_iff] at h by_contra hxt exact h hxt hx @@ -1720,7 +1720,7 @@ theorem normal_separation [NormalSpace α] {s t : Set α} (H1 : IsClosed s) (H2 theorem normal_exists_closure_subset [NormalSpace α] {s t : Set α} (hs : IsClosed s) (ht : IsOpen t) (hst : s ⊆ t) : ∃ u, IsOpen u ∧ s ⊆ u ∧ closure u ⊆ t := by - have : Disjoint s (tᶜ) := Set.disjoint_left.mpr fun x hxs hxt => hxt (hst hxs) + have : Disjoint s tᶜ := Set.disjoint_left.mpr fun x hxs hxt => hxt (hst hxs) rcases normal_separation hs (isClosed_compl_iff.2 ht) this with ⟨s', t', hs', ht', hss', htt', hs't'⟩ refine ⟨s', hs', hss', Subset.trans (closure_minimal ?_ (isClosed_compl_iff.2 ht')) @@ -1952,7 +1952,7 @@ theorem compact_t2_tot_disc_iff_tot_sep : TotallyDisconnectedSpace α ↔ Totall rw [connectedComponent_eq_iInter_clopen, mem_iInter] rintro ⟨w : Set α, hw : IsClopen w, hy : y ∈ w⟩ by_contra hx - exact hyp (wᶜ) w hw.2.isOpen_compl hw.1 hx hy (@isCompl_compl _ w _).symm.codisjoint.top_le + exact hyp wᶜ w hw.2.isOpen_compl hw.1 hx hy (@isCompl_compl _ w _).symm.codisjoint.top_le disjoint_compl_left #align compact_t2_tot_disc_iff_tot_sep compact_t2_tot_disc_iff_tot_sep diff --git a/Mathlib/Topology/Sets/Closeds.lean b/Mathlib/Topology/Sets/Closeds.lean index 888402e5e5022..0fcce8297fa52 100644 --- a/Mathlib/Topology/Sets/Closeds.lean +++ b/Mathlib/Topology/Sets/Closeds.lean @@ -338,7 +338,7 @@ instance : BooleanAlgebra (Clopens α) := @[simp] theorem coe_sdiff (s t : Clopens α) : (↑(s \ t) : Set α) = ↑s \ ↑t := rfl #align topological_space.clopens.coe_sdiff TopologicalSpace.Clopens.coe_sdiff -@[simp] theorem coe_compl (s : Clopens α) : (↑(sᶜ) : Set α) = ↑sᶜ := rfl +@[simp] theorem coe_compl (s : Clopens α) : (↑sᶜ : Set α) = (↑s)ᶜ := rfl #align topological_space.clopens.coe_compl TopologicalSpace.Clopens.coe_compl instance : Inhabited (Clopens α) := ⟨⊥⟩ diff --git a/Mathlib/Topology/Sets/Compacts.lean b/Mathlib/Topology/Sets/Compacts.lean index 9f96c9d80ea94..0509e0d295ad7 100644 --- a/Mathlib/Topology/Sets/Compacts.lean +++ b/Mathlib/Topology/Sets/Compacts.lean @@ -565,7 +565,7 @@ theorem coe_sdiff [T2Space α] (s t : CompactOpens α) : (↑(s \ t) : Set α) = #align topological_space.compact_opens.coe_sdiff TopologicalSpace.CompactOpens.coe_sdiff @[simp] -theorem coe_compl [T2Space α] [CompactSpace α] (s : CompactOpens α) : (↑(sᶜ) : Set α) = ↑sᶜ := +theorem coe_compl [T2Space α] [CompactSpace α] (s : CompactOpens α) : (↑sᶜ : Set α) = (↑s)ᶜ := rfl #align topological_space.compact_opens.coe_compl TopologicalSpace.CompactOpens.coe_compl diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index d8e69c900e69d..a3ede6487d40a 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -176,12 +176,12 @@ theorem le_chainSup {c : Set (PartialRefinement u s)} (hc : IsChain (· ≤ ·) `i ∉ v.carrier`, then there exists a partial refinement that is strictly greater than `v`. -/ theorem exists_gt (v : PartialRefinement u s) (hs : IsClosed s) (i : ι) (hi : i ∉ v.carrier) : ∃ v' : PartialRefinement u s, v < v' := by - have I : (s ∩ ⋂ (j) (_ : j ≠ i), v jᶜ) ⊆ v i := by + have I : (s ∩ ⋂ (j) (_ : j ≠ i), (v j)ᶜ) ⊆ v i := by simp only [subset_def, mem_inter_iff, mem_iInter, and_imp] intro x hxs H rcases mem_iUnion.1 (v.subset_iUnion hxs) with ⟨j, hj⟩ exact (em (j = i)).elim (fun h => h ▸ hj) fun h => (H j h hj).elim - have C : IsClosed (s ∩ ⋂ (j) (_ : j ≠ i), v jᶜ) := + have C : IsClosed (s ∩ ⋂ (j) (_ : j ≠ i), (v j)ᶜ) := IsClosed.inter hs (isClosed_biInter fun _ _ => isClosed_compl_iff.2 <| v.isOpen _) rcases normal_exists_closure_subset C (v.isOpen i) I with ⟨vi, ovi, hvi, cvi⟩ refine' ⟨⟨update v i vi, insert i v.carrier, _, _, _, _⟩, _, _⟩ diff --git a/Mathlib/Topology/StoneCech.lean b/Mathlib/Topology/StoneCech.lean index 23b6594f3f757..637936fb965f8 100644 --- a/Mathlib/Topology/StoneCech.lean +++ b/Mathlib/Topology/StoneCech.lean @@ -60,7 +60,7 @@ theorem ultrafilter_isOpen_basic (s : Set α) : IsOpen { u : Ultrafilter α | s /-- The basic open sets for the topology on ultrafilters are also closed. -/ theorem ultrafilter_isClosed_basic (s : Set α) : IsClosed { u : Ultrafilter α | s ∈ u } := by rw [← isOpen_compl_iff] - convert ultrafilter_isOpen_basic (sᶜ) using 1 + convert ultrafilter_isOpen_basic sᶜ using 1 ext u exact Ultrafilter.compl_mem_iff_not_mem.symm #align ultrafilter_is_closed_basic ultrafilter_isClosed_basic diff --git a/Mathlib/Topology/SubsetProperties.lean b/Mathlib/Topology/SubsetProperties.lean index 42db92a0b46ec..acd1866b8160b 100644 --- a/Mathlib/Topology/SubsetProperties.lean +++ b/Mathlib/Topology/SubsetProperties.lean @@ -100,7 +100,7 @@ theorem IsCompact.induction_on {s : Set α} (hs : IsCompact s) {p : Set α → P (hmono : ∀ ⦃s t⦄, s ⊆ t → p t → p s) (hunion : ∀ ⦃s t⦄, p s → p t → p (s ∪ t)) (hnhds : ∀ x ∈ s, ∃ t ∈ 𝓝[s] x, p t) : p s := by let f : Filter α := - { sets := { t | p (tᶜ) } + { sets := { t | p tᶜ } univ_sets := by simpa sets_of_superset := fun ht₁ ht => hmono (compl_subset_compl.2 ht) ht₁ inter_sets := fun ht₁ ht₂ => by simp [compl_inter, hunion ht₁ ht₂] } @@ -155,8 +155,8 @@ theorem IsCompact.image {f : α → β} (hs : IsCompact s) (hf : Continuous f) : theorem IsCompact.adherence_nhdset {f : Filter α} (hs : IsCompact s) (hf₂ : f ≤ 𝓟 s) (ht₁ : IsOpen t) (ht₂ : ∀ a ∈ s, ClusterPt a f → a ∈ t) : t ∈ f := - Classical.by_cases mem_of_eq_bot fun (this : f ⊓ 𝓟 (tᶜ) ≠ ⊥) => - let ⟨a, ha, (hfa : ClusterPt a <| f ⊓ 𝓟 (tᶜ))⟩ := @hs _ ⟨this⟩ <| inf_le_of_left_le hf₂ + Classical.by_cases mem_of_eq_bot fun (this : f ⊓ 𝓟 tᶜ ≠ ⊥) => + let ⟨a, ha, (hfa : ClusterPt a <| f ⊓ 𝓟 tᶜ)⟩ := @hs _ ⟨this⟩ <| inf_le_of_left_le hf₂ have : a ∈ t := ht₂ a ha hfa.of_inf_left have : tᶜ ∩ t ∈ 𝓝[tᶜ] a := inter_mem_nhdsWithin _ (IsOpen.mem_nhds ht₁ this) have A : 𝓝[tᶜ] a = ⊥ := empty_mem_iff_bot.1 <| compl_inter_self t ▸ this @@ -520,7 +520,7 @@ namespace Filter /-- `Filter.cocompact` is the filter generated by complements to compact sets. -/ def cocompact (α : Type _) [TopologicalSpace α] : Filter α := - ⨅ (s : Set α) (_ : IsCompact s), 𝓟 (sᶜ) + ⨅ (s : Set α) (_ : IsCompact s), 𝓟 sᶜ #align filter.cocompact Filter.cocompact theorem hasBasis_cocompact : (cocompact α).HasBasis IsCompact compl := @@ -588,7 +588,7 @@ theorem Tendsto.isCompact_insert_range {f : ℕ → α} {a} (hf : Tendsto f atTo /-- `Filter.coclosedCompact` is the filter generated by complements to closed compact sets. In a Hausdorff space, this is the same as `Filter.cocompact`. -/ def coclosedCompact (α : Type _) [TopologicalSpace α] : Filter α := - ⨅ (s : Set α) (_ : IsClosed s) (_ : IsCompact s), 𝓟 (sᶜ) + ⨅ (s : Set α) (_ : IsClosed s) (_ : IsCompact s), 𝓟 sᶜ #align filter.coclosed_compact Filter.coclosedCompact theorem hasBasis_coclosedCompact : @@ -1273,7 +1273,7 @@ theorem IsClosed.exists_minimal_nonempty_closed_subset [CompactSpace α] {S : Se refine' ⟨Uᶜ, Set.compl_subset_comm.mp Uc, Ucne, Uo.isClosed_compl, _⟩ intro V' V'sub V'ne V'cls have : V'ᶜ = U := by - refine' h (V'ᶜ) ⟨_, isOpen_compl_iff.mpr V'cls, _⟩ (Set.subset_compl_comm.mp V'sub) + refine' h V'ᶜ ⟨_, isOpen_compl_iff.mpr V'cls, _⟩ (Set.subset_compl_comm.mp V'sub) exact Set.Subset.trans Uc (Set.subset_compl_comm.mp V'sub) simp only [compl_compl, V'ne] rw [← this, compl_compl] @@ -1583,12 +1583,12 @@ theorem IsClopen.inter {s t : Set α} (hs : IsClopen s) (ht : IsClopen t) : IsCl @[simp] theorem isClopen_univ : IsClopen (univ : Set α) := ⟨isOpen_univ, isClosed_univ⟩ #align is_clopen_univ isClopen_univ -theorem IsClopen.compl {s : Set α} (hs : IsClopen s) : IsClopen (sᶜ) := +theorem IsClopen.compl {s : Set α} (hs : IsClopen s) : IsClopen sᶜ := ⟨hs.2.isOpen_compl, hs.1.isClosed_compl⟩ #align is_clopen.compl IsClopen.compl @[simp] -theorem isClopen_compl_iff {s : Set α} : IsClopen (sᶜ) ↔ IsClopen s := +theorem isClopen_compl_iff {s : Set α} : IsClopen sᶜ ↔ IsClopen s := ⟨fun h => compl_compl s ▸ IsClopen.compl h, IsClopen.compl⟩ #align is_clopen_compl_iff isClopen_compl_iff @@ -1971,7 +1971,7 @@ theorem subset_closure_inter_of_isPreirreducible_of_isOpen {S U : Set α} (hS : (hU : IsOpen U) (h : (S ∩ U).Nonempty) : S ⊆ closure (S ∩ U) := by by_contra h' obtain ⟨x, h₁, h₂, h₃⟩ := - hS _ (closure (S ∩ U)ᶜ) hU isClosed_closure.isOpen_compl h (inter_compl_nonempty_iff.mpr h') + hS _ (closure (S ∩ U))ᶜ hU isClosed_closure.isOpen_compl h (inter_compl_nonempty_iff.mpr h') exact h₃ (subset_closure ⟨h₁, h₂⟩) #align subset_closure_inter_of_is_preirreducible_of_is_open subset_closure_inter_of_isPreirreducible_of_isOpen diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index c6a517874c27a..33891fc87fc20 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -318,7 +318,7 @@ theorem exists_finset_nhd_mulSupport_subset {f : ι → X → R} let is := hnf.toFinset.filter fun i => x ∈ U i let js := hnf.toFinset.filter fun j => x ∉ U j refine' - ⟨is, (n ∩ ⋂ j ∈ js, mulTSupport (f j)ᶜ) ∩ ⋂ i ∈ is, U i, inter_mem (inter_mem hn _) _, + ⟨is, (n ∩ ⋂ j ∈ js, (mulTSupport (f j))ᶜ) ∩ ⋂ i ∈ is, U i, inter_mem (inter_mem hn _) _, inter_subset_right _ _, fun z hz => _⟩ · exact (biInter_finset_mem js).mpr fun j hj => IsClosed.compl_mem_nhds (isClosed_mulTSupport _) (Set.not_mem_subset (hso j) (Finset.mem_filter.mp hj).2) diff --git a/Mathlib/Topology/UniformSpace/Compact.lean b/Mathlib/Topology/UniformSpace/Compact.lean index c5fd432d0291b..1cfbf9ad182c0 100644 --- a/Mathlib/Topology/UniformSpace/Compact.lean +++ b/Mathlib/Topology/UniformSpace/Compact.lean @@ -95,14 +95,14 @@ def uniformSpaceOfCompactT2 [TopologicalSpace γ] [CompactSpace γ] [T2Space γ] rw [le_iff_forall_inf_principal_compl] intro V V_in by_contra H - haveI : NeBot (F ⊓ 𝓟 (Vᶜ)) := ⟨H⟩ + haveI : NeBot (F ⊓ 𝓟 Vᶜ) := ⟨H⟩ -- Hence compactness would give us a cluster point (x, y) for F ⊓ 𝓟 Vᶜ - obtain ⟨⟨x, y⟩, hxy⟩ : ∃ p : γ × γ, ClusterPt p (F ⊓ 𝓟 (Vᶜ)) := cluster_point_of_compact _ + obtain ⟨⟨x, y⟩, hxy⟩ : ∃ p : γ × γ, ClusterPt p (F ⊓ 𝓟 Vᶜ) := cluster_point_of_compact _ -- In particular (x, y) is a cluster point of 𝓟 Vᶜ, hence is not in the interior of V, -- and a fortiori not in Δ, so x ≠ y have clV : ClusterPt (x, y) (𝓟 <| Vᶜ) := hxy.of_inf_right have : (x, y) ∉ interior V := by - have : (x, y) ∈ closure (Vᶜ) := by rwa [mem_closure_iff_clusterPt] + have : (x, y) ∈ closure Vᶜ := by rwa [mem_closure_iff_clusterPt] rwa [closure_compl] at this have diag_subset : diagonal γ ⊆ interior V := subset_interior_iff_mem_nhdsSet.2 V_in have x_ne_y : x ≠ y := mt (@diag_subset (x, y)) this diff --git a/Mathlib/Topology/UrysohnsLemma.lean b/Mathlib/Topology/UrysohnsLemma.lean index 4c82d59271ee5..509b82842b7f9 100644 --- a/Mathlib/Topology/UrysohnsLemma.lean +++ b/Mathlib/Topology/UrysohnsLemma.lean @@ -134,7 +134,7 @@ theorem subset_right_C (c : CU X) : c.C ⊆ c.right.C := /-- `n`-th approximation to a continuous function `f : X → ℝ` such that `f = 0` on `c.C` and `f = 1` outside of `c.U`. -/ noncomputable def approx : ℕ → CU X → X → ℝ - | 0, c, x => indicator (c.Uᶜ) 1 x + | 0, c, x => indicator c.Uᶜ 1 x | n + 1, c, x => midpoint ℝ (approx n c.left x) (approx n c.right x) #align urysohns.CU.approx Urysohns.CU.approx