mathlib documentation

algebra.ordered_ring

@[instance]
def ordered_semiring.to_semiring (α : Type u) [s : ordered_semiring α] :

@[class]
structure ordered_semiring (α : Type u) :
Type u
  • add : α → α → α
  • add_assoc : ∀ (a b c_1 : α), a + b + c_1 = a + (b + c_1)
  • zero : α
  • zero_add : ∀ (a : α), 0 + a = a
  • add_zero : ∀ (a : α), a + 0 = a
  • add_comm : ∀ (a b : α), a + b = b + a
  • mul : α → α → α
  • mul_assoc : ∀ (a b c_1 : α), (a * b) * c_1 = a * b * c_1
  • one : α
  • one_mul : ∀ (a : α), 1 * a = a
  • mul_one : ∀ (a : α), a * 1 = a
  • zero_mul : ∀ (a : α), 0 * a = 0
  • mul_zero : ∀ (a : α), a * 0 = 0
  • left_distrib : ∀ (a b c_1 : α), a * (b + c_1) = a * b + a * c_1
  • right_distrib : ∀ (a b c_1 : α), (a + b) * c_1 = a * c_1 + b * c_1
  • add_left_cancel : ∀ (a b c_1 : α), a + b = a + c_1b = c_1
  • add_right_cancel : ∀ (a b c_1 : α), a + b = c_1 + ba = c_1
  • le : α → α → Prop
  • lt : α → α → Prop
  • le_refl : ∀ (a : α), a a
  • le_trans : ∀ (a b c_1 : α), a bb c_1a c_1
  • lt_iff_le_not_le : (∀ (a b : α), a < b a b ¬b a) . "order_laws_tac"
  • le_antisymm : ∀ (a b : α), a bb aa = b
  • add_le_add_left : ∀ (a b : α), a b∀ (c_1 : α), c_1 + a c_1 + b
  • le_of_add_le_add_left : ∀ (a b c_1 : α), a + b a + c_1b c_1
  • mul_lt_mul_of_pos_left : ∀ (a b c_1 : α), a < b0 < c_1c_1 * a < c_1 * b
  • mul_lt_mul_of_pos_right : ∀ (a b c_1 : α), a < b0 < c_1a * c_1 < b * c_1

An ordered_semiring α is a semiring α with a partial order such that multiplication with a positive number and addition are monotone.

Instances
theorem mul_lt_mul_of_pos_left {α : Type u} [ordered_semiring α] {a b c : α} (h₁ : a < b) (h₂ : 0 < c) :
c * a < c * b

theorem mul_lt_mul_of_pos_right {α : Type u} [ordered_semiring α] {a b c : α} (h₁ : a < b) (h₂ : 0 < c) :
a * c < b * c

theorem mul_le_mul_of_nonneg_left {α : Type u} [ordered_semiring α] {a b c : α} (h₁ : a b) (h₂ : 0 c) :
c * a c * b

theorem mul_le_mul_of_nonneg_right {α : Type u} [ordered_semiring α] {a b c : α} (h₁ : a b) (h₂ : 0 c) :
a * c b * c

theorem mul_le_mul {α : Type u} [ordered_semiring α] {a b c d : α} (hac : a c) (hbd : b d) (nn_b : 0 b) (nn_c : 0 c) :
a * b c * d

theorem mul_nonneg {α : Type u} [ordered_semiring α] {a b : α} (ha : 0 a) (hb : 0 b) :
0 a * b

theorem mul_nonpos_of_nonneg_of_nonpos {α : Type u} [ordered_semiring α] {a b : α} (ha : 0 a) (hb : b 0) :
a * b 0

theorem mul_nonpos_of_nonpos_of_nonneg {α : Type u} [ordered_semiring α] {a b : α} (ha : a 0) (hb : 0 b) :
a * b 0

theorem mul_lt_mul {α : Type u} [ordered_semiring α] {a b c d : α} (hac : a < c) (hbd : b d) (pos_b : 0 < b) (nn_c : 0 c) :
a * b < c * d

theorem mul_lt_mul' {α : Type u} [ordered_semiring α] {a b c d : α} (h1 : a c) (h2 : b < d) (h3 : 0 b) (h4 : 0 < c) :
a * b < c * d

theorem mul_pos {α : Type u} [ordered_semiring α] {a b : α} (ha : 0 < a) (hb : 0 < b) :
0 < a * b

theorem mul_neg_of_pos_of_neg {α : Type u} [ordered_semiring α] {a b : α} (ha : 0 < a) (hb : b < 0) :
a * b < 0

theorem mul_neg_of_neg_of_pos {α : Type u} [ordered_semiring α] {a b : α} (ha : a < 0) (hb : 0 < b) :
a * b < 0

theorem mul_self_le_mul_self {α : Type u} [ordered_semiring α] {a b : α} (h1 : 0 a) (h2 : a b) :
a * a b * b

theorem mul_self_lt_mul_self {α : Type u} [ordered_semiring α] {a b : α} (h1 : 0 a) (h2 : a < b) :
a * a < b * b

