Skip to content

Commit 00202e4

Browse files
committed
feat(Topology/Support): add tsupport_smul_{left,right} (#9778)
- rename `Function.support_smul_subset_right` to `Function.support_const_smul_subset` From sphere-eversion; I'm just upstreaming it. Co-authored-by: grunweg <grunweg@posteo.de>
1 parent fac72bc commit 00202e4

File tree

5 files changed

+18
-8
lines changed

5 files changed

+18
-8
lines changed

Archive/Hairer.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -40,7 +40,7 @@ def SmoothSupportedOn (n : ℕ∞) (s : Set E) : Submodule 𝕜 (E → F) where
4040
zero_mem' :=
4141
⟨(tsupport_eq_empty_iff.mpr rfl).subset.trans (empty_subset _), contDiff_const (c := 0)⟩
4242
smul_mem' r f hf :=
43-
⟨(closure_mono <| support_smul_subset_right r f).trans hf.1, contDiff_const.smul hf.2
43+
⟨(closure_mono <| support_const_smul_subset r f).trans hf.1, contDiff_const.smul hf.2
4444

4545
namespace SmoothSupportedOn
4646

Mathlib/Algebra/Module/Basic.lean

Lines changed: 11 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -765,9 +765,16 @@ theorem Int.smul_one_eq_coe {R : Type*} [Ring R] (m : ℤ) : m • (1 : R) = ↑
765765
namespace Function
766766

767767
lemma support_smul_subset_left [Zero R] [Zero M] [SMulWithZero R M] (f : α → R) (g : α → M) :
768-
support (f • g) ⊆ support f := fun x hfg hf ↦ hfg <| by rw [Pi.smul_apply', hf, zero_smul]
768+
support (f • g) ⊆ support f := fun x hfg hf ↦
769+
hfg <| by rw [Pi.smul_apply', hf, zero_smul]
769770
#align function.support_smul_subset_left Function.support_smul_subset_left
770771

772+
-- Changed (2024-01-21): this lemma was generalised;
773+
-- the old version is now called `support_const_smul_subset`.
774+
lemma support_smul_subset_right [Zero M] [SMulZeroClass R M] (f : α → R) (g : α → M) :
775+
support (f • g) ⊆ support g :=
776+
fun x hbf hf ↦ hbf <| by rw [Pi.smul_apply', hf, smul_zero]
777+
771778
lemma support_const_smul_of_ne_zero [Zero R] [Zero M] [SMulWithZero R M] [NoZeroSMulDivisors R M]
772779
(c : R) (g : α → M) (hc : c ≠ 0) : support (c • g) = support g :=
773780
ext fun x ↦ by simp only [hc, mem_support, Pi.smul_apply, Ne.def, smul_eq_zero, false_or_iff]
@@ -778,10 +785,9 @@ lemma support_smul [Zero R] [Zero M] [SMulWithZero R M] [NoZeroSMulDivisors R M]
778785
ext fun _ => smul_ne_zero_iff
779786
#align function.support_smul Function.support_smul
780787

781-
lemma support_smul_subset_right [Zero M] [SMulZeroClass R M] (a : R) (f : α → M) :
782-
support (a • f) ⊆ support f := fun x hbf hf =>
783-
hbf <| by rw [Pi.smul_apply, hf, smul_zero]
784-
#align function.support_smul_subset_right Function.support_smul_subset_right
788+
lemma support_const_smul_subset [Zero M] [SMulZeroClass R M] (a : R) (f : α → M) :
789+
support (a • f) ⊆ support f := support_smul_subset_right (fun _ ↦ a) f
790+
#align function.support_smul_subset_right Function.support_const_smul_subset
785791

786792
end Function
787793

Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1111,7 +1111,7 @@ protected theorem const_smul {𝕜} [TopologicalSpace 𝕜] [AddMonoid β] [Mono
11111111
FinStronglyMeasurable (c • f) μ := by
11121112
refine' ⟨fun n => c • hf.approx n, fun n => _, fun x => (hf.tendsto_approx x).const_smul c⟩
11131113
rw [SimpleFunc.coe_smul]
1114-
refine' (measure_mono (support_smul_subset_right c _)).trans_lt (hf.fin_support_approx n)
1114+
exact (measure_mono (support_const_smul_subset c _)).trans_lt (hf.fin_support_approx n)
11151115
#align measure_theory.fin_strongly_measurable.const_smul MeasureTheory.FinStronglyMeasurable.const_smul
11161116

11171117
end Arithmetic

Mathlib/RingTheory/HahnSeries.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -490,7 +490,7 @@ variable [PartialOrder Γ] {V : Type*} [Monoid R] [AddMonoid V] [DistribMulActio
490490
instance : SMul R (HahnSeries Γ V) :=
491491
fun r x =>
492492
{ coeff := r • x.coeff
493-
isPWO_support' := x.isPWO_support.mono (Function.support_smul_subset_right r x.coeff) }⟩
493+
isPWO_support' := x.isPWO_support.mono (Function.support_const_smul_subset r x.coeff) }⟩
494494

495495
@[simp]
496496
theorem smul_coeff {r : R} {x : HahnSeries Γ V} {a : Γ} : (r • x).coeff a = r • x.coeff a :=

Mathlib/Topology/Support.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -103,6 +103,10 @@ theorem tsupport_smul_subset_left {M α} [TopologicalSpace X] [Zero M] [Zero α]
103103
closure_mono <| support_smul_subset_left f g
104104
#align tsupport_smul_subset_left tsupport_smul_subset_left
105105

106+
theorem tsupport_smul_subset_right {M α} [TopologicalSpace X] [Zero α] [SMulZeroClass M α]
107+
(f : X → M) (g : X → α) : (tsupport fun x => f x • g x) ⊆ tsupport g :=
108+
closure_mono <| support_smul_subset_right f g
109+
106110
@[to_additive]
107111
theorem mulTSupport_mul [TopologicalSpace X] [Monoid α] {f g : X → α} :
108112
(mulTSupport fun x ↦ f x * g x) ⊆ mulTSupport f ∪ mulTSupport g :=

0 commit comments

Comments
 (0)