Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 72789f5

Browse files
committed
feat(linear_algebra/affine_space/affine_subspace): add lemma affine_equiv.span_eq_top_iff (#9695)
Together with supporting lemmas. Formalized as part of the Sphere Eversion project.
1 parent cef78dd commit 72789f5

File tree

3 files changed

+93
-1
lines changed

3 files changed

+93
-1
lines changed

src/algebra/add_torsor.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -193,6 +193,10 @@ lemma vsub_mem_vsub {ps pt : P} (hs : ps ∈ s) (ht : pt ∈ t) :
193193
(ps -ᵥ pt) ∈ s -ᵥ t :=
194194
mem_image2_of_mem hs ht
195195

196+
@[simp] lemma mem_vsub {s t : set P} (g : G) :
197+
g ∈ s -ᵥ t ↔ ∃ (x y : P), x ∈ s ∧ y ∈ t ∧ x -ᵥ y = g :=
198+
mem_image2
199+
196200
/-- `s -ᵥ t` is monotone in both arguments. -/
197201
@[mono] lemma vsub_subset_vsub {s' t' : set P} (hs : s ⊆ s') (ht : t ⊆ t') :
198202
s -ᵥ t ⊆ s' -ᵥ t' :=

src/linear_algebra/affine_space/affine_map.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -308,6 +308,19 @@ begin
308308
rw [h, equiv.comp_surjective, equiv.surjective_comp],
309309
end
310310

311+
lemma image_vsub_image {s t : set P1} (f : P1 →ᵃ[k] P2) :
312+
(f '' s) -ᵥ (f '' t) = f.linear '' (s -ᵥ t) :=
313+
begin
314+
ext v,
315+
simp only [set.mem_vsub, set.mem_image, exists_exists_and_eq_and, exists_and_distrib_left,
316+
← f.linear_map_vsub],
317+
split,
318+
{ rintros ⟨x, hx, y, hy, hv⟩,
319+
exact ⟨x -ᵥ y, ⟨x, hx, y, hy, rfl⟩, hv⟩, },
320+
{ rintros ⟨-, ⟨x, hx, y, hy, rfl⟩, rfl⟩,
321+
exact ⟨x, hx, y, hy, rfl⟩, },
322+
end
323+
311324
omit V2
312325

313326
/-! ### Definition of `affine_map.line_map` and lemmas about it -/

src/linear_algebra/affine_space/affine_subspace.lean

Lines changed: 76 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,7 @@ Copyright (c) 2020 Joseph Myers. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Joseph Myers
55
-/
6-
import linear_algebra.affine_space.basic
6+
import linear_algebra.affine_space.affine_equiv
77
import linear_algebra.tensor_product
88
import data.set.intervals.unordered_interval
99

@@ -178,6 +178,7 @@ variables (k : Type*) {V : Type*} (P : Type*) [ring k] [add_comm_group V] [modul
178178
[affine_space V P]
179179
include V
180180

181+
-- TODO Refactor to use `instance : set_like (affine_subspace k P) P :=` instead
181182
instance : has_coe (affine_subspace k P) (set P) := ⟨carrier⟩
182183
instance : has_mem P (affine_subspace k P) := ⟨λ p s, p ∈ (s : set P)⟩
183184

@@ -1175,3 +1176,77 @@ begin
11751176
end
11761177

11771178
end affine_subspace
1179+
1180+
section maps
1181+
1182+
variables {k V₁ P₁ V₂ P₂ : Type*} [ring k]
1183+
variables [add_comm_group V₁] [module k V₁] [add_torsor V₁ P₁]
1184+
variables [add_comm_group V₂] [module k V₂] [add_torsor V₂ P₂]
1185+
include V₁ V₂
1186+
1187+
variables (f : P₁ →ᵃ[k] P₂)
1188+
1189+
@[simp] lemma affine_map.vector_span_image_eq_submodule_map {s : set P₁} :
1190+
submodule.map f.linear (vector_span k s) = vector_span k (f '' s) :=
1191+
by simp [f.image_vsub_image, vector_span_def]
1192+
1193+
namespace affine_subspace
1194+
1195+
/-- The image of an affine subspace under an affine map as an affine subspace. -/
1196+
def map (s : affine_subspace k P₁) : affine_subspace k P₂ :=
1197+
{ carrier := f '' s,
1198+
smul_vsub_vadd_mem :=
1199+
begin
1200+
rintros t - - - ⟨p₁, h₁, rfl⟩ ⟨p₂, h₂, rfl⟩ ⟨p₃, h₃, rfl⟩,
1201+
use t • (p₁ -ᵥ p₂) +ᵥ p₃,
1202+
suffices : t • (p₁ -ᵥ p₂) +ᵥ p₃ ∈ s, { by simp [this], },
1203+
exact s.smul_vsub_vadd_mem t h₁ h₂ h₃,
1204+
end }
1205+
1206+
@[simp] lemma map_coe (s : affine_subspace k P₁) : (s.map f : set P₂) = f '' s := rfl
1207+
1208+
@[simp] lemma map_bot : (⊥ : affine_subspace k P₁).map f = ⊥ :=
1209+
by { rw ← ext_iff, exact image_empty f, }
1210+
1211+
@[simp] lemma map_direction (s : affine_subspace k P₁) :
1212+
(s.map f).direction = s.direction.map f.linear :=
1213+
by simp [direction_eq_vector_span]
1214+
1215+
lemma map_span (s : set P₁) :
1216+
(affine_span k s).map f = affine_span k (f '' s) :=
1217+
begin
1218+
rcases s.eq_empty_or_nonempty with rfl | ⟨p, hp⟩, { simp, },
1219+
apply ext_of_direction_eq,
1220+
{ simp [direction_affine_span], },
1221+
{ exact ⟨f p, mem_image_of_mem f (subset_affine_span k _ hp),
1222+
subset_affine_span k _ (mem_image_of_mem f hp)⟩, },
1223+
end
1224+
1225+
end affine_subspace
1226+
1227+
namespace affine_map
1228+
1229+
@[simp] lemma map_top_of_surjective (hf : function.surjective f) : affine_subspace.map f ⊤ = ⊤ :=
1230+
begin
1231+
rw ← affine_subspace.ext_iff,
1232+
exact image_univ_of_surjective hf,
1233+
end
1234+
1235+
lemma span_eq_top_of_surjective {s : set P₁}
1236+
(hf : function.surjective f) (h : affine_span k s = ⊤) :
1237+
affine_span k (f '' s) = ⊤ :=
1238+
by rw [← affine_subspace.map_span, h, map_top_of_surjective f hf]
1239+
1240+
end affine_map
1241+
1242+
lemma affine_equiv.span_eq_top_iff {s : set P₁} (e : P₁ ≃ᵃ[k] P₂) :
1243+
affine_span k s = ⊤ ↔ affine_span k (e '' s) = ⊤ :=
1244+
begin
1245+
refine ⟨(e : P₁ →ᵃ[k] P₂).span_eq_top_of_surjective e.surjective, _⟩,
1246+
intros h,
1247+
have : s = e.symm '' (e '' s), { simp [← image_comp], },
1248+
rw this,
1249+
exact (e.symm : P₂ →ᵃ[k] P₁).span_eq_top_of_surjective e.symm.surjective h,
1250+
end
1251+
1252+
end maps

0 commit comments

Comments
 (0)