Documentation

Mathlib.Algebra.Order.Monoid.Unbundled.Defs

Covariants and contravariants #

This file contains general lemmas and instances to work with the interactions between a relation and an action on a Type.

The intended application is the splitting of the ordering from the algebraic assumptions on the operations in the Ordered[...] hierarchy.

The strategy is to introduce two more flexible typeclasses, CovariantClass and ContravariantClass:

Since Co(ntra)variantClass takes as input the operation (typically (+) or (*)) and the order relation (typically (≤) or (<)), these are the only two typeclasses that I have used.

The general approach is to formulate the lemma that you are interested in and prove it, with the Ordered[...] typeclass of your liking. After that, you convert the single typeclass, say [OrderedCancelMonoid M], into three typeclasses, e.g. [CancelMonoid M] [PartialOrder M] [CovariantClass M M (Function.swap (*)) (≤)] and have a go at seeing if the proof still works!

Note that it is possible to combine several Co(ntra)variantClass assumptions together. Indeed, the usual ordered typeclasses arise from assuming the pair [CovariantClass M M (*) (≤)] [ContravariantClass M M (*) (<)] on top of order/algebraic assumptions.

A formal remark is that normally CovariantClass uses the (≤)-relation, while ContravariantClass uses the (<)-relation. This need not be the case in general, but seems to be the most common usage. In the opposite direction, the implication

[Semigroup α] [PartialOrder α] [ContravariantClass α α (*) (≤)] → LeftCancelSemigroup α

holds -- note the Co*ntra* assumption on the (≤)-relation.

Formalization notes #

We stick to the convention of using Function.swap (*) (or Function.swap (+)), for the typeclass assumptions, since Function.swap is slightly better behaved than flip. However, sometimes as a non-typeclass assumption, we prefer flip (*) (or flip (+)), as it is easier to use.

def Covariant (M : Type u_1) (N : Type u_2) (μ : MNN) (r : NNProp) :

Covariant is useful to formulate succinctly statements about the interactions between an action of a Type on another one and a relation on the acted-upon Type.

See the CovariantClass doc-string for its meaning.

Equations
  • Covariant M N μ r = ∀ (m : M) {n₁ n₂ : N}, r n₁ n₂r (μ m n₁) (μ m n₂)
def Contravariant (M : Type u_1) (N : Type u_2) (μ : MNN) (r : NNProp) :

Contravariant is useful to formulate succinctly statements about the interactions between an action of a Type on another one and a relation on the acted-upon Type.

See the ContravariantClass doc-string for its meaning.

Equations
  • Contravariant M N μ r = ∀ (m : M) {n₁ n₂ : N}, r (μ m n₁) (μ m n₂)r n₁ n₂
class CovariantClass (M : Type u_1) (N : Type u_2) (μ : MNN) (r : NNProp) :

Given an action μ of a Type M on a Type N and a relation r on N, informally, the CovariantClass says that "the action μ preserves the relation r."

More precisely, the CovariantClass is a class taking two Types M N, together with an "action" μ : M → N → N and a relation r : N → N → Prop. Its unique field elim is the assertion that for all m ∈ M and all elements n₁, n₂ ∈ N, if the relation r holds for the pair (n₁, n₂), then, the relation r also holds for the pair (μ m n₁, μ m n₂), obtained from (n₁, n₂) by acting upon it by m.

If m : M and h : r n₁ n₂, then CovariantClass.elim m h : r (μ m n₁) (μ m n₂).

  • elim : Covariant M N μ r

    For all m ∈ M and all elements n₁, n₂ ∈ N, if the relation r holds for the pair (n₁, n₂), then, the relation r also holds for the pair (μ m n₁, μ m n₂)

Instances
class ContravariantClass (M : Type u_1) (N : Type u_2) (μ : MNN) (r : NNProp) :

Given an action μ of a Type M on a Type N and a relation r on N, informally, the ContravariantClass says that "if the result of the action μ on a pair satisfies the relation r, then the initial pair satisfied the relation r."

