Skip to content

Commit

Permalink
Rename Immersion.differentiable to contMDiff.
Browse files Browse the repository at this point in the history
  • Loading branch information
grunweg committed Mar 18, 2024
1 parent bfd0afb commit 912bbcf
Show file tree
Hide file tree
Showing 2 changed files with 4 additions and 4 deletions.
6 changes: 3 additions & 3 deletions SphereEversion/Global/Immersion.lean
Original file line number Diff line number Diff line change
Expand Up @@ -122,13 +122,13 @@ variable {n : ℕ} (E : Type*) [NormedAddCommGroup E] [InnerProductSpace ℝ E]
/-- The inclusion of `𝕊^n` into `ℝ^{n+1}` is an immersion. -/
theorem immersion_inclusion_sphere : Immersion (𝓡 n) 𝓘(ℝ, E)
(fun x : sphere (0 : E) 1 ↦ (x : E)) ⊤ where
differentiable := contMDiff_coe_sphere
contMDiff := contMDiff_coe_sphere
diff_injective := mfderiv_coe_sphere_injective

/-- The antipodal map on `𝕊^n ⊆ ℝ^{n+1}` is an immersion. -/
theorem immersion_antipodal_sphere : Immersion (𝓡 n) 𝓘(ℝ, E)
(fun x : sphere (0 : E) 1 ↦ -(x : E)) ⊤ where
differentiable :=
contMDiff :=
-- Write this as the composition of `coe_sphere` and the antipodal map on `E`.
-- The other direction elaborates much worse.
(contDiff_neg.contMDiff).comp contMDiff_coe_sphere
Expand Down Expand Up @@ -334,7 +334,7 @@ theorem sphere_eversion :
rw [this (1, x) (by simp)]
convert formalEversion_one E ω x
· exact fun t ↦ {
differentiable := Smooth.uncurry_left 𝓘(ℝ, ℝ) (𝓡 2) 𝓘(ℝ, E) h₁ t
contMDiff := Smooth.uncurry_left 𝓘(ℝ, ℝ) (𝓡 2) 𝓘(ℝ, E) h₁ t
diff_injective := h₅ t
}

Expand Down
2 changes: 1 addition & 1 deletion SphereEversion/ToMathlib/Geometry/Manifold/Immersion.lean
Original file line number Diff line number Diff line change
Expand Up @@ -53,7 +53,7 @@ section Definition

/-- A `C^n` immersion `f : M → M` is a `C^n` map whose differential is injective at every point. -/
structure Immersion (f : M → M') (n : ℕ∞) : Prop :=
differentiable : ContMDiff I I' n f
contMDiff : ContMDiff I I' n f
diff_injective : ∀ p, Injective (mfderiv I I' f p)

/-- An injective `C^n` immersion -/
Expand Down

0 comments on commit 912bbcf

Please sign in to comment.