@[class]
structure linear_ordered_semiring (α : Type u) :
Type u
  • add : α → α → α
  • add_assoc : ∀ (a b c_1 : α), a + b + c_1 = a + (b + c_1)
  • zero : α
  • zero_add : ∀ (a : α), 0 + a = a
  • add_zero : ∀ (a : α), a + 0 = a
  • add_comm : ∀ (a b : α), a + b = b + a
  • mul : α → α → α
  • mul_assoc : ∀ (a b c_1 : α), (a * b) * c_1 = a * b * c_1
  • one : α
  • one_mul : ∀ (a : α), 1 * a = a
  • mul_one : ∀ (a : α), a * 1 = a
  • zero_mul : ∀ (a : α), 0 * a = 0
  • mul_zero : ∀ (a : α), a * 0 = 0
  • left_distrib : ∀ (a b c_1 : α), a * (b + c_1) = a * b + a * c_1
  • right_distrib : ∀ (a b c_1 : α), (a + b) * c_1 = a * c_1 + b * c_1
  • add_left_cancel : ∀ (a b c_1 : α), a + b = a + c_1b = c_1
  • add_right_cancel : ∀ (a b c_1 : α), a + b = c_1 + ba = c_1
  • le : α → α → Prop
  • lt : α → α → Prop
  • le_refl : ∀ (a : α), a a
  • le_trans : ∀ (a b c_1 : α), a bb c_1a c_1
  • lt_iff_le_not_le : (∀ (a b : α), a < b a b ¬b a) . "order_laws_tac"
  • le_antisymm : ∀ (a b : α), a bb aa = b
  • add_le_add_left : ∀ (a b : α), a b∀ (c_1 : α), c_1 + a c_1 + b
  • le_of_add_le_add_left : ∀ (a b c_1 : α), a + b a + c_1b c_1
  • mul_lt_mul_of_pos_left : ∀ (a b c_1 : α), a < b0 < c_1c_1 * a < c_1 * b
  • mul_lt_mul_of_pos_right : ∀ (a b c_1 : α), a < b0 < c_1a * c_1 < b * c_1
  • le_total : ∀ (a b : α), a b b a
  • zero_lt_one : linear_ordered_semiring.zero < linear_ordered_semiring.one

A linear_ordered_semiring α is a semiring α with a linear order such that multiplication with a positive number and addition are monotone.

Instances
theorem zero_lt_one {α : Type u} [linear_ordered_semiring α] :
0 < 1

theorem zero_le_one {α : Type u} [linear_ordered_semiring α] :
0 1

theorem zero_lt_two {α : Type u} [linear_ordered_semiring α] :
0 < 2

theorem two_ne_zero {α : Type u} [linear_ordered_semiring α] :
2 0

theorem one_lt_two {α : Type u} [linear_ordered_semiring α] :
1 < 2

theorem one_le_two {α : Type u} [linear_ordered_semiring α] :
1 2

theorem zero_lt_four {α : Type u} [linear_ordered_semiring α] :
0 < 4

theorem lt_of_mul_lt_mul_left {α : Type u} [linear_ordered_semiring α] {a b c : α} (h : c * a < c * b) (hc : 0 c) :
a < b

theorem lt_of_mul_lt_mul_right {α : Type u} [linear_ordered_semiring α] {a b c : α} (h : a * c < b * c) (hc : 0 c) :
a < b

theorem le_of_mul_le_mul_left {α : Type u} [linear_ordered_semiring α] {a b c : α} (h : c * a c * b) (hc : 0 < c) :
a b

theorem le_of_mul_le_mul_right {α : Type u} [linear_ordered_semiring α] {a b c : α} (h : a * c b * c) (hc : 0 < c) :
a b

theorem pos_of_mul_pos_left {α : Type u} [linear_ordered_semiring α] {a b : α} (h : 0 < a * b) (h1 : 0 a) :
0 < b

theorem pos_of_mul_pos_right {α : Type u} [linear_ordered_semiring α] {a b : α} (h : 0 < a * b) (h1 : 0 b) :
0 < a

theorem nonneg_of_mul_nonneg_left {α : Type u} [linear_ordered_semiring α] {a b : α} (h : 0 a * b) (h1 : 0 < a) :
0 b

theorem nonneg_of_mul_nonneg_right {α : Type u} [linear_ordered_semiring α] {a b : α} (h : 0 a * b) (h1 : 0 < b) :
0 a

theorem neg_of_mul_neg_left {α : Type u} [linear_ordered_semiring α] {a b : α} (h : a * b < 0) (h1 : 0 a) :
b < 0

theorem neg_of_mul_neg_right {α : Type u} [linear_ordered_semiring α] {a b : α} (h : a * b < 0) (h1 : 0 b) :
a < 0

theorem nonpos_of_mul_nonpos_left {α : Type u} [linear_ordered_semiring α] {a b : α} (h : a * b 0) (h1 : 0 < a) :
b 0

theorem nonpos_of_mul_nonpos_right {α : Type u} [linear_ordered_semiring α] {a b : α} (h : a * b 0) (h1 : 0 < b) :
a 0

@[simp]
theorem mul_le_mul_left {α : Type u} [linear_ordered_semiring α] {a b c : α} (h : 0 < c) :
c * a c * b a b

@[simp]
theorem mul_le_mul_right {α : Type u} [linear_ordered_semiring α] {a b c : α} (h : 0 < c) :
a * c b * c a b

@[simp]
theorem mul_lt_mul_left {α : Type u} [linear_ordered_semiring α] {a b c : α} (h : 0 < c) :
c * a < c * b a < b

@[simp]
theorem mul_lt_mul_right {α : Type u} [linear_ordered_semiring α] {a b c : α} (h : 0 < c) :
a * c < b * c a < b

@[simp]
theorem zero_le_mul_left {α : Type u} [linear_ordered_semiring α] {b c : α} (h : 0 < c) :
0 c * b 0 b

