diff --git a/Nevanlinna/stronglyMeromorphic.lean b/Nevanlinna/stronglyMeromorphic.lean index 8d0f6b7..4140e50 100644 --- a/Nevanlinna/stronglyMeromorphic.lean +++ b/Nevanlinna/stronglyMeromorphic.lean @@ -81,41 +81,49 @@ theorem AnalyticAt.stronglyMeromorphicAt tauto -theorem MeromorphicAt.order_neq_top_iff - {f : ℂ → ℂ} - {z₀ : ℂ} - (hf : MeromorphicAt f z₀) : - hf.order ≠ ⊤ ↔ ∃ (g : ℂ → ℂ), AnalyticAt ℂ g z₀ ∧ g z₀ ≠ 0 ∧ ∀ᶠ (z : ℂ) in nhdsWithin z₀ {z₀}ᶜ, f z = (z - z₀) ^ (hf.order.untop' 0) • g z := by - - rw [← hf.order_eq_int_iff (hf.order.untop' 0)] - constructor - · intro h₁f - apply? - exact Eq.symm (ENat.coe_toNat h₁f) - · intro h₁f - exact ENat.coe_toNat_eq_self.mp (id (Eq.symm h₁f)) - - /- Make strongly MeromorphicAt -/ noncomputable def MeromorphicAt.makeStronglyMeromorphicAt {f : ℂ → ℂ} {z₀ : ℂ} (hf : MeromorphicAt f z₀) : ℂ → ℂ := by - - by_cases h₁f : hf.order = ⊤ - · exact 0 - · - intro z by_cases z = z₀ - · by_cases h₂f : hf.order = 0 - · have : (0 : WithTop ℤ) = (0 : ℤ) := rfl - rw [this, hf.order_eq_int_iff] at h₂f - exact (Classical.choose h₂f) z₀ + · by_cases h₁f : hf.order = (0 : ℤ) + · rw [hf.order_eq_int_iff] at h₁f + exact (Classical.choose h₁f) z₀ · exact 0 · exact f z +lemma m₁ + {f : ℂ → ℂ} + {z₀ : ℂ} + (hf : MeromorphicAt f z₀) : + ∀ z ≠ z₀, f z = hf.makeStronglyMeromorphicAt z := by + intro z hz + unfold MeromorphicAt.makeStronglyMeromorphicAt + simp [hz] + +lemma m₂ + {f : ℂ → ℂ} + {z₀ : ℂ} + (hf : MeromorphicAt f z₀) : + ∀ᶠ (z : ℂ) in nhdsWithin z₀ {z₀}ᶜ, f z = hf.makeStronglyMeromorphicAt z := by + apply eventually_nhdsWithin_of_forall + exact fun x a => m₁ hf x a + + +open Topology + +lemma Mnhds + {f g : ℂ → ℂ} + {z₀ : ℂ} + (h₁ : ∀ᶠ (z : ℂ) in nhdsWithin z₀ {z₀}ᶜ, f z = g z) + (h₂ : f z₀ = g z₀) : + ∀ᶠ (z : ℂ) in nhds z₀, f z = g z := by + rw [eventually_nhds_iff] + rw [eventually_nhdsWithin_iff] at h₁ + sorry theorem StronglyMeromorphicAt_of_makeStronglyMeromorphic {f : ℂ → ℂ} @@ -123,6 +131,26 @@ theorem StronglyMeromorphicAt_of_makeStronglyMeromorphic (hf : MeromorphicAt f z₀) : StronglyMeromorphicAt hf.makeStronglyMeromorphicAt z₀ := by + by_cases h₂f : hf.order = ⊤ + · rw [MeromorphicAt.order_eq_top_iff] at h₂f + let Z : ℂ → ℂ := fun z ↦ 0 + have : hf.makeStronglyMeromorphicAt =ᶠ[𝓝 z₀] 0 := by + --unfold Filter.EventuallyEq + apply Mnhds + · apply eventually_nhdsWithin_of_forall + intro x hx + + sorry + + + sorry + + + + apply AnalyticAt.stronglyMeromorphicAt + rw [analyticAt_congr this] + apply analyticAt_const + by_cases h₂f : hf.order = 0 diff --git a/lake-manifest.json b/lake-manifest.json index 66bfdc5..22c26ed 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "9efd9c267ad7a71c5e3a83e8fbbd446fe61ef119", + "rev": "7c5548eeeb1748da5c2872dbd866d3acbaea4b3f", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -15,7 +15,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "fa3d73a2cf077f4b14c7840352ac7b08aeb6eb41", + "rev": "1357f4f49450abb9dfd4783e38219f4ce84f9785", "name": "Qq", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -35,10 +35,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "cd20dae87c48495f0220663014dff11671597fcf", + "rev": "baa65c6339a56bd22b7292aa4511c54b3cc7a6af", "name": "proofwidgets", "manifestFile": "lake-manifest.json", - "inputRev": "v0.0.43-pre", + "inputRev": "v0.0.43", "inherited": true, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover/lean4-cli", @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "4b61d4abc1659f15ffda5ec24fdebc229d51d066", + "rev": "7bedaed1ef024add1e171cc17706b012a9a37802", "name": "LeanSearchClient", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -75,7 +75,7 @@ "type": "git", "subDir": null, "scope": "", - "rev": "5ab87a979ed4fbe6219ecaf7ca251345ab04ec8f", + "rev": "818db408b33f801f4094614250e6b3cfb9e065c8", "name": "mathlib", "manifestFile": "lake-manifest.json", "inputRev": null,