Documentation

Mathlib.Topology.Algebra.Affine

Topological properties of affine spaces and maps #

This file contains a few facts regarding the continuity of affine maps.

If f is an affine map, then its linear part is continuous iff f is continuous.

@[deprecated AffineMap.continuous_linear_iff (since := "2025-09-13")]

An affine map is continuous iff its underlying linear map is continuous. See also AffineMap.continuous_linear_iff.

If f is an affine map, then its linear part is an open map iff f is an open map.

The line map is continuous.

theorem Filter.Tendsto.lineMap {R : Type u_1} {V : Type u_2} {P : Type u_3} [AddCommGroup V] [TopologicalSpace V] [AddTorsor V P] [TopologicalSpace P] [IsTopologicalAddTorsor P] [Ring R] [Module R V] [TopologicalSpace R] [ContinuousSMul R V] {α : Type u_6} {l : Filter α} {f₁ f₂ : αP} {g : αR} {p₁ p₂ : P} {c : R} (h₁ : Tendsto f₁ l (nhds p₁)) (h₂ : Tendsto f₂ l (nhds p₂)) (hg : Tendsto g l (nhds c)) :
Tendsto (fun (x : α) => (AffineMap.lineMap (f₁ x) (f₂ x)) (g x)) l (nhds ((AffineMap.lineMap p₁ p₂) c))
theorem Filter.Tendsto.midpoint {R : Type u_1} {V : Type u_2} {P : Type u_3} [AddCommGroup V] [TopologicalSpace V] [AddTorsor V P] [TopologicalSpace P] [IsTopologicalAddTorsor P] [Ring R] [Module R V] [TopologicalSpace R] [ContinuousSMul R V] {α : Type u_6} {l : Filter α} [Invertible 2] {f₁ f₂ : αP} {p₁ p₂ : P} (h₁ : Tendsto f₁ l (nhds p₁)) (h₂ : Tendsto f₂ l (nhds p₂)) :
Tendsto (fun (x : α) => _root_.midpoint R (f₁ x) (f₂ x)) l (nhds (_root_.midpoint R p₁ p₂))
theorem eventually_homothety_mem_of_mem_interior (R : Type u_1) {W : Type u_4} {Q : Type u_5} [AddCommGroup W] [TopologicalSpace W] [AddTorsor W Q] [TopologicalSpace Q] [IsTopologicalAddTorsor Q] [CommRing R] [TopologicalSpace R] [Module R W] [ContinuousSMul R W] (x : Q) {s : Set Q} {y : Q} (hy : y interior s) :
∀ᶠ (δ : R) in nhds 1, (AffineMap.homothety x δ) y s
theorem AffineMap.homothety_isOpenMap {R : Type u_1} {V : Type u_2} {P : Type u_3} [AddCommGroup V] [TopologicalSpace V] [AddTorsor V P] [TopologicalSpace P] [IsTopologicalAddTorsor P] [Field R] [Module R V] [ContinuousConstSMul R V] (x : P) (t : R) (ht : t 0) :