Skip to content

Commit

Permalink
chore: bump Std, changes for leanprover-community/batteries#366
Browse files Browse the repository at this point in the history
revert

.

fix a bit

blech

one more

comment out broken and unused results

chore: bump Std

remove redundant simp lemmas

cleanup

fix: make `Nat.testBit_bitwise` a simp-lemma again.

This fixes the Int.testbit_bitwise proofs

fix `castNum_testBit`

fix `testBit_eq_inth`

fixes

golf

golf

Apply suggestions from code review

Co-authored-by: Eric Wieser <wieser.eric@gmail.com>

fix whitespace

fix name conflict

more efficient

restore lemma

feat: Mapping extreme points under continuous maps (#8574)

Prove that extreme points are preserved under affine equivalences, and the less trivial statement that a continuous affine map sends extreme points of a compact set to a superset of the extreme
points of the image of that set.

Also fix a few name and tweak the API a bit.

Co-authored-by: Yury G. Kudryashov <urkud@urkud.name>

chore(Profinite): allow more universe flexibility in Profinite/CofilteredLimit (#8613)

We allow the indexing category for the cofiltered limit in the result `Profinite.exists_locallyConstant` to live in a smaller universe than our profinite sets.

feat(CategoryTheory): `Preregular` and `FinitaryPreExtensive` implies `Precoherent` (#8643)

We prove some results about effective epimorphisms which allow us to deduce that a category which is `FinitaryPreExtensive` and `Preregular`Β is also `Precoherent`.

feat(Analysis/Seminorm): add more `*ball_smul_*ball` (#8724)

We had `ball_smul_ball` and `closedBall_smul_closedBall`.
Add versions with mixed `ball` and `closedBall`.
Also move this lemmas below and golf the proofs.

feat: the annihilator of the kernel of the trace form of a Lie module is contained in the span of its weights (#8739)

fix breakage in BitVec lemmas
  • Loading branch information
kim-em authored and alexkeizer committed Nov 30, 2023
1 parent 6f22008 commit 7cd93d7
Show file tree
Hide file tree
Showing 26 changed files with 437 additions and 286 deletions.
24 changes: 13 additions & 11 deletions Mathlib/Algebra/Lie/Killing.lean
Original file line number Diff line number Diff line change
Expand Up @@ -242,7 +242,8 @@ lemma traceForm_eq_sum_weightSpaceOf [IsTriangularizable R L M] (z : L) :
LinearMap.trace_eq_sum_trace_restrict' hds hfin hxy]
exact Finset.sum_congr (by simp) (fun Ο‡ _ ↦ rfl)

-- In characteristic zero a stronger result holds (no `βŠ“ LieAlgebra.center K L`) TODO prove this!
-- In characteristic zero (or even just `LinearWeights R L M`) a stronger result holds (no
-- `βŠ“ LieAlgebra.center R L`) TODO prove this using `LieModule.traceForm_eq_sum_finrank_nsmul_mul`.
lemma lowerCentralSeries_one_inf_center_le_ker_traceForm :
lowerCentralSeries R L L 1 βŠ“ LieAlgebra.center R L ≀ LinearMap.ker (traceForm R L M) := by
/- Sketch of proof (due to Zassenhaus):
Expand Down Expand Up @@ -490,32 +491,33 @@ lemma LieModule.traceForm_eq_sum_finrank_nsmul_mul [LieAlgebra.IsNilpotent K L]
LinearMap.trace_eq_sum_trace_restrict' hds hfin hxy]
exact Finset.sum_congr (by simp) (fun Ο‡ _ ↦ traceForm_weightSpace_eq K L M Ο‡ x y)

@[simp]
lemma LieModule.span_weight_eq_top_of_ker_traceForm_eq_bot [LieAlgebra.IsNilpotent K L]
[LinearWeights K L M] [IsTriangularizable K L M] [FiniteDimensional K L]
(h : LinearMap.ker (traceForm K L M) = βŠ₯) :
span K (range (weight.toLinear K L M)) = ⊀ := by
-- The reverse inclusion should also hold: TODO prove this!
lemma LieModule.dualAnnihilator_ker_traceForm_le_span_weight [LieAlgebra.IsNilpotent K L]
[LinearWeights K L M] [IsTriangularizable K L M] [FiniteDimensional K L] :
(LinearMap.ker (traceForm K L M)).dualAnnihilator ≀ span K (range (weight.toLinear K L M)) := by
intro g hg
simp only [Submodule.mem_dualAnnihilator, LinearMap.mem_ker] at hg
by_contra contra
obtain ⟨f : Module.Dual K (Module.Dual K L), hf, hf'⟩ :=
Module.exists_dual_map_eq_bot_of_lt_top K (lt_top_iff_ne_top.mpr contra) inferInstance
Submodule.exists_dual_map_eq_bot_of_nmem contra inferInstance
let x : L := (Module.evalEquiv K L).symm f
replace hf' : βˆ€ Ο‡ ∈ weight K L M, Ο‡ x = 0 := by
intro χ hχ
change weight.toLinear K L M βŸ¨Ο‡, hΟ‡βŸ© x = 0
rw [Module.apply_evalEquiv_symm_apply, ← Submodule.mem_bot (R := K), ← hf', Submodule.mem_map]
exact ⟨weight.toLinear K L M βŸ¨Ο‡, hΟ‡βŸ©, Submodule.subset_span (mem_range_self _), rfl⟩
have hx : x β‰  0 := (LinearEquiv.map_ne_zero_iff _).mpr hf
apply hx
rw [← Submodule.mem_bot (R := K), ← h, LinearMap.mem_ker]
have hx : g x β‰  0 := by simpa
refine hx (hg _ ?_)
ext y
rw [LieModule.traceForm_eq_sum_finrank_nsmul_mul, LinearMap.zero_apply]
exact Finset.sum_eq_zero fun Ο‡ hΟ‡ ↦ by simp [hf' Ο‡ hΟ‡]

/-- Given a splitting Cartan subalgebra `H` of a finite-dimensional Lie algebra with non-singular
Killing form, the corresponding roots span the dual space of `H`. -/
@[simp]
lemma LieAlgebra.IsKilling.span_weight_eq_top [FiniteDimensional K L] [IsKilling K L]
(H : LieSubalgebra K L) [H.IsCartanSubalgebra] [IsTriangularizable K H L] :
span K (range (weight.toLinear K H L)) = ⊀ := by
simp
simpa using LieModule.dualAnnihilator_ker_traceForm_le_span_weight K H L

end Field
25 changes: 18 additions & 7 deletions Mathlib/Algebra/Module/LinearMap.lean
Original file line number Diff line number Diff line change
Expand Up @@ -62,10 +62,7 @@ open Function

universe u u' v w x y z

variable {R : Type*} {R₁ : Type*} {Rβ‚‚ : Type*} {R₃ : Type*}
variable {k : Type*} {S : Type*} {S₃ : Type*} {T : Type*}
variable {M : Type*} {M₁ : Type*} {Mβ‚‚ : Type*} {M₃ : Type*}
variable {N₁ : Type*} {Nβ‚‚ : Type*} {N₃ : Type*} {ΞΉ : Type*}
variable {R R₁ Rβ‚‚ R₃ k S S₃ T M M₁ Mβ‚‚ M₃ N₁ Nβ‚‚ N₃ ΞΉ : Type*}

/-- A map `f` between modules over a semiring is linear if it satisfies the two properties
`f (x + y) = f x + f y` and `f (c β€’ x) = c β€’ f x`. The predicate `IsLinearMap R f` asserts this
Expand All @@ -87,7 +84,7 @@ is semilinear if it satisfies the two properties `f (x + y) = f x + f y` and
`M β†’β‚›β‚—[Οƒ] Mβ‚‚`) are bundled versions of such maps. For plain linear maps (i.e. for which
`Οƒ = RingHom.id R`), the notation `M β†’β‚—[R] Mβ‚‚` is available. An unbundled version of plain linear
maps is available with the predicate `IsLinearMap`, but it should be avoided most of the time. -/
structure LinearMap {R : Type*} {S : Type*} [Semiring R] [Semiring S] (Οƒ : R β†’+* S) (M : Type*)
structure LinearMap {R S : Type*} [Semiring R] [Semiring S] (Οƒ : R β†’+* S) (M : Type*)
(Mβ‚‚ : Type*) [AddCommMonoid M] [AddCommMonoid Mβ‚‚] [Module R M] [Module S Mβ‚‚] extends
AddHom M Mβ‚‚ where
/-- A linear map preserves scalar multiplication.
Expand Down Expand Up @@ -150,7 +147,6 @@ namespace SemilinearMapClass
variable (F : Type*)
variable [Semiring R] [Semiring S]
variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid Mβ‚‚] [AddCommMonoid M₃]
variable [AddCommMonoid N₁] [AddCommMonoid Nβ‚‚] [AddCommMonoid N₃]
variable [Module R M] [Module R Mβ‚‚] [Module S M₃]
variable {Οƒ : R β†’+* S}

Expand All @@ -174,8 +170,23 @@ theorem map_smul_inv {Οƒ' : S β†’+* R} [RingHomInvPair Οƒ Οƒ'] (c : S) (x : M) :
c β€’ f x = f (Οƒ' c β€’ x) := by simp
#align semilinear_map_class.map_smul_inv SemilinearMapClass.map_smul_inv

/-- Reinterpret an element of a type of semilinear maps as a semilinear map. -/
abbrev semilinearMap : M β†’β‚›β‚—[Οƒ] M₃ where
toFun := f
map_add' := map_add f
map_smul' := map_smulβ‚›β‚— f

end SemilinearMapClass

namespace LinearMapClass
variable {F : Type*} [Semiring R] [AddCommMonoid M₁] [AddCommMonoid Mβ‚‚] [Module R M₁] [Module R Mβ‚‚]
(f : F) [LinearMapClass F R M₁ Mβ‚‚]

/-- Reinterpret an element of a type of linear maps as a linear map. -/
abbrev linearMap : M₁ β†’β‚—[R] Mβ‚‚ := SemilinearMapClass.semilinearMap f

end LinearMapClass

namespace LinearMap

section AddCommMonoid
Expand Down Expand Up @@ -566,7 +577,7 @@ theorem id_comp : id.comp f = f :=
#align linear_map.id_comp LinearMap.id_comp

theorem comp_assoc
{Rβ‚„ : Type*} {Mβ‚„ : Type*} [Semiring Rβ‚„] [AddCommMonoid Mβ‚„] [Module Rβ‚„ Mβ‚„]
{Rβ‚„ Mβ‚„ : Type*} [Semiring Rβ‚„] [AddCommMonoid Mβ‚„] [Module Rβ‚„ Mβ‚„]
{σ₃₄ : R₃ β†’+* Rβ‚„} {Οƒβ‚‚β‚„ : Rβ‚‚ β†’+* Rβ‚„} {σ₁₄ : R₁ β†’+* Rβ‚„}
[RingHomCompTriple σ₂₃ σ₃₄ Οƒβ‚‚β‚„] [RingHomCompTriple σ₁₃ σ₃₄ σ₁₄] [RingHomCompTriple σ₁₂ Οƒβ‚‚β‚„ σ₁₄]
(f : M₁ β†’β‚›β‚—[σ₁₂] Mβ‚‚) (g : Mβ‚‚ β†’β‚›β‚—[σ₂₃] M₃) (h : M₃ β†’β‚›β‚—[σ₃₄] Mβ‚„) :
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Analysis/Convex/Exposed.lean
Original file line number Diff line number Diff line change
Expand Up @@ -265,7 +265,7 @@ protected theorem isExtreme (hAB : IsExposed π•œ A B) : IsExtreme π•œ A B := b
end IsExposed

theorem exposedPoints_subset_extremePoints : A.exposedPoints π•œ βŠ† A.extremePoints π•œ := fun _ hx =>
mem_extremePoints_iff_extreme_singleton.2 (mem_exposedPoints_iff_exposed_singleton.1 hx).isExtreme
(mem_exposedPoints_iff_exposed_singleton.1 hx).isExtreme.mem_extremePoints
#align exposed_points_subset_extreme_points exposedPoints_subset_extremePoints

end LinearOrderedRing
35 changes: 23 additions & 12 deletions Mathlib/Analysis/Convex/Extreme.lean
Original file line number Diff line number Diff line change
Expand Up @@ -123,12 +123,7 @@ theorem isExtreme_biInter {F : Set (Set E)} (hF : F.Nonempty) (hA : βˆ€ B ∈ F,
#align is_extreme_bInter isExtreme_biInter

theorem isExtreme_sInter {F : Set (Set E)} (hF : F.Nonempty) (hAF : βˆ€ B ∈ F, IsExtreme π•œ A B) :
IsExtreme π•œ A (β‹‚β‚€ F) := by
obtain ⟨B, hB⟩ := hF
refine' ⟨(sInter_subset_of_mem hB).trans (hAF B hB).1, fun x₁ hx₁A xβ‚‚ hxβ‚‚A x hxF hx ↦ _⟩
simp_rw [mem_sInter] at hxF ⊒
have h := fun B hB ↦ (hAF B hB).2 hx₁A hxβ‚‚A (hxF B hB) hx
exact ⟨fun B hB ↦ (h B hB).1, fun B hB ↦ (h B hB).2⟩
IsExtreme π•œ A (β‹‚β‚€ F) := by simpa [sInter_eq_biInter] using isExtreme_biInter hF hAF
#align is_extreme_sInter isExtreme_sInter

theorem mem_extremePoints : x ∈ A.extremePoints π•œ ↔
Expand All @@ -137,13 +132,15 @@ theorem mem_extremePoints : x ∈ A.extremePoints π•œ ↔
#align mem_extreme_points mem_extremePoints

/-- x is an extreme point to A iff {x} is an extreme set of A. -/
theorem mem_extremePoints_iff_extreme_singleton : x ∈ A.extremePoints π•œ ↔ IsExtreme π•œ A {x} := by
refine' ⟨_, fun hx ↦ ⟨singleton_subset_iff.1 hx.1, fun x₁ hx₁ xβ‚‚ hxβ‚‚ ↦ hx.2 hx₁ hxβ‚‚ rfl⟩⟩
@[simp] lemma isExtreme_singleton : IsExtreme π•œ A {x} ↔ x ∈ A.extremePoints π•œ := by
refine ⟨fun hx ↦ ⟨singleton_subset_iff.1 hx.1, fun x₁ hx₁ xβ‚‚ hxβ‚‚ ↦ hx.2 hx₁ hxβ‚‚ rfl⟩, ?_⟩
rintro ⟨hxA, hAx⟩
use singleton_subset_iff.2 hxA
rintro x₁ hx₁A xβ‚‚ hxβ‚‚A y (rfl : y = x)
exact hAx hx₁A hxβ‚‚A
#align mem_extreme_points_iff_extreme_singleton mem_extremePoints_iff_extreme_singleton
#align mem_extreme_points_iff_extreme_singleton isExtreme_singleton

alias ⟨IsExtreme.mem_extremePoints, _⟩ := isExtreme_singleton

theorem extremePoints_subset : A.extremePoints π•œ βŠ† A :=
fun _ hx ↦ hx.1
Expand All @@ -167,8 +164,7 @@ theorem inter_extremePoints_subset_extremePoints_of_subset (hBA : B βŠ† A) :

theorem IsExtreme.extremePoints_subset_extremePoints (hAB : IsExtreme π•œ A B) :
B.extremePoints π•œ βŠ† A.extremePoints π•œ :=
fun _ hx ↦ mem_extremePoints_iff_extreme_singleton.2
(hAB.trans (mem_extremePoints_iff_extreme_singleton.1 hx))
fun _ ↦ by simpa only [←isExtreme_singleton] using hAB.trans
#align is_extreme.extreme_points_subset_extreme_points IsExtreme.extremePoints_subset_extremePoints

theorem IsExtreme.extremePoints_eq (hAB : IsExtreme π•œ A B) :
Expand Down Expand Up @@ -236,6 +232,21 @@ theorem extremePoints_pi (s : βˆ€ i, Set (Ο€ i)) :

end OrderedSemiring

section OrderedRing
variable {L : Type*} [OrderedRing π•œ] [AddCommGroup E] [Module π•œ E] [AddCommGroup F] [Module π•œ F]
[LinearEquivClass L π•œ E F]

lemma image_extremePoints (f : L) (s : Set E) :
f '' extremePoints π•œ s = extremePoints π•œ (f '' s) := by
ext b
obtain ⟨a, rfl⟩ := EquivLike.surjective f b
have : βˆ€ x y, f '' openSegment π•œ x y = openSegment π•œ (f x) (f y) :=
image_openSegment _ (LinearMapClass.linearMap f).toAffineMap
simp only [mem_extremePoints, (EquivLike.surjective f).forall,
(EquivLike.injective f).mem_set_image, (EquivLike.injective f).eq_iff, ← this]

end OrderedRing

section LinearOrderedRing

variable [LinearOrderedRing π•œ] [AddCommGroup E] [Module π•œ E]
Expand All @@ -258,7 +269,7 @@ theorem mem_extremePoints_iff_forall_segment : x ∈ A.extremePoints π•œ ↔

theorem Convex.mem_extremePoints_iff_convex_diff (hA : Convex π•œ A) :
x ∈ A.extremePoints π•œ ↔ x ∈ A ∧ Convex π•œ (A \ {x}) := by
use fun hx ↦ ⟨hx.1, (mem_extremePoints_iff_extreme_singleton.1 hx).convex_diff hA⟩
use fun hx ↦ ⟨hx.1, (isExtreme_singleton.2 hx).convex_diff hA⟩
rintro ⟨hxA, hAx⟩
refine' mem_extremePoints_iff_forall_segment.2 ⟨hxA, fun x₁ hx₁ xβ‚‚ hxβ‚‚ hx ↦ _⟩
rw [convex_iff_segment_subset] at hAx
Expand Down
33 changes: 25 additions & 8 deletions Mathlib/Analysis/Convex/KreinMilman.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,6 +5,7 @@ Authors: YaΓ«l Dillies
-/
import Mathlib.Analysis.Convex.Exposed
import Mathlib.Analysis.NormedSpace.HahnBanach.Separation
import Mathlib.Topology.Algebra.ContinuousAffineMap

#align_import analysis.convex.krein_milman from "leanprover-community/mathlib"@"279297937dede7b1b3451b7b0f1786352ad011fa"

Expand Down Expand Up @@ -51,20 +52,19 @@ See chapter 8 of [Barry Simon, *Convexity*][simon2011]
-/


open Set
open scoped Classical

open Classical

variable {E : Type*} [AddCommGroup E] [Module ℝ E] [TopologicalSpace E] [T2Space E]
variable {E F : Type*} [AddCommGroup E] [Module ℝ E] [TopologicalSpace E] [T2Space E]
[TopologicalAddGroup E] [ContinuousSMul ℝ E] [LocallyConvexSpace ℝ E] {s : Set E}
[AddCommGroup F] [Module ℝ F] [TopologicalSpace F] [T1Space F]

/-- **Krein-Milman lemma**: In a LCTVS, any nonempty compact set has an extreme point. -/
theorem IsCompact.has_extreme_point (hscomp : IsCompact s) (hsnemp : s.Nonempty) :
theorem IsCompact.extremePoints_nonempty (hscomp : IsCompact s) (hsnemp : s.Nonempty) :
(s.extremePoints ℝ).Nonempty := by
let S : Set (Set E) := { t | t.Nonempty ∧ IsClosed t ∧ IsExtreme ℝ s t }
rsuffices ⟨t, ⟨⟨x, hxt⟩, htclos, hst⟩, hBmin⟩ : βˆƒ t ∈ S, βˆ€ u ∈ S, u βŠ† t β†’ u = t
· refine' ⟨x, mem_extremePoints_iff_extreme_singleton.2 _⟩
· refine ⟨x, IsExtreme.mem_extremePoints ?_⟩
rwa [← eq_singleton_iff_unique_mem.2 ⟨hxt, fun y hyB => ?_⟩]
by_contra hyx
obtain ⟨l, hl⟩ := geometric_hahn_banach_point_point hyx
Expand All @@ -88,7 +88,7 @@ theorem IsCompact.has_extreme_point (hscomp : IsCompact s) (hsnemp : s.Nonempty)
(hFS t.mem).2.1
obtain htu | hut := hF.total t.mem u.mem
exacts [⟨t, Subset.rfl, htu⟩, ⟨u, hut, Subset.rfl⟩]
#align is_compact.has_extreme_point IsCompact.has_extreme_point
#align is_compact.has_extreme_point IsCompact.extremePoints_nonempty

/-- **Krein-Milman theorem**: In a LCTVS, any compact convex set is the closure of the convex hull
of its extreme points. -/
Expand All @@ -101,7 +101,24 @@ theorem closure_convexHull_extremePoints (hscomp : IsCompact s) (hAconv : Convex
geometric_hahn_banach_closed_point (convex_convexHull _ _).closure isClosed_closure hxt
have h : IsExposed ℝ s ({ y ∈ s | βˆ€ z ∈ s, l z ≀ l y }) := fun _ => ⟨l, rfl⟩
obtain ⟨z, hzA, hz⟩ := hscomp.exists_forall_ge ⟨x, hxA⟩ l.continuous.continuousOn
obtain ⟨y, hy⟩ := (h.isCompact hscomp).has_extreme_point ⟨z, hzA, hz⟩
obtain ⟨y, hy⟩ := (h.isCompact hscomp).extremePoints_nonempty ⟨z, hzA, hz⟩
linarith [hlr _ (subset_closure <| subset_convexHull _ _ <|
h.isExtreme.extremePoints_subset_extremePoints hy), hy.1.2 x hxA]
#align closure_convex_hull_extreme_points closure_convexHull_extremePoints

/-- A continuous affine map is surjective from the extreme points of a compact set to the extreme
points of the image of that set. This inclusion is in general strict. -/
lemma surjOn_extremePoints_image (f : E β†’A[ℝ] F) (hs : IsCompact s) :
SurjOn f (extremePoints ℝ s) (extremePoints ℝ (f '' s)) := by
rintro w hw
-- The fiber of `w` is nonempty and compact
have ht : IsCompact {x ∈ s | f x = w} := hs.inter_right $ isClosed_singleton.preimage f.continuous
have htβ‚€ : {x ∈ s | f x = w}.Nonempty := by simpa using extremePoints_subset hw
-- Hence by the Krein-Milman lemma it has an extreme point `x`
obtain ⟨x, ⟨hx, rfl⟩, hyt⟩ := ht.extremePoints_nonempty htβ‚€
-- `f x = w` and `x` is an extreme point of `s`, so we're done
refine mem_image_of_mem _ ⟨hx, fun y hy z hz hxyz ↦ ?_⟩
have := by simpa using image_openSegment _ f.toAffineMap y z
have := hw.2 (mem_image_of_mem _ hy) (mem_image_of_mem _ hz) $ by
rw [←this]; exact mem_image_of_mem _ hxyz
exact hyt ⟨hy, this.1⟩ ⟨hz, this.2⟩ hxyz
56 changes: 31 additions & 25 deletions Mathlib/Analysis/Seminorm.lean
Original file line number Diff line number Diff line change
Expand Up @@ -919,31 +919,6 @@ theorem closedBall_finset_sup (p : ΞΉ β†’ Seminorm π•œ E) (s : Finset ΞΉ) (x :
exact closedBall_finset_sup_eq_iInter _ _ _ hr
#align seminorm.closed_ball_finset_sup Seminorm.closedBall_finset_sup

theorem ball_smul_ball (p : Seminorm π•œ E) (r₁ rβ‚‚ : ℝ) :
Metric.ball (0 : π•œ) r₁ β€’ p.ball 0 rβ‚‚ βŠ† p.ball 0 (r₁ * rβ‚‚) := by
rw [Set.subset_def]
intro x hx
rw [Set.mem_smul] at hx
rcases hx with ⟨a, y, ha, hy, hx⟩
rw [← hx, mem_ball_zero, map_smul_eq_mul]
gcongr
Β· exact mem_ball_zero_iff.mp ha
Β· exact p.mem_ball_zero.mp hy
#align seminorm.ball_smul_ball Seminorm.ball_smul_ball

theorem closedBall_smul_closedBall (p : Seminorm π•œ E) (r₁ rβ‚‚ : ℝ) :
Metric.closedBall (0 : π•œ) r₁ β€’ p.closedBall 0 rβ‚‚ βŠ† p.closedBall 0 (r₁ * rβ‚‚) := by
rw [Set.subset_def]
intro x hx
rw [Set.mem_smul] at hx
rcases hx with ⟨a, y, ha, hy, hx⟩
rw [← hx, mem_closedBall_zero, map_smul_eq_mul]
rw [mem_closedBall_zero_iff] at ha
gcongr
Β· exact (norm_nonneg a).trans ha
Β· exact p.mem_closedBall_zero.mp hy
#align seminorm.closed_ball_smul_closed_ball Seminorm.closedBall_smul_closedBall

@[simp]
theorem ball_eq_emptyset (p : Seminorm π•œ E) {x : E} {r : ℝ} (hr : r ≀ 0) : p.ball x r = βˆ… := by
ext
Expand All @@ -959,6 +934,37 @@ theorem closedBall_eq_emptyset (p : Seminorm π•œ E) {x : E} {r : ℝ} (hr : r <
exact hr.trans_le (map_nonneg _ _)
#align seminorm.closed_ball_eq_emptyset Seminorm.closedBall_eq_emptyset

theorem closedBall_smul_ball (p : Seminorm π•œ E) {r₁ : ℝ} (hr₁ : r₁ β‰  0) (rβ‚‚ : ℝ) :
Metric.closedBall (0 : π•œ) r₁ β€’ p.ball 0 rβ‚‚ βŠ† p.ball 0 (r₁ * rβ‚‚) := by
simp only [smul_subset_iff, mem_ball_zero, mem_closedBall_zero_iff, map_smul_eq_mul]
refine fun a ha b hb ↦ mul_lt_mul' ha hb (map_nonneg _ _) ?_
exact hr₁.lt_or_lt.resolve_left <| ((norm_nonneg a).trans ha).not_lt

theorem ball_smul_closedBall (p : Seminorm π•œ E) (r₁ : ℝ) {rβ‚‚ : ℝ} (hrβ‚‚ : rβ‚‚ β‰  0) :
Metric.ball (0 : π•œ) r₁ β€’ p.closedBall 0 rβ‚‚ βŠ† p.ball 0 (r₁ * rβ‚‚) := by
simp only [smul_subset_iff, mem_ball_zero, mem_closedBall_zero, mem_ball_zero_iff,
map_smul_eq_mul]
intro a ha b hb
rw [mul_comm, mul_comm r₁]
refine mul_lt_mul' hb ha (norm_nonneg _) (hrβ‚‚.lt_or_lt.resolve_left ?_)
exact ((map_nonneg p b).trans hb).not_lt

theorem ball_smul_ball (p : Seminorm π•œ E) (r₁ rβ‚‚ : ℝ) :
Metric.ball (0 : π•œ) r₁ β€’ p.ball 0 rβ‚‚ βŠ† p.ball 0 (r₁ * rβ‚‚) := by
rcases eq_or_ne rβ‚‚ 0 with rfl | hrβ‚‚
Β· simp
Β· exact (smul_subset_smul_left (ball_subset_closedBall _ _ _)).trans
(ball_smul_closedBall _ _ hrβ‚‚)
#align seminorm.ball_smul_ball Seminorm.ball_smul_ball

theorem closedBall_smul_closedBall (p : Seminorm π•œ E) (r₁ rβ‚‚ : ℝ) :
Metric.closedBall (0 : π•œ) r₁ β€’ p.closedBall 0 rβ‚‚ βŠ† p.closedBall 0 (r₁ * rβ‚‚) := by
simp only [smul_subset_iff, mem_closedBall_zero, mem_closedBall_zero_iff, map_smul_eq_mul]
intro a ha b hb
gcongr
exact (norm_nonneg _).trans ha
#align seminorm.closed_ball_smul_closed_ball Seminorm.closedBall_smul_closedBall

-- Porting note: TODO: make that an `iff`
theorem neg_mem_ball_zero (r : ℝ) (hx : x ∈ ball p 0 r) : -x ∈ ball p 0 r := by
simpa only [mem_ball_zero, map_neg_eq_map] using hx
Expand Down
5 changes: 5 additions & 0 deletions Mathlib/CategoryTheory/Limits/Shapes/Products.lean
Original file line number Diff line number Diff line change
Expand Up @@ -254,6 +254,11 @@ abbrev Sigma.desc {f : Ξ² β†’ C} [HasCoproduct f] {P : C} (p : βˆ€ b, f b ⟢ P)
colimit.desc _ (Cofan.mk P p)
#align category_theory.limits.sigma.desc CategoryTheory.Limits.Sigma.desc

instance {f : Ξ² β†’ C} [HasCoproduct f] : IsIso (Sigma.desc (fun a ↦ Sigma.ΞΉ f a)) := by
convert IsIso.id _
ext
simp

/-- A version of `Cocones.ext`Β for `Cofan`s. -/
@[simps!]
def Cofan.ext {f : Ξ² β†’ C} {c₁ cβ‚‚ : Cofan f} (e : c₁.pt β‰… cβ‚‚.pt)
Expand Down
Loading

0 comments on commit 7cd93d7

Please sign in to comment.