Documentation

Mathlib.Algebra.AddTorsor.Basic

Torsors of additive group actions #

Further results for torsors, that are not in Mathlib.Algebra.AddTorsor.Defs to avoid increasing imports there.

theorem Set.singleton_vsub_self {G : Type u_1} {P : Type u_2} [AddGroup G] [T : AddTorsor G P] (p : P) :
theorem vsub_left_cancel {G : Type u_1} {P : Type u_2} [AddGroup G] [T : AddTorsor G P] {p₁ p₂ p : P} (h : p₁ -ᵥ p = p₂ -ᵥ p) :
p₁ = p₂

If the same point subtracted from two points produces equal results, those points are equal.

@[simp]
theorem vsub_left_cancel_iff {G : Type u_1} {P : Type u_2} [AddGroup G] [T : AddTorsor G P] {p₁ p₂ p : P} :
p₁ -ᵥ p = p₂ -ᵥ p p₁ = p₂

The same point subtracted from two points produces equal results if and only if those points are equal.

theorem vsub_left_injective {G : Type u_1} {P : Type u_2} [AddGroup G] [T : AddTorsor G P] (p : P) :
Function.Injective fun (x : P) => x -ᵥ p

Subtracting the point p is an injective function.

theorem vsub_right_cancel {G : Type u_1} {P : Type u_2} [AddGroup G] [T : AddTorsor G P] {p₁ p₂ p : P} (h : p -ᵥ p₁ = p -ᵥ p₂) :
p₁ = p₂

If subtracting two points from the same point produces equal results, those points are equal.

@[simp]
theorem vsub_right_cancel_iff {G : Type u_1} {P : Type u_2} [AddGroup G] [T : AddTorsor G P] {p₁ p₂ p : P} :
p -ᵥ p₁ = p -ᵥ p₂ p₁ = p₂

Subtracting two points from the same point produces equal results if and only if those points are equal.

theorem vsub_right_injective {G : Type u_1} {P : Type u_2} [AddGroup G] [T : AddTorsor G P] (p : P) :
Function.Injective fun (x : P) => p -ᵥ x

Subtracting a point from the point p is an injective function.

@[simp]
theorem vsub_sub_vsub_cancel_left {G : Type u_1} {P : Type u_2} [AddCommGroup G] [AddTorsor G P] (p₁ p₂ p₃ : P) :
p₃ -ᵥ p₂ - (p₃ -ᵥ p₁) = p₁ -ᵥ p₂

Cancellation subtracting the results of two subtractions.

