Skip to content

Commit

Permalink
feat(linear_algebra/orientation): eq_neg_iff_eq_neg (#11629)
Browse files Browse the repository at this point in the history
Add two more `module.ray` lemmas about negation.
  • Loading branch information
jsm28 committed Jan 24, 2022
1 parent 7ddb9a3 commit 9ef7f6b
Showing 1 changed file with 11 additions and 0 deletions.
11 changes: 11 additions & 0 deletions src/linear_algebra/orientation.lean
Expand Up @@ -341,6 +341,17 @@ variables {R}
@[simp] protected lemma neg_neg [nontrivial R] (x : module.ray R M) : -(-x) = x :=
quotient.ind (λ a, congr_arg quotient.mk $ ray_vector.neg_neg _) x

variables (R M)

/-- Negating a ray is involutive. -/
lemma neg_involutive [nontrivial R] : function.involutive (λ x : module.ray R M, -x) :=
λ x, module.ray.neg_neg x

variables {R M}

protected lemma eq_neg_iff_eq_neg [nontrivial R] (x y : module.ray R M) : x = -y ↔ y = -x :=
by rw [←module.ray.neg_neg x, (neg_involutive R M).injective.eq_iff, module.ray.neg_neg x, eq_comm]

/-- A ray does not equal its own negation. -/
lemma ne_neg_self [nontrivial R] [no_zero_smul_divisors R M] (x : module.ray R M) : x ≠ -x :=
begin
Expand Down

0 comments on commit 9ef7f6b

Please sign in to comment.