Working…

This commit is contained in:
Stefan Kebekus 2024-12-13 12:26:05 +01:00
parent 84abab6b78
commit 39c70be68c

View File

@ -6,6 +6,8 @@ import Nevanlinna.meromorphicOn_divisor
import Nevanlinna.stronglyMeromorphicOn import Nevanlinna.stronglyMeromorphicOn
import Nevanlinna.mathlibAddOn import Nevanlinna.mathlibAddOn
import Mathlib.MeasureTheory.Integral.CircleIntegral import Mathlib.MeasureTheory.Integral.CircleIntegral
import Mathlib.MeasureTheory.Integral.IntervalIntegral
open scoped Interval Topology open scoped Interval Topology
@ -43,27 +45,85 @@ lemma c
exact TopologicalSpace.SecondCountableTopology.to_separableSpace exact TopologicalSpace.SecondCountableTopology.to_separableSpace
theorem integrability_congr_changeDiscrete theorem d
{U S : Set }
{r : }
(hU : Metric.sphere 0 |r| ⊆ U)
(hS : S ∈ Filter.codiscreteWithin U) :
Countable ((circleMap 0 r)⁻¹' (S Uᶜ)ᶜ) := by
have : (circleMap 0 r)⁻¹' U = := by
simpa
simp [this]
sorry
theorem integrability_congr_changeDiscrete₀
{f₁ f₂ : } {f₁ f₂ : }
{U : Set } {U : Set }
{r : } {r : }
(hU : Metric.sphere 0 |r| ⊆ U)
(hf : f₁ =ᶠ[Filter.codiscreteWithin U] f₂) : (hf : f₁ =ᶠ[Filter.codiscreteWithin U] f₂) :
IntervalIntegrable (f₁ ∘ (circleMap 0 r)) MeasureTheory.volume 0 (2 * π) → IntervalIntegrable (f₂ ∘ (circleMap 0 r)) MeasureTheory.volume 0 (2 * π) := by IntervalIntegrable (f₁ ∘ (circleMap 0 r)) MeasureTheory.volume 0 (2 * π) → IntervalIntegrable (f₂ ∘ (circleMap 0 r)) MeasureTheory.volume 0 (2 * π) := by
intro hf₁ intro hf₁
apply IntervalIntegrable.congr hf₁ by_cases hr : r = 0
· unfold circleMap
rw [hr]
simp
have : f₂ ∘ (fun (θ : ) ↦ 0) = (fun r ↦ f₂ 0) := by
exact rfl
rw [this]
simp
· apply IntervalIntegrable.congr hf₁
rw [Filter.eventuallyEq_iff_exists_mem] rw [Filter.eventuallyEq_iff_exists_mem]
use (circleMap 0 r)⁻¹' ({z | f₁ z = f₂ z} ∩ U) use (circleMap 0 r)⁻¹' ({z | f₁ z = f₂ z} Uᶜ)
constructor constructor
· apply Set.Countable.measure_zero · apply Set.Countable.measure_zero
have : (circleMap 0 r ⁻¹' ({z | f₁ z = f₂ z} ∩ U))ᶜ = (circleMap 0 r ⁻¹' ({z | f₁ z = f₂ z} ∩ U)ᶜ) := by have : (circleMap 0 r ⁻¹' ({z | f₁ z = f₂ z} Uᶜ))ᶜ = (circleMap 0 r ⁻¹' ({z | f₁ z = f₂ z} Uᶜ)ᶜ) := by
exact rfl exact rfl
rw [this] rw [this]
apply Set.Countable.preimage_circleMap apply Set.Countable.preimage_circleMap
apply c let A := c {z | f₁ z = f₂ z} U hf
sorry simp at A
sorry simpa
exact hr
· intro x hx · intro x hx
simp at hx simp at hx
simp simp
exact hx.1 rcases hx with h|h
· assumption
· let A := hU (circleMap_mem_sphere' 0 r x)
tauto
theorem integrability_congr_changeDiscrete
{f₁ f₂ : }
{U : Set }
{r : }
(hU : Metric.sphere 0 |r| ⊆ U)
(hf : f₁ =ᶠ[Filter.codiscreteWithin U] f₂) :
IntervalIntegrable (f₁ ∘ (circleMap 0 r)) MeasureTheory.volume 0 (2 * π) ↔ IntervalIntegrable (f₂ ∘ (circleMap 0 r)) MeasureTheory.volume 0 (2 * π) := by
constructor
· exact integrability_congr_changeDiscrete₀ hU hf
· exact integrability_congr_changeDiscrete₀ hU (EventuallyEq.symm hf)
theorem integral_congr_changeDiscrete
{f₁ f₂ : }
{U : Set }
{r : }
(hU : Metric.sphere 0 |r| ⊆ U)
(hf : f₁ =ᶠ[Filter.codiscreteWithin U] f₂) :
∫ (x : ) in (0)..(2 * π), f₁ (circleMap 0 r x) = ∫ (x : ) in (0)..(2 * π), f₂ (circleMap 0 r x) := by
apply intervalIntegral.integral_congr_ae
rw [eventually_iff_exists_mem]
use (circleMap 0 r)⁻¹' ({z | f₁ z = f₂ z} Uᶜ)
sorry