@[simp]
theorem zero_le_mul_right {α : Type u} [linear_ordered_semiring α] {b c : α} (h : 0 < c) :
0 b * c 0 b

@[simp]
theorem zero_lt_mul_left {α : Type u} [linear_ordered_semiring α] {b c : α} (h : 0 < c) :
0 < c * b 0 < b

@[simp]
theorem zero_lt_mul_right {α : Type u} [linear_ordered_semiring α] {b c : α} (h : 0 < c) :
0 < b * c 0 < b

@[simp]
theorem bit0_le_bit0 {α : Type u} [linear_ordered_semiring α] {a b : α} :
bit0 a bit0 b a b

@[simp]
theorem bit0_lt_bit0 {α : Type u} [linear_ordered_semiring α] {a b : α} :
bit0 a < bit0 b a < b

@[simp]
theorem bit1_le_bit1 {α : Type u} [linear_ordered_semiring α] {a b : α} :
bit1 a bit1 b a b

@[simp]
theorem bit1_lt_bit1 {α : Type u} [linear_ordered_semiring α] {a b : α} :
bit1 a < bit1 b a < b

@[simp]
theorem one_le_bit1 {α : Type u} [linear_ordered_semiring α] {a : α} :
1 bit1 a 0 a

@[simp]
theorem one_lt_bit1 {α : Type u} [linear_ordered_semiring α] {a : α} :
1 < bit1 a 0 < a

@[simp]
theorem zero_le_bit0 {α : Type u} [linear_ordered_semiring α] {a : α} :
0 bit0 a 0 a

@[simp]
theorem zero_lt_bit0 {α : Type u} [linear_ordered_semiring α] {a : α} :
0 < bit0 a 0 < a

theorem mul_lt_mul'' {α : Type u} [linear_ordered_semiring α] {a b c d : α} (h1 : a < c) (h2 : b < d) (h3 : 0 a) (h4 : 0 b) :
a * b < c * d

theorem le_mul_iff_one_le_left {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 < b) :
b a * b 1 a

theorem lt_mul_iff_one_lt_left {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 < b) :
b < a * b 1 < a

theorem le_mul_iff_one_le_right {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 < b) :
b b * a 1 a

theorem lt_mul_iff_one_lt_right {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 < b) :
b < b * a 1 < a

theorem lt_mul_of_one_lt_right {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 < b) (a_1 : 1 < a) :
b < b * a

theorem le_mul_of_one_le_right {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 b) (h : 1 a) :
b b * a

theorem le_mul_of_one_le_left {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 b) (h : 1 a) :
b a * b

theorem mul_nonneg_iff_right_nonneg_of_pos {α : Type u} [linear_ordered_semiring α] {a b : α} (h : 0 < a) :
0 b * a 0 b

theorem bit1_pos {α : Type u} [linear_ordered_semiring α] {a : α} (h : 0 a) :
0 < bit1 a

theorem bit1_pos' {α : Type u} [linear_ordered_semiring α] {a : α} (h : 0 < a) :
0 < bit1 a

theorem lt_add_one {α : Type u} [linear_ordered_semiring α] (a : α) :
a < a + 1

theorem lt_one_add {α : Type u} [linear_ordered_semiring α] (a : α) :
a < 1 + a

theorem one_lt_mul {α : Type u} [linear_ordered_semiring α] {a b : α} (ha : 1 a) (hb : 1 < b) :
1 < a * b

