Skip to content

Commit decdbad

Browse files
committed
Tweaks
1 parent d05f38f commit decdbad

File tree

2 files changed

+7
-15
lines changed

2 files changed

+7
-15
lines changed

RMT4.lean

Lines changed: 2 additions & 15 deletions
Original file line numberDiff line numberDiff line change
@@ -43,11 +43,7 @@ lemma ContinuousOn_uderiv (hU : IsOpen U) : ContinuousOn uderiv (𝓗 U) := by
4343

4444
def 𝓜 (U : Set ℂ) := {f ∈ 𝓗 U | MapsTo f U (closedBall (0 : ℂ) 1)}
4545

46-
lemma UniformlyBounded_𝓜 : UniformlyBoundedOn ((↑) : 𝓜 U → ℂ →ᵤ[compacts U] ℂ) U := by
47-
rintro K ⟨hK1, _⟩
48-
refine ⟨1, zero_lt_one, ?_⟩
49-
rintro z hz x ⟨⟨f, hf⟩, rfl⟩
50-
exact hf.2 (hK1 hz)
46+
example : 𝓜 U = 𝓕K U (fun _ => closedBall 0 1) := 𝓕K_const.symm
5147

5248
lemma IsClosed_𝓜 (hU : IsOpen U) : IsClosed (𝓜 U) := by
5349
suffices : IsClosed {f : ℂ →ᵤ[compacts U] ℂ | MapsTo f U (closedBall 0 1)}
@@ -58,16 +54,7 @@ lemma IsClosed_𝓜 (hU : IsOpen U) : IsClosed (𝓜 U) := by
5854
(mem_singleton z) ⟨singleton_subset_iff.2 hz, isCompact_singleton⟩).continuous)
5955

6056
lemma IsCompact_𝓜 (hU : IsOpen U) : IsCompact (𝓜 U) := by
61-
have l1 (K) (hK : K ∈ compacts U) : EquicontinuousOn ((↑) : 𝓜 U → ℂ →ᵤ[compacts U] ℂ) K :=
62-
UniformlyBounded_𝓜.equicontinuous_on hU (·.2.1) hK
63-
have l2 : ∀ K ∈ compacts U, ∀ x ∈ K, ∃ Q, IsCompact Q ∧ ∀ (f : 𝓜 U), f.val x ∈ Q := by
64-
intro K hK x hx
65-
refine ⟨closedBall 0 1, isCompact_of_isClosed_isBounded isClosed_ball isBounded_closedBall, ?_⟩
66-
exact fun f => f.prop.2 (hK.1 hx)
67-
rw [isCompact_iff_compactSpace]
68-
refine ArzelaAscoli.compactSpace_of_closedEmbedding (fun _ hK => hK.2) ?_ l1 l2
69-
refine ⟨⟨by tauto, fun f g => Subtype.ext⟩, ?_⟩
70-
simpa [UniformOnFun.ofFun, range] using IsClosed_𝓜 hU
57+
simpa only [𝓕K_const] using isCompact_𝓕K hU (fun _ _ => isCompact_closedBall 0 1)
7158

7259
-- `𝓘 U` : holomorphic injections into the unit ball
7360

RMT4/montel.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -107,6 +107,11 @@ lemma UniformlyBoundedOn.equicontinuous_on
107107
def 𝓕K (U : Set ℂ) (Q : Set ℂ → Set ℂ) : Set (ℂ →ᵤ[compacts U] ℂ) :=
108108
{f ∈ 𝓗 U | ∀ K ∈ compacts U, MapsTo f K (Q K)}
109109

110+
lemma 𝓕K_const {Q : Set ℂ} : 𝓕K U (fun _ => Q) = {f ∈ 𝓗 U | MapsTo f U Q} := by
111+
ext f ; simp [𝓕K, 𝓗] ; rintro - ; constructor <;> intro h
112+
· exact fun z hz => h {z} ⟨by { rintro w rfl ; exact hz }, isCompact_singleton⟩ (mem_singleton z)
113+
· exact fun K ⟨h1, _⟩ => h.mono_left h1
114+
110115
theorem isClosed_𝓕K (hU : IsOpen U) (hQ : ∀ K ∈ compacts U, IsCompact (Q K)) :
111116
IsClosed (𝓕K U Q) := by
112117
rw [𝓕K, setOf_and] ; apply (isClosed_𝓗 hU).inter

0 commit comments

Comments
 (0)