diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index 548d9f4b345bc6..4ffb4d9295ffc0 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -119,24 +119,28 @@ theorem AnalyticOn.comp₂ {h : F × G → H} {f : E → F} {g : E → G} {s : S fun _ xt ↦ (ha _ (m _ xt)).comp₂ (fa _ xt) (ga _ xt) /-- Analytic functions on products are analytic in the first coordinate -/ -theorem AnalyticAt.along_fst {f : E × F → G} {p : E × F} (fa : AnalyticAt 𝕜 f p) : +theorem AnalyticAt.curry_left {f : E × F → G} {p : E × F} (fa : AnalyticAt 𝕜 f p) : AnalyticAt 𝕜 (fun x ↦ f (x, p.2)) p.1 := AnalyticAt.comp₂ fa (analyticAt_id _ _) analyticAt_const +alias AnalyticAt.along_fst := AnalyticAt.curry_left /-- Analytic functions on products are analytic in the second coordinate -/ -theorem AnalyticAt.along_snd {f : E × F → G} {p : E × F} (fa : AnalyticAt 𝕜 f p) : +theorem AnalyticAt.curry_right {f : E × F → G} {p : E × F} (fa : AnalyticAt 𝕜 f p) : AnalyticAt 𝕜 (fun y ↦ f (p.1, y)) p.2 := AnalyticAt.comp₂ fa analyticAt_const (analyticAt_id _ _) +alias AnalyticAt.along_snd := AnalyticAt.curry_right /-- Analytic functions on products are analytic in the first coordinate -/ -theorem AnalyticOn.along_fst {f : E × F → G} {s : Set (E × F)} {y : F} (fa : AnalyticOn 𝕜 f s) : +theorem AnalyticOn.curry_left {f : E × F → G} {s : Set (E × F)} {y : F} (fa : AnalyticOn 𝕜 f s) : AnalyticOn 𝕜 (fun x ↦ f (x, y)) {x | (x, y) ∈ s} := fun x m ↦ (fa (x, y) m).along_fst +alias AnalyticOn.along_fst := AnalyticOn.curry_left /-- Analytic functions on products are analytic in the second coordinate -/ -theorem AnalyticOn.along_snd {f : E × F → G} {x : E} {s : Set (E × F)} (fa : AnalyticOn 𝕜 f s) : +theorem AnalyticOn.curry_right {f : E × F → G} {x : E} {s : Set (E × F)} (fa : AnalyticOn 𝕜 f s) : AnalyticOn 𝕜 (fun y ↦ f (x, y)) {y | (x, y) ∈ s} := fun y m ↦ (fa (x, y) m).along_snd +alias AnalyticOn.along_snd := AnalyticOn.curry_right /-! ### Arithmetic on analytic functions diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index 1736af41690a8a..fec455c896e4ba 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -853,7 +853,7 @@ lemma isOpen_A_with_param {r s : ℝ} (hf : Continuous f.uncurry) (L : E →L[ simp only [A, half_lt_self_iff, not_lt, mem_Ioc, mem_ball, map_sub, mem_setOf_eq] apply isOpen_iff_mem_nhds.2 rintro ⟨a, x⟩ ⟨r', ⟨Irr', Ir'r⟩, hr⟩ - have ha : Continuous (f a) := continuous_uncurry_left a hf + have ha : Continuous (f a) := hf.uncurry_left a rcases exists_between Irr' with ⟨t, hrt, htr'⟩ rcases exists_between hrt with ⟨t', hrt', ht't⟩ obtain ⟨b, b_lt, hb⟩ : ∃ b, b < s * r ∧ ∀ y ∈ closedBall x t, ∀ z ∈ closedBall x t, diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 329a2bb37c3b36..23fa227ec49b13 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -508,18 +508,22 @@ lemma isClosedMap_swap : IsClosedMap (Prod.swap : X × Y → Y × X) := fun s hs rw [image_swap_eq_preimage_swap] exact hs.preimage continuous_swap -theorem continuous_uncurry_left {f : X → Y → Z} (x : X) (h : Continuous (uncurry f)) : +theorem Continuous.uncurry_left {f : X → Y → Z} (x : X) (h : Continuous (uncurry f)) : Continuous (f x) := h.comp (Continuous.Prod.mk _) -#align continuous_uncurry_left continuous_uncurry_left +#align continuous_uncurry_left Continuous.uncurry_left -theorem continuous_uncurry_right {f : X → Y → Z} (y : Y) (h : Continuous (uncurry f)) : +theorem Continuous.uncurry_right {f : X → Y → Z} (y : Y) (h : Continuous (uncurry f)) : Continuous fun a => f a y := h.comp (Continuous.Prod.mk_left _) -#align continuous_uncurry_right continuous_uncurry_right +#align continuous_uncurry_right Continuous.uncurry_right + +-- 2024-03-09 +@[deprecated] alias continuous_uncurry_left := Continuous.uncurry_left +@[deprecated] alias continuous_uncurry_right := Continuous.uncurry_right theorem continuous_curry {g : X × Y → Z} (x : X) (h : Continuous g) : Continuous (curry g x) := - continuous_uncurry_left x h + Continuous.uncurry_left x h #align continuous_curry continuous_curry theorem IsOpen.prod {s : Set X} {t : Set Y} (hs : IsOpen s) (ht : IsOpen t) : IsOpen (s ×ˢ t) := @@ -636,6 +640,29 @@ theorem ContinuousAt.prod_map' {f : X → Z} {g : Y → W} {x : X} {y : Y} (hf : hf.fst'.prod hg.snd' #align continuous_at.prod_map' ContinuousAt.prod_map' +theorem ContinuousAt.comp₂ {f : Y × Z → W} {g : X → Y} {h : X → Z} {x : X} + (hf : ContinuousAt f (g x, h x)) (hg : ContinuousAt g x) (hh : ContinuousAt h x) : + ContinuousAt (fun x ↦ f (g x, h x)) x := + ContinuousAt.comp hf (hg.prod hh) + +theorem ContinuousAt.comp₂_of_eq {f : Y × Z → W} {g : X → Y} {h : X → Z} {x : X} {y : Y × Z} + (hf : ContinuousAt f y) (hg : ContinuousAt g x) (hh : ContinuousAt h x) (e : (g x, h x) = y) : + ContinuousAt (fun x ↦ f (g x, h x)) x := by + rw [← e] at hf + exact hf.comp₂ hg hh + +/-- Continuous functions on products are continuous in their first argument -/ +theorem Continuous.curry_left {f : X × Y → Z} (hf : Continuous f) {y : Y} : + Continuous fun x ↦ f (x, y) := + hf.comp (continuous_id.prod_mk continuous_const) +alias Continuous.along_fst := Continuous.curry_left + +/-- Continuous functions on products are continuous in their second argument -/ +theorem Continuous.curry_right {f : X × Y → Z} (hf : Continuous f) {x : X} : + Continuous fun y ↦ f (x, y) := + hf.comp (continuous_const.prod_mk continuous_id) +alias Continuous.along_snd := Continuous.curry_right + -- todo: reformulate using `Set.image2` -- todo: prove a version of `generateFrom_union` with `image2 (∩) s t` in the LHS and use it here theorem prod_generateFrom_generateFrom_eq {X Y : Type*} {s : Set (Set X)} {t : Set (Set Y)} diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 87536ea221bea1..b43122e8e89316 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -1152,6 +1152,19 @@ theorem ContinuousOn.prod {f : α → β} {g : α → γ} {s : Set α} (hf : Con ContinuousWithinAt.prod (hf x hx) (hg x hx) #align continuous_on.prod ContinuousOn.prod +theorem ContinuousAt.comp₂_continuousWithinAt {f : β × γ → δ} {g : α → β} {h : α → γ} {x : α} + {s : Set α} (hf : ContinuousAt f (g x, h x)) (hg : ContinuousWithinAt g s x) + (hh : ContinuousWithinAt h s x) : + ContinuousWithinAt (fun x ↦ f (g x, h x)) s x := + ContinuousAt.comp_continuousWithinAt hf (hg.prod hh) + +theorem ContinuousAt.comp₂_continuousWithinAt_of_eq {f : β × γ → δ} {g : α → β} + {h : α → γ} {x : α} {s : Set α} {y : β × γ} (hf : ContinuousAt f y) + (hg : ContinuousWithinAt g s x) (hh : ContinuousWithinAt h s x) (e : (g x, h x) = y) : + ContinuousWithinAt (fun x ↦ f (g x, h x)) s x := by + rw [← e] at hf + exact hf.comp₂_continuousWithinAt hg hh + theorem Inducing.continuousWithinAt_iff {f : α → β} {g : β → γ} (hg : Inducing g) {s : Set α} {x : α} : ContinuousWithinAt f s x ↔ ContinuousWithinAt (g ∘ f) s x := by simp_rw [ContinuousWithinAt, Inducing.tendsto_nhds_iff hg]; rfl diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 2b6a730af14d50..b2df6df367ae95 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -1605,8 +1605,7 @@ theorem eqOn_closure₂' [T2Space Z] {s : Set X} {t : Set Y} {f g : X → Y → theorem eqOn_closure₂ [T2Space Z] {s : Set X} {t : Set Y} {f g : X → Y → Z} (h : ∀ x ∈ s, ∀ y ∈ t, f x y = g x y) (hf : Continuous (uncurry f)) (hg : Continuous (uncurry g)) : ∀ x ∈ closure s, ∀ y ∈ closure t, f x y = g x y := - eqOn_closure₂' h (fun x => continuous_uncurry_left x hf) (fun x => continuous_uncurry_right x hf) - (fun y => continuous_uncurry_left y hg) fun y => continuous_uncurry_right y hg + eqOn_closure₂' h hf.uncurry_left hf.uncurry_right hg.uncurry_left hg.uncurry_right #align eq_on_closure₂ eqOn_closure₂ /-- If `f x = g x` for all `x ∈ s` and `f`, `g` are continuous on `t`, `s ⊆ t ⊆ closure s`, then