Documentation

Mathlib.Algebra.Group.Even

Squares and even elements #

This file defines square and even elements in a monoid.

Main declarations #

Note #

TODO #

See also #

Mathlib/Algebra/Ring/Parity.lean for the definition of odd elements as well as facts about Even / IsSquare in rings.

def IsSquare {α : Type u_2} [Mul α] (a : α) :

An element a of a type α with multiplication satisfies IsSquare a if a = r * r, for some root r : α.

Equations
    Instances For
      def Even {α : Type u_2} [Add α] (a : α) :

      An element a of a type α with addition satisfies Even a if a = r + r, for some r : α.

      Equations
        Instances For
          theorem isSquare_iff_exists_mul_self {α : Type u_2} [Mul α] (a : α) :
          IsSquare a (r : α), a = r * r
          theorem even_iff_exists_add_self {α : Type u_2} [Add α] (a : α) :
          Even a (r : α), a = r + r
          theorem IsSquare.exists_mul_self {α : Type u_2} [Mul α] (a : α) :
          IsSquare a (r : α), a = r * r

          Alias of the forward direction of isSquare_iff_exists_mul_self.

          theorem Even.exists_add_self {α : Type u_2} [Add α] (a : α) :
          Even a (r : α), a = r + r
          @[simp]
          theorem IsSquare.mul_self {α : Type u_2} [Mul α] (r : α) :
          IsSquare (r * r)
          @[simp]
          theorem Even.add_self {α : Type u_2} [Add α] (r : α) :
          Even (r + r)
          theorem isSquare_op_iff {α : Type u_2} [Mul α] {a : α} :
          theorem even_op_iff {α : Type u_2} [Add α] {a : α} :
          theorem even_unop_iff {α : Type u_2} [Add α] {a : αᵃᵒᵖ} :
          @[simp]
          theorem even_ofMul_iff {α : Type u_2} [Mul α] {a : α} :
          @[simp]
          theorem isSquare_toMul_iff {α : Type u_2} [Mul α] {a : Additive α} :
          @[simp]
          theorem isSquare_ofAdd_iff {α : Type u_2} [Add α] {a : α} :
          @[simp]
          theorem even_toAdd_iff {α : Type u_2} [Add α] {a : Multiplicative α} :
          @[simp]
          theorem IsSquare.one {α : Type u_2} [MulOneClass α] :
          @[simp]
          theorem Even.zero {α : Type u_2} [AddZeroClass α] :
          theorem IsSquare.map {F : Type u_1} {α : Type u_2} {β : Type u_3} [MulOneClass α] [MulOneClass β] [FunLike F α β] [MonoidHomClass F α β] {a : α} (f : F) :
          IsSquare aIsSquare (f a)
          theorem Even.map {F : Type u_1} {α : Type u_2} {β : Type u_3} [AddZeroClass α] [AddZeroClass β] [FunLike F α β] [AddMonoidHomClass F α β] {a : α} (f : F) :
          Even aEven (f a)
          theorem isSquare_subset_image_isSquare {F : Type u_1} {α : Type u_2} {β : Type u_3} [MulOneClass α] [MulOneClass β] [FunLike F α β] [MonoidHomClass F α β] {f : F} (hf : Function.Surjective f) :
          {b : β | IsSquare b} f '' {a : α | IsSquare a}
          theorem even_subset_image_even {F : Type u_1} {α : Type u_2} {β : Type u_3} [AddZeroClass α] [AddZeroClass β] [FunLike F α β] [AddMonoidHomClass F α β] {f : F} (hf : Function.Surjective f) :
          {b : β | Even b} f '' {a : α | Even a}
          theorem isSquare_iff_exists_sq {α : Type u_2} [Monoid α] (a : α) :
          IsSquare a (r : α), a = r ^ 2
          theorem even_iff_exists_two_nsmul {α : Type u_2} [AddMonoid α] (a : α) :
          Even a (r : α), a = 2 r
          theorem IsSquare.exists_sq {α : Type u_2} [Monoid α] (a : α) :
          IsSquare a (r : α), a = r ^ 2

          Alias of the forward direction of isSquare_iff_exists_sq.

          theorem Even.exists_two_nsmul {α : Type u_2} [AddMonoid α] (a : α) :
          Even a (r : α), a = 2 r

          Alias of the forwards direction of even_iff_exists_two_nsmul.

          theorem IsSquare.sq {α : Type u_2} [Monoid α] (r : α) :
          IsSquare (r ^ 2)
          theorem Even.two_nsmul {α : Type u_2} [AddMonoid α] (r : α) :
          Even (2 r)
          theorem IsSquare.pow {α : Type u_2} [Monoid α] {a : α} (n : ) (ha : IsSquare a) :
          IsSquare (a ^ n)
          theorem Even.nsmul_right {α : Type u_2} [AddMonoid α] {a : α} (n : ) (ha : Even a) :
          Even (n a)
          theorem Even.isSquare_pow {α : Type u_2} [Monoid α] {n : } (hn : Even n) (a : α) :
          IsSquare (a ^ n)
          theorem Even.nsmul_left {α : Type u_2} [AddMonoid α] {n : } (hn : Even n) (a : α) :
          Even (n a)
          theorem IsSquare.mul {α : Type u_2} [CommSemigroup α] {a b : α} :
          IsSquare aIsSquare bIsSquare (a * b)
          theorem Even.add {α : Type u_2} [AddCommSemigroup α] {a b : α} :
          Even aEven bEven (a + b)
          @[simp]
          theorem isSquare_inv {α : Type u_2} [DivisionMonoid α] {a : α} :
          @[simp]
          theorem even_neg {α : Type u_2} [SubtractionMonoid α] {a : α} :
          Even (-a) Even a
          theorem IsSquare.inv {α : Type u_2} [DivisionMonoid α] {a : α} :

          Alias of the reverse direction of isSquare_inv.

          theorem Even.neg {α : Type u_2} [SubtractionMonoid α] {a : α} :
          Even aEven (-a)
          theorem IsSquare.zpow {α : Type u_2} [DivisionMonoid α] {a : α} (n : ) :
          IsSquare aIsSquare (a ^ n)
          theorem Even.zsmul_right {α : Type u_2} [SubtractionMonoid α] {a : α} (n : ) :
          Even aEven (n a)
          theorem IsSquare.div {α : Type u_2} [DivisionCommMonoid α] {a b : α} (ha : IsSquare a) (hb : IsSquare b) :
          IsSquare (a / b)
          theorem Even.sub {α : Type u_2} [SubtractionCommMonoid α] {a b : α} (ha : Even a) (hb : Even b) :
          Even (a - b)
          @[simp]
          theorem Even.isSquare_zpow {α : Type u_2} [Group α] {n : } :
          Even n∀ (a : α), IsSquare (a ^ n)
          @[simp]
          theorem Even.zsmul_left {α : Type u_2} [AddGroup α] {n : } :
          Even n∀ (a : α), Even (n a)