scilib documentation

algebra.order.group.order_iso

Inverse and multiplication as order isomorphisms in ordered groups #

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

x ↦ x⁻¹ as an order-reversing equivalence.

Equations

x ↦ -x as an order-reversing equivalence.

Equations
theorem neg_le {α : Type u} [add_group α] [has_le α] [covariant_class α α has_add.add has_le.le] [covariant_class α α (function.swap has_add.add) has_le.le] {a b : α} :
-a b -b a
theorem inv_le' {α : Type u} [group α] [has_le α] [covariant_class α α has_mul.mul has_le.le] [covariant_class α α (function.swap has_mul.mul) has_le.le] {a b : α} :
theorem inv_le_of_inv_le' {α : Type u} [group α] [has_le α] [covariant_class α α has_mul.mul has_le.le] [covariant_class α α (function.swap has_mul.mul) has_le.le] {a b : α} :
a⁻¹ b b⁻¹ a

Alias of the forward direction of inv_le'.

theorem neg_le_of_neg_le {α : Type u} [add_group α] [has_le α] [covariant_class α α has_add.add has_le.le] [covariant_class α α (function.swap has_add.add) has_le.le] {a b : α} :
-a b -b a

Alias of the forward direction of inv_le'.

theorem le_neg {α : Type u} [add_group α] [has_le α] [covariant_class α α has_add.add has_le.le] [covariant_class α α (function.swap has_add.add) has_le.le] {a b : α} :
a -b b -a
theorem le_inv' {α : Type u} [group α] [has_le α] [covariant_class α α has_mul.mul has_le.le] [covariant_class α α (function.swap has_mul.mul) has_le.le] {a b : α} :

x ↦ a / x as an order-reversing equivalence.

Equations

x ↦ a - x as an order-reversing equivalence.

Equations
theorem le_inv_of_le_inv {α : Type u} [group α] [has_le α] [covariant_class α α has_mul.mul has_le.le] [covariant_class α α (function.swap has_mul.mul) has_le.le] {a b : α} :
a b⁻¹ b a⁻¹

Alias of the forward direction of le_inv'.

theorem le_neg_of_le_neg {α : Type u} [add_group α] [has_le α] [covariant_class α α has_add.add has_le.le] [covariant_class α α (function.swap has_add.add) has_le.le] {a b : α} :
a -b b -a

Alias of the forward direction of le_inv'.

@[simp]
theorem order_iso.add_right_apply {α : Type u} [add_group α] [has_le α] [covariant_class α α (function.swap has_add.add) has_le.le] (a x : α) :
@[simp]
theorem order_iso.mul_right_apply {α : Type u} [group α] [has_le α] [covariant_class α α (function.swap has_mul.mul) has_le.le] (a x : α) :
@[simp]
def order_iso.div_right {α : Type u} [group α] [has_le α] [covariant_class α α (function.swap has_mul.mul) has_le.le] (a : α) :
α ≃o α

x ↦ x / a as an order isomorphism.

Equations
@[simp]
theorem order_iso.sub_right_apply {α : Type u} [add_group α] [has_le α] [covariant_class α α (function.swap has_add.add) has_le.le] (a b : α) :
@[simp]
theorem order_iso.div_right_apply {α : Type u} [group α] [has_le α] [covariant_class α α (function.swap has_mul.mul) has_le.le] (a b : α) :
def order_iso.sub_right {α : Type u} [add_group α] [has_le α] [covariant_class α α (function.swap has_add.add) has_le.le] (a : α) :
α ≃o α

x ↦ x - a as an order isomorphism.

Equations
@[simp]
theorem order_iso.add_left_apply {α : Type u} [add_group α] [has_le α] [covariant_class α α has_add.add has_le.le] (a x : α) :
@[simp]
theorem order_iso.mul_left_apply {α : Type u} [group α] [has_le α] [covariant_class α α has_mul.mul has_le.le] (a x : α) :