-
Notifications
You must be signed in to change notification settings - Fork 298
feat(analysis/cauchy_equation): Add Cauchy's Functional Equation #12933
Conversation
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
Thank you for working on this! The proof can be golfed a lot using lemmas I mention in my comments.
src/analysis/cauchy_equation.lean
Outdated
|
||
/-- **Cauchy's functional equation**. An additive monoid homomorphism automatically preserves `ℚ`. | ||
-/ | ||
theorem add_monoid_hom.is_linear_map_rat (f : ℝ →+ ℝ) : is_linear_map ℚ f := |
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
See also add_monoid_hom.to_rat_linear_map
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
See #12933 (comment)
src/analysis/cauchy_equation.lean
Outdated
λ h, is_open_univ.measure_ne_zero μ univ_nonempty $ by rw [h, coe_zero, pi.zero_apply] | ||
|
||
lemma exists_zero_nhds_bounded (f : ℝ →+ ℝ) (h : measurable f) : | ||
∃ s, s ∈ 𝓝 (0 : ℝ) ∧ bounded (f '' s) := |
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
What is the right generality for this lemma? Probably, it's true at least for all finite dimensional real normed spaces.
src/analysis/cauchy_equation.lean
Outdated
end | ||
|
||
lemma additive_continuous_at_zero_of_bounded_nhds_zero (f : ℝ →+ ℝ) (hs : s ∈ 𝓝 (0 : ℝ)) | ||
(hbounded : bounded (f '' s)) : continuous_at f 0 := |
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
If you upgrade f
to a ℚ
-linear map, then you can use linear_map.bound_of_shell
or linear_map.bound_of_ball_bound
. Again, you can upgrade your lemma to a real TVS.
src/analysis/cauchy_equation.lean
Outdated
{ rw [smul_eq_mul, smul_eq_mul, h (c * x), h x, ←mul_assoc, mul_comm _ c, mul_assoc] } | ||
end | ||
|
||
lemma is_linear_rat (f : ℝ →+ ℝ) : ∀ (q : ℚ), f q = f 1 * q := |
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
This lemma and most lemmas below are already in mathlib
, aren't they?
@YaelDillies, are you interested in finishing this, or are your own PRs more important to you? |
Yes, I in fact started fixing it earlier today. It needs some more cleanup but the gist is here. |
|
||
local notation `ℝⁿ` := ι → ℝ | ||
|
||
lemma add_monoid_hom.measurable_of_continuous (f : ℝ →+ ℝ) (h : measurable f) : continuous f := |
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
Name says measurable_of_continuous
but the theorem says measurable f -> continuous f
.
@@ -2329,6 +2329,8 @@ by rw [← empty_mem_iff_bot, mem_ae_iff, compl_empty, measure_univ_eq_zero] | |||
@[simp] lemma ae_ne_bot : μ.ae.ne_bot ↔ μ ≠ 0 := | |||
ne_bot_iff.trans (not_congr ae_eq_bot) | |||
|
|||
instance ae_ne_bot.to_ne_zero [μ.ae.ne_bot] : ne_zero μ := ⟨ae_ne_bot.1 ‹_›⟩ |
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
We have [NeZero μ] : μ.ae.NeBot
in Lean 4.
open topological_space | ||
variables {G H : Type*} [seminormed_add_group G] [topological_add_group G] [is_R_or_C H] {s : set G} | ||
|
||
lemma add_monoid_hom.continuous_of_bounded_nhds_zero (f : G →+ H) (hs : s ∈ 𝓝 (0 : G)) |
There was a problem hiding this comment.
Choose a reason for hiding this comment
The reason will be displayed to describe this comment to others. Learn more.
Isn't [is_R_or_C H]
too strong? Is it true for any normed space over rationals?
Ported to LeanCamCombi |
Cauchy's functional equation over the real-line for a Lebesgue measurable function and some of the usual variants from which linearity of a given function is deduced.