Squares, even and odd elements #
This file proves some general facts about squares, even and odd elements of semirings.
In the implementation, we define IsSquare
and we let Even
be the notion transported by
to_additive
. The definition are therefore as follows:
IsSquare a ↔ ∃ r, a = r * r
Even a ↔ ∃ r, a = r + r
Odd elements are not unified with a multiplicative notion.
Future work #
- TODO: Try to generalize further the typeclass assumptions on
IsSquare/Even
. For instance, in some cases, there areSemiring
assumptions that I (DT) am not convinced are necessary. - TODO: Consider moving the definition and lemmas about
Odd
to a separate file. - TODO: The "old" definition of
Even a
asked for the existence of an elementc
such thata = 2 * c
. For this reason, several fixes introduce an extratwo_mul
or← two_mul
. It might be the case that by making a careful choice ofsimp
lemma, this can be avoided.
abbrev
even_op_iff.match_1
{α : Type u_1}
[Add α]
{a : α}
(motive : Even (AddOpposite.op a) → Prop)
:
∀ (x : Even (AddOpposite.op a)), (∀ (c : αᵃᵒᵖ) (hc : AddOpposite.op a = c + c), motive ⋯) → motive x
Equations
- ⋯ = ⋯
Instances For
theorem
isSquare_unop_iff
{α : Type u_2}
[Mul α]
{a : αᵐᵒᵖ}
:
IsSquare (MulOpposite.unop a) ↔ IsSquare a
instance
instDecidablePredAddOppositeEvenAdd
{α : Type u_2}
[Add α]
[DecidablePred Even]
:
DecidablePred Even
Equations
instance
instDecidablePredMulOppositeIsSquareMul
{α : Type u_2}
[Mul α]
[DecidablePred IsSquare]
:
DecidablePred IsSquare
Equations
instance
instDecidablePredAdditiveEvenAdd
{α : Type u_2}
[Mul α]
[DecidablePred IsSquare]
:
DecidablePred Even
Equations
- instDecidablePredAdditiveEvenAdd x = decidable_of_iff (IsSquare (Additive.toMul x)) ⋯
@[simp]
instance
instDecidablePredMultiplicativeIsSquareMul
{α : Type u_2}
[Add α]
[DecidablePred Even]
:
DecidablePred IsSquare
Equations
- instDecidablePredMultiplicativeIsSquareMul x = decidable_of_iff (Even (Multiplicative.toAdd x)) ⋯
theorem
Even.map
{F : Type u_1}
{α : Type u_2}
{β : Type u_3}
[AddZeroClass α]
[AddZeroClass β]
[FunLike F α β]
[AddMonoidHomClass F α β]
{m : α}
(f : F)
:
theorem
IsSquare.map
{F : Type u_1}
{α : Type u_2}
{β : Type u_3}
[MulOneClass α]
[MulOneClass β]
[FunLike F α β]
[MonoidHomClass F α β]
{m : α}
(f : F)
:
Alias of the forward direction of isSquare_iff_exists_sq
.
Alias of the reverse direction of isSquare_iff_exists_sq
.
Alias of the forwards direction of even_iff_exists_two_nsmul
.
@[simp]
Alias of the reverse direction of isSquare_inv
.
theorem
Even.neg_one_zpow
{α : Type u_2}
[DivisionMonoid α]
[HasDistribNeg α]
{n : ℤ}
(h : Even n)
:
theorem
IsSquare.div
{α : Type u_2}
[DivisionCommMonoid α]
{a : α}
{b : α}
(ha : IsSquare a)
(hb : IsSquare b)
:
theorem
Even.tsub
{α : Type u_2}
[CanonicallyLinearOrderedAddCommMonoid α]
[Sub α]
[OrderedSub α]
[ContravariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
{m : α}
{n : α}
(hm : Even m)
(hn : Even n)
:
Alias of the forward direction of even_iff_exists_bit0
.
Alias of the forward direction of even_iff_two_dvd
.
Alias of the forward direction of odd_iff_exists_bit1
.
@[simp]
theorem
Odd.pos
{α : Type u_2}
[CanonicallyOrderedCommSemiring α]
[Nontrivial α]
{n : α}
(hn : Odd n)
:
0 < n
theorem
Even.pow_pos
{R : Type u_4}
[LinearOrderedRing R]
{a : R}
{n : ℕ}
(hn : Even n)
(ha : a ≠ 0)
:
theorem
Odd.pow_nonpos
{R : Type u_4}
[LinearOrderedRing R]
{a : R}
{n : ℕ}
(hn : Odd n)
(ha : a ≤ 0)
:
theorem
Odd.pow_neg
{R : Type u_4}
[LinearOrderedRing R]
{a : R}
{n : ℕ}
(hn : Odd n)
(ha : a < 0)
:
@[simp]
theorem
Odd.strictMono_pow
{R : Type u_4}
[LinearOrderedRing R]
{n : ℕ}
(hn : Odd n)
:
StrictMono fun (a : R) => a ^ n