theorem mul_le_one {α : Type u} [linear_ordered_semiring α] {a b : α} (ha : a 1) (hb' : 0 b) (hb : b 1) :
a * b 1

theorem one_lt_mul_of_le_of_lt {α : Type u} [linear_ordered_semiring α] {a b : α} (ha : 1 a) (hb : 1 < b) :
1 < a * b

theorem one_lt_mul_of_lt_of_le {α : Type u} [linear_ordered_semiring α] {a b : α} (ha : 1 < a) (hb : 1 b) :
1 < a * b

theorem mul_le_of_le_one_right {α : Type u} [linear_ordered_semiring α] {a b : α} (ha : 0 a) (hb1 : b 1) :
a * b a

theorem mul_le_of_le_one_left {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 b) (ha1 : a 1) :
a * b b

theorem mul_lt_one_of_nonneg_of_lt_one_left {α : Type u} [linear_ordered_semiring α] {a b : α} (ha0 : 0 a) (ha : a < 1) (hb : b 1) :
a * b < 1

theorem mul_lt_one_of_nonneg_of_lt_one_right {α : Type u} [linear_ordered_semiring α] {a b : α} (ha : a 1) (hb0 : 0 b) (hb : b < 1) :
a * b < 1

theorem mul_le_iff_le_one_left {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 < b) :
a * b b a 1

theorem mul_lt_iff_lt_one_left {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 < b) :
a * b < b a < 1

theorem mul_le_iff_le_one_right {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 < b) :
b * a b a 1

theorem mul_lt_iff_lt_one_right {α : Type u} [linear_ordered_semiring α] {a b : α} (hb : 0 < b) :
b * a < b a < 1

theorem nonpos_of_mul_nonneg_left {α : Type u} [linear_ordered_semiring α] {a b : α} (h : 0 a * b) (hb : b < 0) :
a 0

theorem nonpos_of_mul_nonneg_right {α : Type u} [linear_ordered_semiring α] {a b : α} (h : 0 a * b) (ha : a < 0) :
b 0

theorem neg_of_mul_pos_left {α : Type u} [linear_ordered_semiring α] {a b : α} (h : 0 < a * b) (hb : b 0) :
a < 0

theorem neg_of_mul_pos_right {α : Type u} [linear_ordered_semiring α] {a b : α} (h : 0 < a * b) (ha : a 0) :
b < 0

theorem monotone_mul_left_of_nonneg {α : Type u} [linear_ordered_semiring α] {a : α} (ha : 0 a) :
monotone (λ (x : α), a * x)

theorem monotone_mul_right_of_nonneg {α : Type u} [linear_ordered_semiring α] {a : α} (ha : 0 a) :
monotone (λ (x : α), x * a)

theorem monotone.mul_const {α : Type u} {β : Type u_1} [linear_ordered_semiring α] [preorder β] {f : β → α} {a : α} (hf : monotone f) (ha : 0 a) :
monotone (λ (x : β), (f x) * a)

theorem monotone.const_mul {α : Type u} {β : Type u_1} [linear_ordered_semiring α] [preorder β] {f : β → α} {a : α} (hf : monotone f) (ha : 0 a) :
monotone (λ (x : β), a * f x)

theorem monotone.mul {α : Type u} {β : Type u_1} [linear_ordered_semiring α] [preorder β] {f g : β → α} (hf : monotone f) (hg : monotone g) (hf0 : ∀ (x : β), 0 f x) (hg0 : ∀ (x : β), 0 g x) :
monotone (λ (x : β), (f x) * g x)

theorem strict_mono_mul_left_of_pos {α : Type u} [linear_ordered_semiring α] {a : α} (ha : 0 < a) :
strict_mono (λ (x : α), a * x)

theorem strict_mono_mul_right_of_pos {α : Type u} [linear_ordered_semiring α] {a : α} (ha : 0 < a) :
strict_mono (λ (x : α), x * a)

theorem strict_mono.mul_const {α : Type u} {β : Type u_1} [linear_ordered_semiring α] [preorder β] {f : β → α} {a : α} (hf : strict_mono f) (ha : 0 < a) :
strict_mono (λ (x : β), (f x) * a)

theorem strict_mono.const_mul {α : Type u} {β : Type u_1} [linear_ordered_semiring α] [preorder β] {f : β → α} {a : α} (hf : strict_mono f) (ha : 0 < a) :
strict_mono (λ (x : β), a * f x)

theorem strict_mono.mul_monotone {α : Type u} {β : Type u_1} [linear_ordered_semiring α] [preorder β] {f g : β → α} (hf : strict_mono f) (hg : monotone g) (hf0 : ∀ (x : β), 0 f x) (hg0 : ∀ (x : β), 0 < g x) :
strict_mono (λ (x : β), (f x) * g x)

theorem monotone.mul_strict_mono {α : Type u} {β : Type u_1} [linear_ordered_semiring α] [preorder β] {f g : β → α} (hf : monotone f) (hg : strict_mono g) (hf0 : ∀ (x : β), 0 < f x) (hg0 : ∀ (x : β), 0 g x) :
strict_mono (λ (x : β), (f x) * g x)

theorem strict_mono.mul {α : Type u} {β : Type u_1} [linear_ordered_semiring α] [preorder β] {f g : β → α} (hf : strict_mono f) (hg : strict_mono g) (hf0 : ∀ (x : β), 0 f x) (hg0 : ∀ (x : β), 0 g x) :
strict_mono (λ (x : β), (f x) * g x)

@[class]
structure decidable_linear_ordered_semiring (α : Type u) :
Type u
  • add : α → α → α
  • add_assoc : ∀ (a b c_1 : α), a + b + c_1 = a + (b + c_1)
  • zero : α
  • zero_add : ∀ (a : α), 0 + a = a
  • add_zero : ∀ (a : α), a + 0 = a
  • add_comm : ∀ (a b : α), a + b = b + a
  • mul : α → α → α
  • mul_assoc : ∀ (a b c_1 : α), (a * b) * c_1 = a * b * c_1
  • one : α
  • one_mul : ∀ (a : α), 1 * a = a
  • mul_one : ∀ (a : α), a * 1 = a
  • zero_mul : ∀ (a : α), 0 * a = 0
  • mul_zero : ∀ (a : α), a * 0 = 0
  • left_distrib : ∀ (a b c_1 : α), a * (b + c_1) = a * b + a * c_1
  • right_distrib : ∀ (a b c_1 : α), (a + b) * c_1 = a * c_1 + b * c_1
  • add_left_cancel : ∀ (a b c_1 : α), a + b = a + c_1b = c_1
  • add_right_cancel : ∀ (a b c_1 : α), a + b = c_1 + ba = c_1
  • le : α → α → Prop
  • lt : α → α → Prop
  • le_refl : ∀ (a : α), a a
  • le_trans : ∀ (a b c_1 : α), a bb c_1a c_1
  • lt_iff_le_not_le : (∀ (a b : α), a < b a b ¬b a) . "order_laws_tac"
  • le_antisymm : ∀ (a b : α), a bb aa = b
  • add_le_add_left : ∀ (a b : α), a b∀ (c_1 : α), c_1 + a c_1 + b
  • le_of_add_le_add_left : ∀ (a b c_1 : α), a + b a + c_1b c_1
  • mul_lt_mul_of_pos_left : ∀ (a b c_1 : α), a < b0 < c_1c_1 * a < c_1 * b
  • mul_lt_mul_of_pos_right : ∀ (a b c_1 : α), a < b0 < c_1a * c_1 < b * c_1
  • le_total : ∀ (a b : α), a b b a
  • zero_lt_one : decidable_linear_ordered_semiring.zero < decidable_linear_ordered_semiring.one
  • decidable_le : decidable_rel has_le.le
  • decidable_eq : decidable_eq α
  • decidable_lt : decidable_rel has_lt.lt

A decidable_linear_ordered_semiring α is a semiring α with a decidable linear order such that multiplication with a positive number and addition are monotone.

Instances
@[simp]
theorem decidable.mul_le_mul_left {α : Type u} [decidable_linear_ordered_semiring α] {a b c : α} (h : 0 < c) :
c * a c * b a b

@[simp]
theorem decidable.mul_le_mul_right {α : Type u} [decidable_linear_ordered_semiring α] {a b c : α} (h : 0 < c) :
a * c b * c a b

@[instance]
def ordered_ring.to_nontrivial (α : Type u) [s : ordered_ring α] :

@[class]
structure ordered_ring (α : Type u) :
Type u
  • add : α → α → α
  • add_assoc : ∀ (a b c_1 : α), a + b + c_1 = a + (b + c_1)
  • zero : α
  • zero_add : ∀ (a : α), 0 + a = a
  • add_zero : ∀ (a : α), a + 0 = a
  • neg : α → α
  • add_left_neg : ∀ (a : α), -a + a = 0
  • add_comm : ∀ (a b : α), a + b = b + a
  • mul : α → α → α
  • mul_assoc : ∀ (a b c_1 : α), (a * b) * c_1 = a * b * c_1
  • one : α
  • one_mul : ∀ (a : α), 1 * a = a
  • mul_one : ∀ (a : α), a * 1 = a
  • left_distrib : ∀ (a b c_1 : α), a * (b + c_1) = a * b + a * c_1
  • right_distrib : ∀ (a b c_1 : α), (a + b) * c_1 = a * c_1 + b * c_1
  • le : α → α → Prop
  • lt : α → α → Prop
  • le_refl : ∀ (a : α), a a
  • le_trans : ∀ (a b c_1 : α), a bb c_1a c_1
  • lt_iff_le_not_le : (∀ (a b : α), a < b a b ¬b a) . "order_laws_tac"
  • le_antisymm : ∀ (a b : α), a bb aa = b
  • add_le_add_left : ∀ (a b : α), a b∀ (c_1 : α), c_1 + a c_1 + b
  • exists_pair_ne : ∃ (x y : α), x y
  • mul_pos : ∀ (a b : α), 0 < a0 < b0 < a * b

An ordered_ring α is a ring α with a partial order such that multiplication with a positive number and addition are monotone.

Instances
@[instance]
def ordered_ring.to_ring (α : Type u) [s : ordered_ring α] :
ring α

theorem ordered_ring.mul_nonneg {α : Type u} [ordered_ring α] (a b : α) (h₁ : 0 a) (h₂ : 0 b) :
0 a * b

theorem ordered_ring.mul_le_mul_of_nonneg_left {α : Type u} [ordered_ring α] {a b c : α} (h₁ : a b) (h₂ : 0 c) :
c * a c * b

theorem ordered_ring.mul_le_mul_of_nonneg_right {α : Type u} [ordered_ring α] {a b c : α} (h₁ : a b) (h₂ : 0 c) :
a * c b * c

theorem ordered_ring.mul_lt_mul_of_pos_left {α : Type u} [ordered_ring α] {a b c : α} (h₁ : a < b) (h₂ : 0 < c) :
c * a < c * b

theorem ordered_ring.mul_lt_mul_of_pos_right {α : Type u} [ordered_ring α] {a b c : α} (h₁ : a < b) (h₂ : 0 < c) :
a * c < b * c

theorem mul_le_mul_of_nonpos_left {α : Type u} [ordered_ring α] {a b c : α} (h : b a) (hc : c 0) :
c * a c * b

theorem mul_le_mul_of_nonpos_right {α : Type u} [ordered_ring α] {a b c : α} (h : b a) (hc : c 0) :
a * c b * c

theorem mul_nonneg_of_nonpos_of_nonpos {α : Type u} [ordered_ring α] {a b : α} (ha : a 0) (hb : b 0) :
0 a * b

theorem mul_lt_mul_of_neg_left {α : Type u} [ordered_ring α] {a b c : α} (h : b < a) (hc : c < 0) :
c * a < c * b

theorem mul_lt_mul_of_neg_right {α : Type u} [ordered_ring α] {a b c : α} (h : b < a) (hc : c < 0) :
a * c < b * c

theorem mul_pos_of_neg_of_neg {α : Type u} [ordered_ring α] {a b : α} (ha : a < 0) (hb : b < 0) :
0 < a * b

@[instance]

@[class]
structure linear_ordered_ring (α : Type u) :
Type u
  • add : α → α → α
  • add_assoc : ∀ (a b c_1 : α), a + b + c_1 = a + (b + c_1)
  • zero : α
  • zero_add : ∀ (a : α), 0 + a = a
  • add_zero : ∀ (a : α), a + 0 = a
  • neg : α → α
  • add_left_neg : ∀ (a : α), -a + a = 0
  • add_comm : ∀ (a b : α), a + b = b + a
  • mul : α → α → α
  • mul_assoc : ∀ (a b c_1 : α), (a * b) * c_1 = a * b * c_1
  • one : α
  • one_mul : ∀ (a : α), 1 * a = a
  • mul_one : ∀ (a : α), a * 1 = a
  • left_distrib : ∀ (a b c_1 : α), a * (b + c_1) = a * b + a * c_1
  • right_distrib : ∀ (a b c_1 : α), (a + b) * c_1 = a * c_1 + b * c_1
  • le : α → α → Prop
  • lt : α → α → Prop
  • le_refl : ∀ (a : α), a a
  • le_trans : ∀ (a b c_1 : α), a bb c_1a c_1
  • lt_iff_le_not_le : (∀ (a b : α), a < b a b ¬b a) . "order_laws_tac"
  • le_antisymm : ∀ (a b : α), a bb aa = b
  • add_le_add_left : ∀ (a b : α), a b∀ (c_1 : α), c_1 + a c_1 + b
  • exists_pair_ne : ∃ (x y : α), x y
  • mul_pos : ∀ (a b : α), 0 < a0 < b0 < a * b
  • le_total : ∀ (a b : α), a b b a
  • zero_lt_one : linear_ordered_ring.zero < linear_ordered_ring.one

A linear_ordered_ring α is a ring α with a linear order such that multiplication with a positive number and addition are monotone.

Instances
@[instance]

theorem mul_self_nonneg {α : Type u} [linear_ordered_ring α] (a : α) :
0 a * a

theorem pos_and_pos_or_neg_and_neg_of_mul_pos {α : Type u} [linear_ordered_ring α] {a b : α} (hab : 0 < a * b) :
0 < a 0 < b a < 0 b < 0

theorem gt_of_mul_lt_mul_neg_left {α : Type u} [linear_ordered_ring α] {a b c : α} (h : c * a < c * b) (hc : c 0) :
b < a

theorem neg_one_lt_zero {α : Type u} [linear_ordered_ring α] :
-1 < 0

theorem le_of_mul_le_of_one_le {α : Type u} [linear_ordered_ring α] {a b c : α} (h : a * c b) (hb : 0 b) (hc : 1 c) :
a b

theorem nonneg_le_nonneg_of_squares_le {α : Type u} [linear_ordered_ring α] {a b : α} (hb : 0 b) (h : a * a b * b) :
a b

theorem mul_self_le_mul_self_iff {α : Type u} [linear_ordered_ring α] {a b : α} (h1 : 0 a) (h2 : 0 b) :
a b a * a b * b

theorem mul_self_lt_mul_self_iff {α : Type u} [linear_ordered_ring α] {a b : α} (h1 : 0 a) (h2 : 0 b) :
a < b a * a < b * b

theorem linear_ordered_ring.eq_zero_or_eq_zero_of_mul_eq_zero {α : Type u} [linear_ordered_ring α] {a b : α} (h : a * b = 0) :
a = 0 b = 0

@[simp]
theorem mul_le_mul_left_of_neg {α : Type u} [linear_ordered_ring α] {a b c : α} (h : c < 0) :
c * a c * b b a

@[simp]
theorem mul_le_mul_right_of_neg {α : Type u} [linear_ordered_ring α] {a b c : α} (h : c < 0) :
a * c b * c b a

@[simp]
theorem mul_lt_mul_left_of_neg {α : Type u} [linear_ordered_ring α] {a b c : α} (h : c < 0) :
c * a < c * b b < a

@[simp]
theorem mul_lt_mul_right_of_neg {α : Type u} [linear_ordered_ring α] {a b c : α} (h : c < 0) :
a * c < b * c b < a

theorem sub_one_lt {α : Type u} [linear_ordered_ring α] (a : α) :
a - 1 < a

theorem mul_self_pos {α : Type u} [linear_ordered_ring α] {a : α} (ha : a 0) :
0 < a * a

theorem mul_self_le_mul_self_of_le_of_neg_le {α : Type u} [linear_ordered_ring α] {x y : α} (h₁ : x y) (h₂ : -x y) :
x * x y * y

theorem nonneg_of_mul_nonpos_left {α : Type u} [linear_ordered_ring α] {a b : α} (h : a * b 0) (hb : b < 0) :
0 a

theorem nonneg_of_mul_nonpos_right {α : Type u} [linear_ordered_ring α] {a b : α} (h : a * b 0) (ha : a < 0) :
0 b

theorem pos_of_mul_neg_left {α : Type u} [linear_ordered_ring α] {a b : α} (h : a * b < 0) (hb : b 0) :
0 < a

theorem pos_of_mul_neg_right {α : Type u} [linear_ordered_ring α] {a b : α} (h : a * b < 0) (ha : a 0) :
0 < b

theorem mul_self_add_mul_self_eq_zero {α : Type u} [linear_ordered_ring α] {x y : α} :
x * x + y * y = 0 x = 0 y = 0

@[class]
structure linear_ordered_comm_ring (α : Type u) :
Type u
  • add : α → α → α
  • add_assoc : ∀ (a b c_1 : α), a + b + c_1 = a + (b + c_1)
  • zero : α
  • zero_add : ∀ (a : α), 0 + a = a
  • add_zero : ∀ (a : α), a + 0 = a
  • neg : α → α
  • add_left_neg : ∀ (a : α), -a + a = 0
  • add_comm : ∀ (a b : α), a + b = b + a
  • mul : α → α → α
  • mul_assoc : ∀ (a b c_1 : α), (a * b) * c_1 = a * b * c_1
  • one : α
  • one_mul : ∀ (a : α), 1 * a = a
  • mul_one : ∀ (a : α), a * 1 = a
  • left_distrib : ∀ (a b c_1 : α), a * (b + c_1) = a * b + a * c_1
  • right_distrib : ∀ (a b c_1 : α), (a + b) * c_1 = a * c_1 + b * c_1
  • le : α → α → Prop
  • lt : α → α → Prop
  • le_refl : ∀ (a : α), a a
  • le_trans : ∀ (a b c_1 : α), a bb c_1a c_1
  • lt_iff_le_not_le : (∀ (a b : α), a < b a b ¬b a) . "order_laws_tac"
  • le_antisymm : ∀ (a b : α), a bb aa = b
  • add_le_add_left : ∀ (a b : α), a b∀ (c_1 : α), c_1 + a c_1 + b
  • exists_pair_ne : ∃ (x y : α), x y
  • mul_pos : ∀ (a b : α), 0 < a0 < b0 < a * b
  • le_total : ∀ (a b : α), a b b a
  • zero_lt_one : linear_ordered_comm_ring.zero < linear_ordered_comm_ring.one
  • mul_comm : ∀ (a b : α), a * b = b * a

A linear_ordered_comm_ring α is a commutative ring α with a linear order such that multiplication with a positive number and addition are monotone.

Instances
@[class]
structure decidable_linear_ordered_comm_ring (α : Type u) :
Type u
  • add : α → α → α
  • add_assoc : ∀ (a b c_1 : α), a + b + c_1 = a + (b + c_1)
  • zero : α
  • zero_add : ∀ (a : α), 0 + a = a
  • add_zero : ∀ (a : α), a + 0 = a
  • neg : α → α
  • add_left_neg : ∀ (a : α), -a + a = 0
  • add_comm : ∀ (a b : α), a + b = b + a
  • mul : α → α → α
  • mul_assoc : ∀ (a b c_1 : α), (a * b) * c_1 = a * b * c_1
  • one : α
  • one_mul : ∀ (a : α), 1 * a = a
  • mul_one : ∀ (a : α), a * 1 = a
  • left_distrib : ∀ (a b c_1 : α), a * (b + c_1) = a * b + a * c_1
  • right_distrib : ∀ (a b c_1 : α), (a + b) * c_1 = a * c_1 + b * c_1
  • le : α → α → Prop
  • lt : α → α → Prop
  • le_refl : ∀ (a : α), a a
  • le_trans : ∀ (a b c_1 : α), a bb c_1a c_1
  • lt_iff_le_not_le : (∀ (a b : α), a < b a b ¬b a) . "order_laws_tac"
  • le_antisymm : ∀ (a b : α), a bb aa = b
  • add_le_add_left : ∀ (a b : α), a b∀ (c_1 : α), c_1 + a c_1 + b
  • exists_pair_ne : ∃ (x y : α), x y
  • mul_pos : ∀ (a b : α), 0 < a0 < b0 < a * b
  • le_total : ∀ (a b : α), a b b a
  • zero_lt_one : decidable_linear_ordered_comm_ring.zero < decidable_linear_ordered_comm_ring.one
  • mul_comm : ∀ (a b : α), a * b = b * a
  • decidable_le : decidable_rel has_le.le
  • decidable_eq : decidable_eq α
  • decidable_lt : decidable_rel has_lt.lt

A decidable_linear_ordered_comm_ring α is a commutative ring α with a decidable linear order such that multiplication with a positive number and addition are monotone.

Instances
theorem abs_mul {α : Type u} [decidable_linear_ordered_comm_ring α] (a b : α) :
abs (a * b) = (abs a) * abs b

theorem abs_mul_abs_self {α : Type u} [decidable_linear_ordered_comm_ring α] (a : α) :
(abs a) * abs a = a * a

theorem abs_mul_self {α : Type u} [decidable_linear_ordered_comm_ring α] (a : α) :
abs (a * a) = a * a

theorem sub_le_of_abs_sub_le_left {α : Type u} [decidable_linear_ordered_comm_ring α] {a b c : α} (h : abs (a - b) c) :
b - c a

theorem sub_le_of_abs_sub_le_right {α : Type u} [decidable_linear_ordered_comm_ring α] {a b c : α} (h : abs (a - b) c) :
a - c b

theorem sub_lt_of_abs_sub_lt_left {α : Type u} [decidable_linear_ordered_comm_ring α] {a b c : α} (h : abs (a - b) < c) :
b - c < a

theorem sub_lt_of_abs_sub_lt_right {α : Type u} [decidable_linear_ordered_comm_ring α] {a b c : α} (h : abs (a - b) < c) :
a - c < b

theorem abs_sub_square {α : Type u} [decidable_linear_ordered_comm_ring α] (a b : α) :
(abs (a - b)) * abs (a - b) = a * a + b * b - ((1 + 1) * a) * b

theorem eq_zero_of_mul_self_add_mul_self_eq_zero {α : Type u} [decidable_linear_ordered_comm_ring α] {a b : α} (h : a * a + b * b = 0) :
a = 0

theorem abs_abs_sub_abs_le_abs_sub {α : Type u} [decidable_linear_ordered_comm_ring α] (a b : α) :
abs (abs a - abs b) abs (a - b)

@[simp]
theorem abs_two {α : Type u} [decidable_linear_ordered_comm_ring α] :
abs 2 = 2

@[instance]
def nonneg_ring.to_ring (α : Type u_1) [s : nonneg_ring α] :
ring α

@[class]
structure nonneg_ring (α : Type u_1) :
Type u_1

Extend nonneg_add_comm_group to support ordered rings specified by their nonnegative elements

Instances
@[instance]

@[instance]
def nonneg_ring.to_nontrivial (α : Type u_1) [s : nonneg_ring α] :

@[class]
structure linear_nonneg_ring (α : Type u_1) :
Type u_1

Extend nonneg_add_comm_group to support linearly ordered rings specified by their nonnegative elements

@[instance]
def linear_nonneg_ring.to_domain (α : Type u_1) [s : linear_nonneg_ring α] :

def nonneg_ring.to_linear_nonneg_ring {α : Type u} [nonneg_ring α] (nonneg_total : ∀ (a : α), nonneg_ring.nonneg a nonneg_ring.nonneg (-a)) :

to_linear_nonneg_ring shows that a nonneg_ring with a total order is a domain, hence a linear_nonneg_ring.

Equations

Convert a linear_nonneg_ring with a commutative multiplication and decidable non-negativity into a decidable_linear_ordered_comm_ring

Equations
@[class]
structure canonically_ordered_comm_semiring (α : Type u_1) :
Type u_1
  • add : α → α → α
  • add_assoc : ∀ (a b c_1 : α), a + b + c_1 = a + (b + c_1)
  • zero : α
  • zero_add : ∀ (a : α), 0 + a = a
  • add_zero : ∀ (a : α), a + 0 = a
  • add_comm : ∀ (a b : α), a + b = b + a
  • le : α → α → Prop
  • lt : α → α → Prop
  • le_refl : ∀ (a : α), a a
  • le_trans : ∀ (a b c_1 : α), a bb c_1a c_1
  • lt_iff_le_not_le : (∀ (a b : α), a < b a b ¬b a) . "order_laws_tac"
  • le_antisymm : ∀ (a b : α), a bb aa = b
  • add_le_add_left : ∀ (a b : α), a b∀ (c_1 : α), c_1 + a c_1 + b
  • lt_of_add_lt_add_left : ∀ (a b c_1 : α), a + b < a + c_1b < c_1
  • bot : α
  • bot_le : ∀ (a : α), a
  • le_iff_exists_add : ∀ (a b : α), a b ∃ (c_1 : α), b = a + c_1
  • mul : α → α → α
  • mul_assoc : ∀ (a b c_1 : α), (a * b) * c_1 = a * b * c_1
  • one : α
  • one_mul : ∀ (a : α), 1 * a = a
  • mul_one : ∀ (a : α), a * 1 = a
  • zero_mul : ∀ (a : α), 0 * a = 0
  • mul_zero : ∀ (a : α), a * 0 = 0
  • left_distrib : ∀ (a b c_1 : α), a * (b + c_1) = a * b + a * c_1
  • right_distrib : ∀ (a b c_1 : α), (a + b) * c_1 = a * c_1 + b * c_1
  • mul_comm : ∀ (a b : α), a * b = b * a
  • exists_pair_ne : ∃ (x y : α), x y
  • eq_zero_or_eq_zero_of_mul_eq_zero : ∀ (a b : α), a * b = 0a = 0 b = 0

A canonically ordered commutative semiring is an ordered, commutative semiring in which a ≤ b iff there exists c with b = a + c. This is satisfied by the natural numbers, for example, but not the integers or other ordered groups.

Instances
theorem canonically_ordered_semiring.mul_le_mul {α : Type u} [canonically_ordered_comm_semiring α] {a b c d : α} (hab : a b) (hcd : c d) :
a * c b * d

theorem canonically_ordered_semiring.mul_le_mul_left' {α : Type u} [canonically_ordered_comm_semiring α] {b c : α} (h : b c) (a : α) :
a * b a * c

theorem canonically_ordered_semiring.mul_le_mul_right' {α : Type u} [canonically_ordered_comm_semiring α] {b c : α} (h : b c) (a : α) :
b * a c * a

theorem canonically_ordered_semiring.mul_pos {α : Type u} [canonically_ordered_comm_semiring α] {a b : α} :
0 < a * b 0 < a 0 < b

@[instance]
def with_top.nontrivial {α : Type u} [nonempty α] :

@[instance]
def with_top.mul_zero_class {α : Type u} [decidable_eq α] [has_zero α] [has_mul α] :

Equations
theorem with_top.mul_def {α : Type u} [decidable_eq α] [has_zero α] [has_mul α] {a b : with_top α} :
a * b = ite (a = 0 b = 0) 0 (option.bind a (λ (a : α), option.bind b (λ (b : α), a * b)))

@[simp]
theorem with_top.mul_top {α : Type u} [decidable_eq α] [has_zero α] [has_mul α] {a : with_top α} (h : a 0) :

@[simp]
theorem with_top.top_mul {α : Type u} [decidable_eq α] [has_zero α] [has_mul α] {a : with_top α} (h : a 0) :

@[simp]
theorem with_top.top_mul_top {α : Type u} [decidable_eq α] [has_zero α] [has_mul α] :

theorem with_top.coe_mul {α : Type u} [decidable_eq α] [mul_zero_class α] {a b : α} :
a * b = (a) * b

theorem with_top.mul_coe {α : Type u} [decidable_eq α] [mul_zero_class α] {b : α} (hb : b 0) {a : with_top α} :
a * b = option.bind a (λ (a : α), a * b)

@[simp]
theorem with_top.mul_eq_top_iff {α : Type u} [decidable_eq α] [mul_zero_class α] {a b : with_top α} :
a * b = a 0 b = a = b 0