Documentation

Mathlib.Algebra.Ring.Parity

Even and odd elements in rings #

This file defines odd elements and proves some general facts about even and odd elements of rings.

As opposed to Even, Odd does not have a multiplicative counterpart.

TODO #

Try to generalize Even lemmas further. For example, there are still a few lemmas whose Semiring assumptions I (DT) am not convinced are necessary. If that turns out to be true, they could be moved to Algebra.Group.Even.

See also #

Algebra.Group.Even for the definition of even elements.

@[simp]
theorem Even.neg_pow {α : Type u_2} [Monoid α] [HasDistribNeg α] {n : } :
Even n∀ (a : α), (-a) ^ n = a ^ n
theorem Even.neg_one_pow {α : Type u_2} [Monoid α] [HasDistribNeg α] {n : } (h : Even n) :
(-1) ^ n = 1
theorem Even.neg_zpow {α : Type u_2} [DivisionMonoid α] [HasDistribNeg α] {n : } :
Even n∀ (a : α), (-a) ^ n = a ^ n
theorem Even.neg_one_zpow {α : Type u_2} [DivisionMonoid α] [HasDistribNeg α] {n : } (h : Even n) :
(-1) ^ n = 1
@[simp]
theorem isSquare_zero {α : Type u_2} [MulZeroClass α] :
theorem even_iff_exists_two_mul {α : Type u_2} [Semiring α] {a : α} :
Even a ∃ (b : α), a = 2 * b
theorem even_iff_two_dvd {α : Type u_2} [Semiring α] {a : α} :
Even a 2 a
theorem Even.two_dvd {α : Type u_2} [Semiring α] {a : α} :
Even a2 a

Alias of the forward direction of even_iff_two_dvd.

theorem Even.trans_dvd {α : Type u_2} [Semiring α] {a : α} {b : α} (ha : Even a) (hab : a b) :
theorem Dvd.dvd.even {α : Type u_2} [Semiring α] {a : α} {b : α} (hab : a b) (ha : Even a) :
@[simp]
theorem range_two_mul (α : Type u_5) [Semiring α] :
(Set.range fun (x : α) => 2 * x) = {a : α | Even a}
@[simp]
theorem even_bit0 {α : Type u_2} [Semiring α] (a : α) :
@[simp]
theorem even_two {α : Type u_2} [Semiring α] :
@[simp]
theorem Even.mul_left {α : Type u_2} [Semiring α] {a : α} (ha : Even a) (b : α) :
Even (b * a)
@[simp]
theorem Even.mul_right {α : Type u_2} [Semiring α] {a : α} (ha : Even a) (b : α) :
Even (a * b)
theorem even_two_mul {α : Type u_2} [Semiring α] (a : α) :
Even (2 * a)
theorem Even.pow_of_ne_zero {α : Type u_2} [Semiring α] {a : α} (ha : Even a) {n : } :
n 0Even (a ^ n)
def Odd {α : Type u_2} [Semiring α] (a : α) :

An element a of a semiring is odd if there exists k such a = 2*k + 1.

Equations
  • Odd a = ∃ (k : α), a = 2 * k + 1
Instances For
    theorem odd_iff_exists_bit1 {α : Type u_2} [Semiring α] {a : α} :
    Odd a ∃ (b : α), a = bit1 b
    theorem Odd.exists_bit1 {α : Type u_2} [Semiring α] {a : α} :
    Odd a∃ (b : α), a = bit1 b

    Alias of the forward direction of odd_iff_exists_bit1.

    @[simp]
    theorem odd_bit1 {α : Type u_2} [Semiring α] (a : α) :
    Odd (bit1 a)
    @[simp]
    theorem range_two_mul_add_one (α : Type u_5) [Semiring α] :
    (Set.range fun (x : α) => 2 * x + 1) = {a : α | Odd a}
    theorem Even.add_odd {α : Type u_2} [Semiring α] {a : α} {b : α} :
    Even aOdd bOdd (a + b)
    theorem Even.odd_add {α : Type u_2} [Semiring α] {a : α} {b : α} (ha : Even a) (hb : Odd b) :
    Odd (b + a)
    theorem Odd.add_even {α : Type u_2} [Semiring α] {a : α} {b : α} (ha : Odd a) (hb : Even b) :
    Odd (a + b)
    theorem Odd.add_odd {α : Type u_2} [Semiring α] {a : α} {b : α} :
    Odd aOdd bEven (a + b)
    @[simp]
    theorem odd_one {α : Type u_2} [Semiring α] :
    Odd 1
    @[simp]
    theorem Even.add_one {α : Type u_2} [Semiring α] {a : α} (h : Even a) :
    Odd (a + 1)
    @[simp]
    theorem Even.one_add {α : Type u_2} [Semiring α] {a : α} (h : Even a) :
    Odd (1 + a)
    theorem odd_two_mul_add_one {α : Type u_2} [Semiring α] (a : α) :
    Odd (2 * a + 1)
    @[simp]
    theorem odd_add_self_one' {α : Type u_2} [Semiring α] {a : α} :
    Odd (a + (a + 1))
    @[simp]
    theorem odd_add_one_self {α : Type u_2} [Semiring α] {a : α} :
    Odd (a + 1 + a)
    @[simp]
    theorem odd_add_one_self' {α : Type u_2} [Semiring α] {a : α} :
    Odd (a + (1 + a))
    theorem Odd.map {F : Type u_1} {α : Type u_2} {β : Type u_3} [Semiring α] [Semiring β] {a : α} [FunLike F α β] [RingHomClass F α β] (f : F) :
    Odd aOdd (f a)
    @[simp]
    theorem Odd.mul {α : Type u_2} [Semiring α] {a : α} {b : α} :
    Odd aOdd bOdd (a * b)
    theorem Odd.pow {α : Type u_2} [Semiring α] {a : α} (ha : Odd a) {n : } :
    Odd (a ^ n)
    theorem Odd.pow_add_pow_eq_zero {α : Type u_2} [Semiring α] {a : α} {b : α} {n : } [IsCancelAdd α] (hn : Odd n) (hab : a + b = 0) :
    a ^ n + b ^ n = 0
    theorem Odd.neg_pow {α : Type u_2} [Monoid α] [HasDistribNeg α] {n : } :
    Odd n∀ (a : α), (-a) ^ n = -a ^ n
    @[simp]
    theorem Odd.neg_one_pow {α : Type u_2} [Monoid α] [HasDistribNeg α] {n : } (h : Odd n) :
    (-1) ^ n = -1
    theorem even_neg_two {α : Type u_2} [Ring α] :
    Even (-2)
    theorem Odd.neg {α : Type u_2} [Ring α] {a : α} (hp : Odd a) :
    Odd (-a)
    @[simp]
    theorem odd_neg {α : Type u_2} [Ring α] {a : α} :
    Odd (-a) Odd a
    theorem odd_neg_one {α : Type u_2} [Ring α] :
    Odd (-1)
    theorem Odd.sub_even {α : Type u_2} [Ring α] {a : α} {b : α} (ha : Odd a) (hb : Even b) :
    Odd (a - b)
    theorem Even.sub_odd {α : Type u_2} [Ring α] {a : α} {b : α} (ha : Even a) (hb : Odd b) :
    Odd (a - b)
    theorem Odd.sub_odd {α : Type u_2} [Ring α] {a : α} {b : α} (ha : Odd a) (hb : Odd b) :
    Even (a - b)