More precisely, the ContravariantClass is a class taking two Types M N, together with an "action" μ : M → N → N and a relation r : N → N → Prop. Its unique field elim is the assertion that for all m ∈ M and all elements n₁, n₂ ∈ N, if the relation r holds for the pair (μ m n₁, μ m n₂) obtained from (n₁, n₂) by acting upon it by m, then, the relation r also holds for the pair (n₁, n₂).

If m : M and h : r (μ m n₁) (μ m n₂), then ContravariantClass.elim m h : r n₁ n₂.

  • elim : Contravariant M N μ r

    For all m ∈ M and all elements n₁, n₂ ∈ N, if the relation r holds for the pair (μ m n₁, μ m n₂) obtained from (n₁, n₂) by acting upon it by m, then, the relation r also holds for the pair (n₁, n₂).

Instances
@[reducible, inline]
abbrev MulLeftMono (M : Type u_1) [Mul M] [LE M] :

Typeclass for monotonicity of multiplication on the left, namely b₁ ≤ b₂ → a * b₁ ≤ a * b₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedCommMonoid.

Equations
@[reducible, inline]
abbrev MulRightMono (M : Type u_1) [Mul M] [LE M] :

Typeclass for monotonicity of multiplication on the right, namely a₁ ≤ a₂ → a₁ * b ≤ a₂ * b.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedCommMonoid.

Equations
@[reducible, inline]
abbrev AddLeftMono (M : Type u_1) [Add M] [LE M] :

Typeclass for monotonicity of addition on the left, namely b₁ ≤ b₂ → a + b₁ ≤ a + b₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedAddCommMonoid.

Equations
@[reducible, inline]
abbrev AddRightMono (M : Type u_1) [Add M] [LE M] :

Typeclass for monotonicity of addition on the right, namely a₁ ≤ a₂ → a₁ + b ≤ a₂ + b.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedAddCommMonoid.

Equations
@[reducible, inline]
abbrev MulLeftStrictMono (M : Type u_1) [Mul M] [LT M] :

Typeclass for monotonicity of multiplication on the left, namely b₁ < b₂ → a * b₁ < a * b₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedCommGroup.

Equations
@[reducible, inline]
abbrev MulRightStrictMono (M : Type u_1) [Mul M] [LT M] :

Typeclass for monotonicity of multiplication on the right, namely a₁ < a₂ → a₁ * b < a₂ * b.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedCommGroup.

Equations
@[reducible, inline]
abbrev AddLeftStrictMono (M : Type u_1) [Add M] [LT M] :

Typeclass for monotonicity of addition on the left, namely b₁ < b₂ → a + b₁ < a + b₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedAddCommGroup.

Equations
@[reducible, inline]
abbrev AddRightStrictMono (M : Type u_1) [Add M] [LT M] :

Typeclass for monotonicity of addition on the right, namely a₁ < a₂ → a₁ + b < a₂ + b.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedAddCommGroup.

Equations
@[reducible, inline]
abbrev MulLeftReflectLT (M : Type u_1) [Mul M] [LT M] :

Typeclass for strict reverse monotonicity of multiplication on the left, namely a * b₁ < a * b₂ → b₁ < b₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedCommGroup.

Equations
@[reducible, inline]
abbrev MulRightReflectLT (M : Type u_1) [Mul M] [LT M] :

Typeclass for strict reverse monotonicity of multiplication on the right, namely a₁ * b < a₂ * b → a₁ < a₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedCommGroup.

Equations
@[reducible, inline]
abbrev AddLeftReflectLT (M : Type u_1) [Add M] [LT M] :

Typeclass for strict reverse monotonicity of addition on the left, namely a + b₁ < a + b₂ → b₁ < b₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedAddCommGroup.

Equations
@[reducible, inline]
abbrev AddRightReflectLT (M : Type u_1) [Add M] [LT M] :

Typeclass for strict reverse monotonicity of addition on the right, namely a₁ * b < a₂ * b → a₁ < a₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedAddCommGroup.

Equations
@[reducible, inline]
abbrev MulLeftReflectLE (M : Type u_1) [Mul M] [LE M] :

Typeclass for reverse monotonicity of multiplication on the left, namely a * b₁ ≤ a * b₂ → b₁ ≤ b₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedCancelCommMonoid.

Equations
@[reducible, inline]
abbrev MulRightReflectLE (M : Type u_1) [Mul M] [LE M] :

