mathlib3 documentation

algebra.order.group.inj_surj

Pull back ordered groups along injective maps. #

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

@[reducible]
def function.injective.ordered_add_comm_group {α : Type u_1} [ordered_add_comm_group α] {β : Type u_2} [has_zero β] [has_add β] [has_neg β] [has_sub β] [has_smul β] [has_smul β] (f : β α) (hf : function.injective f) (one : f 0 = 0) (mul : (x y : β), f (x + y) = f x + f y) (inv : (x : β), f (-x) = -f x) (div : (x y : β), f (x - y) = f x - f y) (npow : (x : β) (n : ), f (n x) = n f x) (zpow : (x : β) (n : ), f (n x) = n f x) :

Pullback an ordered_add_comm_group under an injective map.

Equations
@[reducible]
def function.injective.ordered_comm_group {α : Type u_1} [ordered_comm_group α] {β : Type u_2} [has_one β] [has_mul β] [has_inv β] [has_div β] [has_pow β ] [has_pow β ] (f : β α) (hf : function.injective f) (one : f 1 = 1) (mul : (x y : β), f (x * y) = f x * f y) (inv : (x : β), f x⁻¹ = (f x)⁻¹) (div : (x y : β), f (x / y) = f x / f y) (npow : (x : β) (n : ), f (x ^ n) = f x ^ n) (zpow : (x : β) (n : ), f (x ^ n) = f x ^ n) :

Pullback an ordered_comm_group under an injective map. See note [reducible non-instances].

Equations
@[reducible]
def function.injective.linear_ordered_comm_group {α : Type u_1} [linear_ordered_comm_group α] {β : Type u_2} [has_one β] [has_mul β] [has_inv β] [has_div β] [has_pow β ] [has_pow β ] [has_sup β] [has_inf β] (f : β α) (hf : function.injective f) (one : f 1 = 1) (mul : (x y : β), f (x * y) = f x * f y) (inv : (x : β), f x⁻¹ = (f x)⁻¹) (div : (x y : β), f (x / y) = f x / f y) (npow : (x : β) (n : ), f (x ^ n) = f x ^ n) (zpow : (x : β) (n : ), f (x ^ n) = f x ^ n) (hsup : (x y : β), f (x y) = linear_order.max (f x) (f y)) (hinf : (x y : β), f (x y) = linear_order.min (f x) (f y)) :

Pullback a linear_ordered_comm_group under an injective map. See note [reducible non-instances].

Equations
@[reducible]
def function.injective.linear_ordered_add_comm_group {α : Type u_1} [linear_ordered_add_comm_group α] {β : Type u_2} [has_zero β] [has_add β] [has_neg β] [has_sub β] [has_smul β] [has_smul β] [has_sup β] [has_inf β] (f : β α) (hf : function.injective f) (one : f 0 = 0) (mul : (x y : β), f (x + y) = f x + f y) (inv : (x : β), f (-x) = -f x) (div : (x y : β), f (x - y) = f x - f y) (npow : (x : β) (n : ), f (n x) = n f x) (zpow : (x : β) (n : ), f (n x) = n f x) (hsup : (x y : β), f (x y) = linear_order.max (f x) (f y)) (hinf : (x y : β), f (x y) = linear_order.min (f x) (f y)) :

Pullback a linear_ordered_add_comm_group under an injective map.

Equations