Documentation

Mathlib.Algebra.Order.Group.InjSurj

Pull back ordered groups along injective maps. #

def Function.Injective.orderedAddCommGroup {α : Type u_1} [inst : OrderedAddCommGroup α] {β : Type u_2} [inst : Zero β] [inst : Add β] [inst : Neg β] [inst : Sub β] [inst : SMul β] [inst : 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 OrderedAddCommGroup under an injective map.

Equations
  • One or more equations did not get rendered due to their size.
def Function.Injective.orderedAddCommGroup.proof_1 {α : Type u_2} [inst : OrderedAddCommGroup α] {β : Type u_1} [inst : Zero β] [inst : Add β] [inst : SMul β] (f : βα) (hf : Function.Injective f) (one : f 0 = 0) (mul : ∀ (x y : β), f (x + y) = f x + f y) (npow : ∀ (x : β) (n : ), f (n x) = n f x) (a : β) (b : β) :
a + b = b + a
Equations
  • (_ : ∀ (a b : β), a + b = b + a) = (_ : ∀ (a b : β), a + b = b + a)
def Function.Injective.orderedCommGroup {α : Type u_1} [inst : OrderedCommGroup α] {β : Type u_2} [inst : One β] [inst : Mul β] [inst : Inv β] [inst : Div β] [inst : Pow β ] [inst : 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 OrderedCommGroup under an injective map. See note [reducible non-instances].

Equations
  • One or more equations did not get rendered due to their size.
def Function.Injective.linearOrderedAddCommGroup {α : Type u_1} [inst : LinearOrderedAddCommGroup α] {β : Type u_2} [inst : Zero β] [inst : Add β] [inst : Neg β] [inst : Sub β] [inst : SMul β] [inst : SMul β] [inst : Sup β] [inst : 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) = max (f x) (f y)) (hinf : ∀ (x y : β), f (x y) = min (f x) (f y)) :

Pullback a LinearOrderedAddCommGroup under an injective map.

Equations
  • One or more equations did not get rendered due to their size.
def Function.Injective.linearOrderedCommGroup {α : Type u_1} [inst : LinearOrderedCommGroup α] {β : Type u_2} [inst : One β] [inst : Mul β] [inst : Inv β] [inst : Div β] [inst : Pow β ] [inst : Pow β ] [inst : Sup β] [inst : 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) = max (f x) (f y)) (hinf : ∀ (x y : β), f (x y) = min (f x) (f y)) :

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

Equations
  • One or more equations did not get rendered due to their size.