Typeclass for reverse monotonicity of multiplication on the right, namely a₁ * b ≤ a₂ * b → a₁ ≤ a₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedCancelCommMonoid.

Equations
@[reducible, inline]
abbrev AddLeftReflectLE (M : Type u_1) [Add M] [LE M] :

Typeclass for reverse monotonicity of addition on the left, namely a + b₁ ≤ a + b₂ → b₁ ≤ b₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedCancelAddCommMonoid.

Equations
@[reducible, inline]
abbrev AddRightReflectLE (M : Type u_1) [Add M] [LE M] :

Typeclass for reverse monotonicity of addition on the right, namely a₁ + b ≤ a₂ + b → a₁ ≤ a₂.

You should usually not use this very granular typeclass directly, but rather a typeclass like OrderedCancelAddCommMonoid.

Equations
theorem rel_iff_cov (M : Type u_1) (N : Type u_2) (μ : MNN) (r : NNProp) [CovariantClass M N μ r] [ContravariantClass M N μ r] (m : M) {a b : N} :
r (μ m a) (μ m b) r a b
theorem Covariant.flip {M : Type u_1} {N : Type u_2} {μ : MNN} {r : NNProp} (h : Covariant M N μ r) :
theorem Contravariant.flip {M : Type u_1} {N : Type u_2} {μ : MNN} {r : NNProp} (h : Contravariant M N μ r) :
theorem act_rel_act_of_rel {M : Type u_1} {N : Type u_2} {μ : MNN} {r : NNProp} [CovariantClass M N μ r] (m : M) {a b : N} (ab : r a b) :
r (μ m a) (μ m b)
theorem Group.covariant_iff_contravariant {N : Type u_2} {r : NNProp} [Group N] :
Covariant N N (fun (x1 x2 : N) => x1 * x2) r Contravariant N N (fun (x1 x2 : N) => x1 * x2) r
theorem AddGroup.covariant_iff_contravariant {N : Type u_2} {r : NNProp} [AddGroup N] :
Covariant N N (fun (x1 x2 : N) => x1 + x2) r Contravariant N N (fun (x1 x2 : N) => x1 + x2) r
@[instance 100]
instance Group.covconv {N : Type u_2} {r : NNProp} [Group N] [CovariantClass N N (fun (x1 x2 : N) => x1 * x2) r] :
ContravariantClass N N (fun (x1 x2 : N) => x1 * x2) r
@[instance 100]
instance AddGroup.covconv {N : Type u_2} {r : NNProp} [AddGroup N] [CovariantClass N N (fun (x1 x2 : N) => x1 + x2) r] :
ContravariantClass N N (fun (x1 x2 : N) => x1 + x2) r
theorem Group.covariant_swap_iff_contravariant_swap {N : Type u_2} {r : NNProp} [Group N] :
Covariant N N (Function.swap fun (x1 x2 : N) => x1 * x2) r Contravariant N N (Function.swap fun (x1 x2 : N) => x1 * x2) r
theorem AddGroup.covariant_swap_iff_contravariant_swap {N : Type u_2} {r : NNProp} [AddGroup N] :
Covariant N N (Function.swap fun (x1 x2 : N) => x1 + x2) r Contravariant N N (Function.swap fun (x1 x2 : N) => x1 + x2) r
@[instance 100]
instance Group.covconv_swap {N : Type u_2} {r : NNProp} [Group N] [CovariantClass N N (Function.swap fun (x1 x2 : N) => x1 * x2) r] :
ContravariantClass N N (Function.swap fun (x1 x2 : N) => x1 * x2) r
@[instance 100]
instance AddGroup.covconv_swap {N : Type u_2} {r : NNProp} [AddGroup N] [CovariantClass N N (Function.swap fun (x1 x2 : N) => x1 + x2) r] :
ContravariantClass N N (Function.swap fun (x1 x2 : N) => x1 + x2) r
theorem act_rel_of_rel_of_act_rel {M : Type u_1} {N : Type u_2} {μ : MNN} {r : NNProp} [CovariantClass M N μ r] [IsTrans N r] (m : M) {a b c : N} (ab : r a b) (rl : r (μ m b) c) :
r (μ m a) c
theorem rel_act_of_rel_of_rel_act {M : Type u_1} {N : Type u_2} {μ : MNN} {r : NNProp} [CovariantClass M N μ r] [IsTrans N r] (m : M) {a b c : N} (ab : r a b) (rr : r c (μ m a)) :
r c (μ m b)
theorem act_rel_act_of_rel_of_rel {N : Type u_2} {r : NNProp} {mu : NNN} [IsTrans N r] [i : CovariantClass N N mu r] [i' : CovariantClass N N (Function.swap mu) r] {a b c d : N} (ab : r a b) (cd : r c d) :
r (mu a c) (mu b d)
theorem rel_of_act_rel_act {M : Type u_1} {N : Type u_2} {μ : MNN} {r : NNProp} [ContravariantClass M N μ r] (m : M) {a b : N} (ab : r (μ m a) (μ m b)) :
r a b
theorem act_rel_of_act_rel_of_rel_act_rel {M : Type u_1} {N : Type u_2} {μ : MNN} {r : NNProp} [ContravariantClass M N μ r] [IsTrans N r] (m : M) {a b c : N} (ab : r (μ m a) b) (rl : r (μ m b) (μ m c)) :
r (μ m a) c
theorem rel_act_of_act_rel_act_of_rel_act {M : Type u_1} {N : Type u_2} {μ : MNN} {r : NNProp} [ContravariantClass M N μ r] [IsTrans N r] (m : M) {a b c : N} (ab : r (μ m a) (μ m b)) (rr : r b (μ m c)) :
r a (μ m c)
theorem Covariant.monotone_of_const {M : Type u_1} {N : Type u_2} {μ : MNN} [Preorder N] [CovariantClass M N μ fun (x1 x2 : N) => x1 x2] (m : M) :
Monotone (μ m)

The partial application of a constant to a covariant operator is monotone.

theorem Monotone.covariant_of_const {M : Type u_1} {N : Type u_2} {μ : MNN} {α : Type u_3} [Preorder α] [Preorder N] {f : Nα} [CovariantClass M N μ fun (x1 x2 : N) => x1 x2] (hf : Monotone f) (m : M) :
Monotone fun (x : N) => f (μ m x)

A monotone function remains monotone when composed with the partial application of a covariant operator. E.g., ∀ (m : ℕ), Monotone f → Monotone (fun n ↦ f (m + n)).

theorem Monotone.covariant_of_const' {N : Type u_2} {α : Type u_3} [Preorder α] [Preorder N] {f : Nα} {μ : NNN} [CovariantClass N N (Function.swap μ) fun (x1 x2 : N) => x1 x2] (hf : Monotone f) (m : N) :
Monotone fun (x : N) => f (μ x m)

Same as Monotone.covariant_of_const, but with the constant on the other side of the operator. E.g., ∀ (m : ℕ), Monotone f → Monotone (fun n ↦ f (n + m)).

theorem Antitone.covariant_of_const {M : Type u_1} {N : Type u_2} {μ : MNN} {α : Type u_3} [Preorder α] [Preorder N] {f : Nα} [CovariantClass M N μ fun (x1 x2 : N) => x1 x2] (hf : Antitone f) (m : M) :
Antitone fun (x : N) => f (μ m x)

Dual of Monotone.covariant_of_const

theorem Antitone.covariant_of_const' {N : Type u_2} {α : Type u_3} [Preorder α] [Preorder N] {f : Nα} {μ : NNN} [CovariantClass N N (Function.swap μ) fun (x1 x2 : N) => x1 x2] (hf : Antitone f) (m : N) :
Antitone fun (x : N) => f (μ x m)

Dual of Monotone.covariant_of_const'

theorem covariant_le_of_covariant_lt (M : Type u_1) (N : Type u_2) (μ : MNN) [PartialOrder N] :
(Covariant M N μ fun (x1 x2 : N) => x1 < x2)Covariant M N μ fun (x1 x2 : N) => x1 x2
theorem covariantClass_le_of_lt (M : Type u_1) (N : Type u_2) (μ : MNN) [PartialOrder N] [CovariantClass M N μ fun (x1 x2 : N) => x1 < x2] :
CovariantClass M N μ fun (x1 x2 : N) => x1 x2
theorem contravariant_le_iff_contravariant_lt_and_eq (M : Type u_1) (N : Type u_2) (μ : MNN) [PartialOrder N] :
(Contravariant M N μ fun (x1 x2 : N) => x1 x2) (Contravariant M N μ fun (x1 x2 : N) => x1 < x2) Contravariant M N μ fun (x1 x2 : N) => x1 = x2
theorem contravariant_lt_of_contravariant_le (M : Type u_1) (N : Type u_2) (μ : MNN) [PartialOrder N] :
(Contravariant M N μ fun (x1 x2 : N) => x1 x2)Contravariant M N μ fun (x1 x2 : N) => x1 < x2
theorem covariant_le_iff_contravariant_lt (M : Type u_1) (N : Type u_2) (μ : MNN) [LinearOrder N] :
(Covariant M N μ fun (x1 x2 : N) => x1 x2) Contravariant M N μ fun (x1 x2 : N) => x1 < x2
theorem covariant_lt_iff_contravariant_le (M : Type u_1) (N : Type u_2) (μ : MNN) [LinearOrder N] :
(Covariant M N μ fun (x1 x2 : N) => x1 < x2) Contravariant M N μ fun (x1 x2 : N) => x1 x2
theorem covariant_flip_iff (N : Type u_2) (r : NNProp) (mu : NNN) [h : Std.Commutative mu] :
Covariant N N (flip mu) r Covariant N N mu r
theorem contravariant_flip_iff (N : Type u_2) (r : NNProp) (mu : NNN) [h : Std.Commutative mu] :
Contravariant N N (flip mu) r Contravariant N N mu r
instance contravariant_lt_of_covariant_le (N : Type u_2) (mu : NNN) [LinearOrder N] [CovariantClass N N mu fun (x1 x2 : N) => x1 x2] :
ContravariantClass N N mu fun (x1 x2 : N) => x1 < x2
instance covariant_lt_of_contravariant_le (N : Type u_2) (mu : NNN) [LinearOrder N] [ContravariantClass N N mu fun (x1 x2 : N) => x1 x2] :
CovariantClass N N mu fun (x1 x2 : N) => x1 < x2
instance covariant_swap_mul_of_covariant_mul (N : Type u_2) (r : NNProp) [CommSemigroup N] [CovariantClass N N (fun (x1 x2 : N) => x1 * x2) r] :
CovariantClass N N (Function.swap fun (x1 x2 : N) => x1 * x2) r
instance covariant_swap_add_of_covariant_add (N : Type u_2) (r : NNProp) [AddCommSemigroup N] [CovariantClass N N (fun (x1 x2 : N) => x1 + x2) r] :
CovariantClass N N (Function.swap fun (x1 x2 : N) => x1 + x2) r
instance contravariant_swap_mul_of_contravariant_mul (N : Type u_2) (r : NNProp) [CommSemigroup N] [ContravariantClass N N (fun (x1 x2 : N) => x1 * x2) r] :
ContravariantClass N N (Function.swap fun (x1 x2 : N) => x1 * x2) r
instance contravariant_swap_add_of_contravariant_add (N : Type u_2) (r : NNProp) [AddCommSemigroup N] [ContravariantClass N N (fun (x1 x2 : N) => x1 + x2) r] :
ContravariantClass N N (Function.swap fun (x1 x2 : N) => x1 + x2) r
theorem covariant_lt_of_covariant_le_of_contravariant_eq (M : Type u_1) (N : Type u_2) (μ : MNN) [ContravariantClass M N μ fun (x1 x2 : N) => x1 = x2] [PartialOrder N] [CovariantClass M N μ fun (x1 x2 : N) => x1 x2] :
CovariantClass M N μ fun (x1 x2 : N) => x1 < x2
theorem contravariant_le_of_contravariant_eq_and_lt (M : Type u_1) (N : Type u_2) (μ : MNN) [PartialOrder N] [ContravariantClass M N μ fun (x1 x2 : N) => x1 = x2] [ContravariantClass M N μ fun (x1 x2 : N) => x1 < x2] :
ContravariantClass M N μ fun (x1 x2 : N) => x1 x2