Skip to content

Commit 01a486f

Browse files
committed
chore(LinearAlgebra/AffineSpace/AffineSubspace): trivial affine morphisms (#8434)
We already have the `LinearEquiv` versions of these. This also copies the lemmas around the definition
1 parent cf21ac0 commit 01a486f

File tree

1 file changed

+61
-0
lines changed

1 file changed

+61
-0
lines changed

Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean

Lines changed: 61 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -828,6 +828,17 @@ theorem card_pos_of_affineSpan_eq_top {ι : Type*} [Fintype ι] {p : ι → P}
828828
exact Fintype.card_pos_iff.mpr ⟨i⟩
829829
#align affine_subspace.card_pos_of_affine_span_eq_top AffineSubspace.card_pos_of_affineSpan_eq_top
830830

831+
attribute [local instance] toAddTorsor
832+
833+
/-- The top affine subspace is linearly equivalent to the affine space.
834+
835+
This is the affine version of `Submodule.topEquiv`. -/
836+
@[simps! linear apply symm_apply_coe]
837+
def topEquiv : (⊤ : AffineSubspace k P) ≃ᵃ[k] P where
838+
toEquiv := Equiv.Set.univ P
839+
linear := .ofEq _ _ (direction_top _ _ _) ≪≫ₗ Submodule.topEquiv
840+
map_vadd' _p _v := rfl
841+
831842
variable {P}
832843

833844
/-- No points are in `⊥`. -/
@@ -1588,6 +1599,29 @@ theorem map_span (s : Set P₁) : (affineSpan k s).map f = affineSpan k (f '' s)
15881599
subset_affineSpan k _ (mem_image_of_mem f hp)⟩
15891600
#align affine_subspace.map_span AffineSubspace.map_span
15901601

1602+
section inclusion
1603+
variable {S₁ S₂ : AffineSubspace k P₁} [Nonempty S₁] [Nonempty S₂]
1604+
1605+
attribute [local instance] AffineSubspace.toAddTorsor
1606+
1607+
/-- Affine map from a smaller to a larger subspace of the same space.
1608+
1609+
This is the affine version of `Submodule.inclusion`. -/
1610+
@[simps linear]
1611+
def inclusion (h : S₁ ≤ S₂) : S₁ →ᵃ[k] S₂ where
1612+
toFun := Set.inclusion h
1613+
linear := Submodule.inclusion <| AffineSubspace.direction_le h
1614+
map_vadd' _ _ := rfl
1615+
1616+
@[simp]
1617+
theorem coe_inclusion_apply (h : S₁ ≤ S₂) (x : S₁) : (inclusion h x : P₁) = x :=
1618+
rfl
1619+
1620+
@[simp]
1621+
theorem inclusion_rfl : inclusion (le_refl S₁) = AffineMap.id k S₁ := rfl
1622+
1623+
end inclusion
1624+
15911625
end AffineSubspace
15921626

15931627
namespace AffineMap
@@ -1607,6 +1641,33 @@ end AffineMap
16071641

16081642
namespace AffineEquiv
16091643

1644+
section ofEq
1645+
variable (S₁ S₂ : AffineSubspace k P₁) [Nonempty S₁] [Nonempty S₂]
1646+
1647+
attribute [local instance] AffineSubspace.toAddTorsor
1648+
1649+
/-- Affine equivalence between two equal affine subspace.
1650+
1651+
This is the affine version of `LinearEquiv.ofEq`. -/
1652+
@[simps linear]
1653+
def ofEq (h : S₁ = S₂) : S₁ ≃ᵃ[k] S₂ where
1654+
toEquiv := Equiv.Set.ofEq <| congr_arg _ h
1655+
linear := .ofEq _ _ <| congr_arg _ h
1656+
map_vadd' _ _ := rfl
1657+
1658+
@[simp]
1659+
theorem coe_ofEq_apply (h : S₁ = S₂) (x : S₁) : (ofEq S₁ S₂ h x : P₁) = x :=
1660+
rfl
1661+
1662+
@[simp]
1663+
theorem ofEq_symm (h : S₁ = S₂) : (ofEq S₁ S₂ h).symm = ofEq S₂ S₁ h.symm :=
1664+
rfl
1665+
1666+
@[simp]
1667+
theorem ofEq_rfl : ofEq S₁ S₁ rfl = AffineEquiv.refl k S₁ := rfl
1668+
1669+
end ofEq
1670+
16101671
theorem span_eq_top_iff {s : Set P₁} (e : P₁ ≃ᵃ[k] P₂) :
16111672
affineSpan k s = ⊤ ↔ affineSpan k (e '' s) = ⊤ := by
16121673
refine' ⟨(e : P₁ →ᵃ[k] P₂).span_eq_top_of_surjective e.surjective, _⟩

0 commit comments

Comments
 (0)