mathlib3 documentation

algebra.order.sub.canonical

Lemmas about subtraction in canonically ordered monoids #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

@[simp]
theorem add_tsub_cancel_of_le {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b : α} (h : a b) :
a + (b - a) = b
theorem tsub_add_cancel_of_le {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b : α} (h : a b) :
b - a + a = b
theorem add_le_of_le_tsub_right_of_le {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (h : b c) (h2 : a c - b) :
a + b c
theorem add_le_of_le_tsub_left_of_le {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (h : a c) (h2 : b c - a) :
a + b c
theorem tsub_le_tsub_iff_right {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (h : c b) :
a - c b - c a b
theorem tsub_left_inj {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (h1 : c a) (h2 : c b) :
a - c = b - c a = b
theorem tsub_inj_left {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (h₁ : a b) (h₂ : a c) :
b - a = c - a b = c
theorem lt_of_tsub_lt_tsub_right_of_le {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (h : c b) (h2 : a - c < b - c) :
a < b

See lt_of_tsub_lt_tsub_right for a stronger statement in a linear order.

theorem tsub_add_tsub_cancel {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (hab : b a) (hcb : c b) :
a - b + (b - c) = a - c
theorem tsub_tsub_tsub_cancel_right {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (h : c b) :
a - c - (b - c) = a - b

Lemmas that assume that an element is add_le_cancellable. #

@[protected]
@[protected]
@[protected]
theorem add_le_cancellable.add_tsub_assoc_of_le {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {b c : α} (hc : add_le_cancellable c) (h : c b) (a : α) :
a + b - c = a + (b - c)
@[protected]
@[protected]
theorem add_le_cancellable.tsub_tsub_assoc {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (hbc : add_le_cancellable (b - c)) (h₁ : b a) (h₂ : c b) :
a - (b - c) = a - b + c
@[protected]
theorem add_le_cancellable.tsub_add_tsub_comm {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c d : α} (hb : add_le_cancellable b) (hd : add_le_cancellable d) (hba : b a) (hdc : d c) :
a - b + (c - d) = a + c - (b + d)
@[protected]
@[protected]
@[protected]
theorem add_le_cancellable.tsub_lt_iff_left {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (hb : add_le_cancellable b) (hba : b a) :
a - b < c a < b + c
@[protected]
theorem add_le_cancellable.tsub_lt_iff_right {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (hb : add_le_cancellable b) (hba : b a) :
a - b < c a < c + b
@[protected]
theorem add_le_cancellable.tsub_lt_iff_tsub_lt {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (hb : add_le_cancellable b) (hc : add_le_cancellable c) (h₁ : b a) (h₂ : c a) :
a - b < c a - c < b
@[protected]
theorem add_le_cancellable.le_tsub_iff_le_tsub {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (ha : add_le_cancellable a) (hc : add_le_cancellable c) (h₁ : a b) (h₂ : c b) :
a b - c c b - a
@[protected]
@[protected]
@[protected]
theorem add_le_cancellable.tsub_inj_right {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (hab : add_le_cancellable (a - b)) (h₁ : b a) (h₂ : c a) (h₃ : a - b = a - c) :
b = c
@[protected]
theorem add_le_cancellable.tsub_lt_tsub_left_of_le {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (hab : add_le_cancellable (a - b)) (h₁ : b a) (h : c < b) :
a - b < a - c
@[protected]
theorem add_le_cancellable.tsub_lt_tsub_right_of_le {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (hc : add_le_cancellable c) (h : c a) (h2 : a < b) :
a - c < b - c
@[protected]
@[protected, simp]
theorem add_le_cancellable.add_tsub_tsub_cancel {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (hac : add_le_cancellable (a - c)) (h : c a) :
a + b - (a - c) = b + c
@[protected]
@[protected]
theorem add_le_cancellable.tsub_tsub_tsub_cancel_left {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} (hab : add_le_cancellable (a - b)) (h : b a) :
a - c - (a - b) = b - c

Lemmas where addition is order-reflecting. #

theorem add_tsub_assoc_of_le {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {b c : α} [contravariant_class α α has_add.add has_le.le] (h : c b) (a : α) :
a + b - c = a + (b - c)

See add_tsub_le_assoc for an inequality.

theorem tsub_tsub_assoc {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} [contravariant_class α α has_add.add has_le.le] (h₁ : b a) (h₂ : c b) :
a - (b - c) = a - b + c
theorem tsub_add_tsub_comm {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c d : α} [contravariant_class α α has_add.add has_le.le] (hba : b a) (hdc : d c) :
a - b + (c - d) = a + c - (b + d)
theorem tsub_lt_iff_tsub_lt {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} [contravariant_class α α has_add.add has_le.le] (h₁ : b a) (h₂ : c a) :
a - b < c a - c < b
theorem le_tsub_iff_le_tsub {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} [contravariant_class α α has_add.add has_le.le] (h₁ : a b) (h₂ : c b) :
a b - c c b - a

See lt_tsub_iff_right for a stronger statement in a linear order.

See lt_tsub_iff_left for a stronger statement in a linear order.

See lt_of_tsub_lt_tsub_left for a stronger statement in a linear order.

theorem tsub_inj_right {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} [contravariant_class α α has_add.add has_le.le] (h₁ : b a) (h₂ : c a) (h₃ : a - b = a - c) :
b = c

See tsub_lt_tsub_iff_left_of_le for a stronger statement in a linear order.

@[simp]
theorem add_tsub_tsub_cancel {α : Type u_1} [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] [covariant_class α α has_add.add has_le.le] [has_sub α] [has_ordered_sub α] {a b c : α} [contravariant_class α α has_add.add has_le.le] (h : c a) :
a + b - (a - c) = b + c

See tsub_tsub_le for an inequality.

Lemmas in a canonically ordered monoid. #

theorem add_tsub_cancel_iff_le {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
a + (b - a) = b a b
theorem tsub_add_cancel_iff_le {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
b - a + a = b a b
@[simp]
theorem tsub_eq_zero_iff_le {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
a - b = 0 a b
@[simp]
theorem tsub_eq_zero_of_le {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
a b a - b = 0

Alias of the reverse direction of tsub_eq_zero_iff_le.

@[simp]
theorem tsub_self {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] (a : α) :
a - a = 0
@[simp]
theorem tsub_le_self {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
a - b a
@[simp]
theorem zero_tsub {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] (a : α) :
0 - a = 0
theorem tsub_self_add {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] (a b : α) :
a - (a + b) = 0
theorem tsub_pos_iff_not_le {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
0 < a - b ¬a b
theorem tsub_pos_of_lt {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} (h : a < b) :
0 < b - a
theorem tsub_lt_of_lt {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} (h : a < b) :
a - c < b
@[protected]
theorem add_le_cancellable.tsub_le_tsub_iff_left {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} (ha : add_le_cancellable a) (hc : add_le_cancellable c) (h : c a) :
a - b a - c c b
@[protected]
theorem add_le_cancellable.tsub_right_inj {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} (ha : add_le_cancellable a) (hb : add_le_cancellable b) (hc : add_le_cancellable c) (hba : b a) (hca : c a) :
a - b = a - c b = c

Lemmas where addition is order-reflecting. #

theorem tsub_le_tsub_iff_left {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} [contravariant_class α α has_add.add has_le.le] (h : c a) :
a - b a - c c b
theorem tsub_right_inj {α : Type u_1} [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} [contravariant_class α α has_add.add has_le.le] (hba : b a) (hca : c a) :
a - b = a - c b = c

Lemmas in a linearly canonically ordered monoid. #

@[simp]
theorem tsub_pos_iff_lt {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
0 < a - b b < a
@[protected]
theorem add_le_cancellable.lt_tsub_iff_right {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} (hc : add_le_cancellable c) :
a < b - c a + c < b
@[protected]
theorem add_le_cancellable.lt_tsub_iff_left {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} (hc : add_le_cancellable c) :
a < b - c c + a < b
@[protected]
theorem add_le_cancellable.tsub_lt_tsub_iff_right {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} (hc : add_le_cancellable c) (h : c a) :
a - c < b - c a < b
@[protected]
theorem add_le_cancellable.tsub_lt_self {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} (ha : add_le_cancellable a) (h₁ : 0 < a) (h₂ : 0 < b) :
a - b < a
@[protected]
theorem add_le_cancellable.tsub_lt_self_iff {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} (ha : add_le_cancellable a) :
a - b < a 0 < a 0 < b
@[protected]
theorem add_le_cancellable.tsub_lt_tsub_iff_left_of_le {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} (ha : add_le_cancellable a) (hb : add_le_cancellable b) (h : b a) :
a - b < a - c c < b

See lt_tsub_iff_left_of_le_of_le for a weaker statement in a partial order.

theorem tsub_lt_tsub_iff_right {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} [contravariant_class α α has_add.add has_le.le] (h : c a) :
a - c < b - c a < b

This lemma also holds for ennreal, but we need a different proof for that.

theorem tsub_lt_tsub_iff_left_of_le {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b c : α} [contravariant_class α α has_add.add has_le.le] (h : b a) :
a - b < a - c c < b

See lt_tsub_iff_left_of_le_of_le for a weaker statement in a partial order.

Lemmas about max and min. #

theorem tsub_add_eq_max {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
a - b + b = linear_order.max a b
theorem add_tsub_eq_max {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
a + (b - a) = linear_order.max a b
theorem tsub_min {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
a - linear_order.min a b = a - b
theorem tsub_add_min {α : Type u_1} [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] {a b : α} :
a - b + linear_order.min a b = a