@[simp]
theorem vadd_vsub_vadd_cancel_left {G : Type u_1} {P : Type u_2} [AddCommGroup G] [AddTorsor G P] (v : G) (p₁ p₂ : P) :
(v +ᵥ p₁) -ᵥ (v +ᵥ p₂) = p₁ -ᵥ p₂
theorem vsub_vadd_comm {G : Type u_1} {P : Type u_2} [AddCommGroup G] [AddTorsor G P] (p₁ p₂ p₃ : P) :
(p₁ -ᵥ p₂) +ᵥ p₃ = (p₃ -ᵥ p₂) +ᵥ p₁
theorem vadd_eq_vadd_iff_sub_eq_vsub {G : Type u_1} {P : Type u_2} [AddCommGroup G] [AddTorsor G P] {v₁ v₂ : G} {p₁ p₂ : P} :
v₁ +ᵥ p₁ = v₂ +ᵥ p₂ v₂ - v₁ = p₁ -ᵥ p₂
theorem vsub_sub_vsub_comm {G : Type u_1} {P : Type u_2} [AddCommGroup G] [AddTorsor G P] (p₁ p₂ p₃ p₄ : P) :
p₁ -ᵥ p₂ - (p₃ -ᵥ p₄) = p₁ -ᵥ p₃ - (p₂ -ᵥ p₄)
@[simp]
theorem Set.vadd_set_vsub_vadd_set {G : Type u_1} {P : Type u_2} [AddCommGroup G] [AddTorsor G P] (v : G) (s t : Set P) :
(v +ᵥ s) -ᵥ (v +ᵥ t) = s -ᵥ t
instance Prod.instAddTorsor {G : Type u_1} {G' : Type u_2} {P : Type u_3} {P' : Type u_4} [AddGroup G] [AddGroup G'] [AddTorsor G P] [AddTorsor G' P'] :
AddTorsor (G × G') (P × P')
Equations
@[simp]
theorem Prod.fst_vadd {G : Type u_1} {G' : Type u_2} {P : Type u_3} {P' : Type u_4} [AddGroup G] [AddGroup G'] [AddTorsor G P] [AddTorsor G' P'] (v : G × G') (p : P × P') :
(v +ᵥ p).1 = v.1 +ᵥ p.1
@[simp]
theorem Prod.snd_vadd {G : Type u_1} {G' : Type u_2} {P : Type u_3} {P' : Type u_4} [AddGroup G] [AddGroup G'] [AddTorsor G P] [AddTorsor G' P'] (v : G × G') (p : P × P') :
(v +ᵥ p).2 = v.2 +ᵥ p.2
@[simp]
theorem Prod.mk_vadd_mk {G : Type u_1} {G' : Type u_2} {P : Type u_3} {P' : Type u_4} [AddGroup G] [AddGroup G'] [AddTorsor G P] [AddTorsor G' P'] (v : G) (v' : G') (p : P) (p' : P') :
(v, v') +ᵥ (p, p') = (v +ᵥ p, v' +ᵥ p')
@[simp]
theorem Prod.fst_vsub {G : Type u_1} {G' : Type u_2} {P : Type u_3} {P' : Type u_4} [AddGroup G] [AddGroup G'] [AddTorsor G P] [AddTorsor G' P'] (p₁ p₂ : P × P') :
(p₁ -ᵥ p₂).1 = p₁.1 -ᵥ p₂.1
@[simp]
theorem Prod.snd_vsub {G : Type u_1} {G' : Type u_2} {P : Type u_3} {P' : Type u_4} [AddGroup G] [AddGroup G'] [AddTorsor G P] [AddTorsor G' P'] (p₁ p₂ : P × P') :
(p₁ -ᵥ p₂).2 = p₁.2 -ᵥ p₂.2
@[simp]
theorem Prod.mk_vsub_mk {G : Type u_1} {G' : Type u_2} {P : Type u_3} {P' : Type u_4} [AddGroup G] [AddGroup G'] [AddTorsor G P] [AddTorsor G' P'] (p₁ p₂ : P) (p₁' p₂' : P') :
(p₁, p₁') -ᵥ (p₂, p₂') = (p₁ -ᵥ p₂, p₁' -ᵥ p₂')
instance Pi.instAddTorsor {I : Type u} {fg : IType v} [(i : I) → AddGroup (fg i)] {fp : IType w} [(i : I) → AddTorsor (fg i) (fp i)] :
AddTorsor ((i : I) → fg i) ((i : I) → fp i)

A product of AddTorsors is an AddTorsor.

Equations
@[simp]
theorem Equiv.constVAdd_zero (G : Type u_1) (P : Type u_2) [AddGroup G] [AddTorsor G P] :
constVAdd P 0 = 1
@[simp]
theorem Equiv.constVAdd_add {G : Type u_1} (P : Type u_2) [AddGroup G] [AddTorsor G P] (v₁ v₂ : G) :
constVAdd P (v₁ + v₂) = constVAdd P v₁ * constVAdd P v₂
def Equiv.constVAddHom {G : Type u_1} (P : Type u_2) [AddGroup G] [AddTorsor G P] :

Equiv.constVAdd as a homomorphism from Multiplicative G to Equiv.perm P

Equations
Instances For
    @[simp]
    theorem Equiv.left_vsub_pointReflection {G : Type u_1} {P : Type u_2} [AddGroup G] [AddTorsor G P] (x y : P) :
    @[simp]
    theorem Equiv.right_vsub_pointReflection {G : Type u_1} {P : Type u_2} [AddGroup G] [AddTorsor G P] (x y : P) :
    y -ᵥ (pointReflection x) y = 2 (y -ᵥ x)
    theorem Equiv.pointReflection_fixed_iff_of_injective_two_nsmul {G : Type u_1} {P : Type u_2} [AddGroup G] [AddTorsor G P] {x y : P} (h : Function.Injective fun (x : G) => 2 x) :
    (pointReflection x) y = y y = x

    x is the only fixed point of pointReflection x. This lemma requires x + x = y + y ↔ x = y. There is no typeclass to use here, so we add it as an explicit argument.

    @[deprecated Equiv.pointReflection_fixed_iff_of_injective_two_nsmul (since := "2024-11-18")]
    theorem Equiv.pointReflection_fixed_iff_of_injective_bit0 {G : Type u_1} {P : Type u_2} [AddGroup G] [AddTorsor G P] {x y : P} (h : Function.Injective fun (x : G) => 2 x) :
    (pointReflection x) y = y y = x

    Alias of Equiv.pointReflection_fixed_iff_of_injective_two_nsmul.


    x is the only fixed point of pointReflection x. This lemma requires x + x = y + y ↔ x = y. There is no typeclass to use here, so we add it as an explicit argument.

    theorem Equiv.injective_pointReflection_left_of_injective_two_nsmul {G : Type u_3} {P : Type u_4} [AddCommGroup G] [AddTorsor G P] (h : Function.Injective fun (x : G) => 2 x) (y : P) :
    Function.Injective fun (x : P) => (pointReflection x) y
    @[deprecated Equiv.injective_pointReflection_left_of_injective_two_nsmul (since := "2024-11-18")]
    theorem Equiv.injective_pointReflection_left_of_injective_bit0 {G : Type u_3} {P : Type u_4} [AddCommGroup G] [AddTorsor G P] (h : Function.Injective fun (x : G) => 2 x) (y : P) :
    Function.Injective fun (x : P) => (pointReflection x) y

    Alias of Equiv.injective_pointReflection_left_of_injective_two_nsmul.