Algebraic structures on the set of positive numbers #
In this file we define various instances (AddSemigroup
, OrderedCommMonoid
etc) on the
type {x : R // 0 < x}
. In each case we try to require the weakest possible typeclass
assumptions on R
but possibly, there is a room for improvements.
Equations
- Positive.addSemigroup = { toAdd := Positive.instAddSubtypeLtOfNat_mathlib, add_assoc := ⋯ }
instance
Positive.addCommSemigroup
{M : Type u_3}
[AddCommMonoid M]
[Preorder M]
[AddLeftStrictMono M]
:
Equations
- Positive.addCommSemigroup = { toAddSemigroup := Positive.addSemigroup, add_comm := ⋯ }
instance
Positive.addLeftCancelSemigroup
{M : Type u_3}
[AddLeftCancelMonoid M]
[Preorder M]
[AddLeftStrictMono M]
:
Equations
- Positive.addLeftCancelSemigroup = { toAddSemigroup := Positive.addSemigroup, add_left_cancel := ⋯ }
instance
Positive.addRightCancelSemigroup
{M : Type u_3}
[AddRightCancelMonoid M]
[Preorder M]
[AddLeftStrictMono M]
:
Equations
- Positive.addRightCancelSemigroup = { toAddSemigroup := Positive.addSemigroup, add_right_cancel := ⋯ }
instance
Positive.addLeftStrictMono
{M : Type u_1}
[AddMonoid M]
[Preorder M]
[AddLeftStrictMono M]
:
instance
Positive.addRightStrictMono
{M : Type u_1}
[AddMonoid M]
[Preorder M]
[AddLeftStrictMono M]
[AddRightStrictMono M]
:
instance
Positive.addLeftReflectLT
{M : Type u_1}
[AddMonoid M]
[Preorder M]
[AddLeftStrictMono M]
[AddLeftReflectLT M]
:
instance
Positive.addRightReflectLT
{M : Type u_1}
[AddMonoid M]
[Preorder M]
[AddLeftStrictMono M]
[AddRightReflectLT M]
:
instance
Positive.addLeftReflectLE
{M : Type u_1}
[AddMonoid M]
[Preorder M]
[AddLeftStrictMono M]
[AddLeftReflectLE M]
:
instance
Positive.addRightReflectLE
{M : Type u_1}
[AddMonoid M]
[Preorder M]
[AddLeftStrictMono M]
[AddRightReflectLE M]
:
Equations
- Positive.instSemigroupSubtypeLtOfNat = { toMul := Positive.instMulSubtypeLtOfNat_mathlib, mul_assoc := ⋯ }
Equations
- Positive.instDistribSubtypeLtOfNat = { toMul := Positive.instMulSubtypeLtOfNat_mathlib, toAdd := Positive.instAddSubtypeLtOfNat_mathlib, left_distrib := ⋯, right_distrib := ⋯ }
Equations
- Positive.instOneSubtypeLtOfNat = { one := ⟨1, ⋯⟩ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- Positive.orderedCommMonoid = { toMonoid := Positive.instMonoidSubtypeLtOfNat, mul_comm := ⋯, toPartialOrder := Subtype.partialOrder fun (x : R) => 0 < x, mul_le_mul_left := ⋯ }
If R
is a nontrivial linear ordered commutative semiring, then {x : R // 0 < x}
is a linear
ordered cancellative commutative monoid.
Equations
- One or more equations did not get rendered due to their size.