Documentation

Mathlib.Order.PiLex

Lexicographic order on Pi types #

This file defines the lexicographic order for Pi types. a is less than b if a i = b i for all i up to some point k, and a k < b k.

Notation #

See also #

Related files are:

def Pi.Lex {ι : Type u_1} {β : ιType u_2} (r : ιιProp) (s : {i : ι} → β iβ iProp) (x : (i : ι) → β i) (y : (i : ι) → β i) :

The lexicographic relation on Π i : ι, β i, where ι is ordered by r, and each β i is ordered by s.

Equations
  • Pi.Lex r s x y = ∃ (i : ι), (∀ (j : ι), r j ix j = y j) s (x i) (y i)
Instances For

    The notation Πₗ i, α i refers to a pi type equipped with the lexicographic order.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem Pi.toLex_apply {ι : Type u_1} {β : ιType u_2} (x : (i : ι) → β i) (i : ι) :
      toLex x i = x i
      @[simp]
      theorem Pi.ofLex_apply {ι : Type u_1} {β : ιType u_2} (x : Lex ((i : ι) → β i)) (i : ι) :
      ofLex x i = x i
      theorem Pi.lex_lt_of_lt_of_preorder {ι : Type u_1} {β : ιType u_2} [(i : ι) → Preorder (β i)] {r : ιιProp} (hwf : WellFounded r) {x : (i : ι) → β i} {y : (i : ι) → β i} (hlt : x < y) :
      ∃ (i : ι), (∀ (j : ι), r j ix j y j y j x j) x i < y i
      theorem Pi.lex_lt_of_lt {ι : Type u_1} {β : ιType u_2} [(i : ι) → PartialOrder (β i)] {r : ιιProp} (hwf : WellFounded r) {x : (i : ι) → β i} {y : (i : ι) → β i} (hlt : x < y) :
      Pi.Lex r (fun (i : ι) (x x_1 : β i) => x < x_1) x y
      theorem Pi.isTrichotomous_lex {ι : Type u_1} {β : ιType u_2} (r : ιιProp) (s : {i : ι} → β iβ iProp) [∀ (i : ι), IsTrichotomous (β i) s] (wf : WellFounded r) :
      IsTrichotomous ((i : ι) → β i) (Pi.Lex r s)
      instance Pi.instLTLexForall {ι : Type u_1} {β : ιType u_2} [LT ι] [(a : ι) → LT (β a)] :
      LT (Lex ((i : ι) → β i))
      Equations
      • Pi.instLTLexForall = { lt := Pi.Lex (fun (x x_1 : ι) => x < x_1) fun (x : ι) (x_1 x_2 : β x) => x_1 < x_2 }
      instance Pi.Lex.isStrictOrder {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [(a : ι) → PartialOrder (β a)] :
      IsStrictOrder (Lex ((i : ι) → β i)) fun (x x_1 : Lex ((i : ι) → β i)) => x < x_1
      Equations
      • =
      instance Pi.instPartialOrderLexForallOfLinearOrder {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [(a : ι) → PartialOrder (β a)] :
      PartialOrder (Lex ((i : ι) → β i))
      Equations
      • Pi.instPartialOrderLexForallOfLinearOrder = partialOrderOfSO fun (x x_1 : Lex ((i : ι) → β i)) => x < x_1
      noncomputable instance Pi.instLinearOrderLexForallOfIsWellOrderLt {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [(a : ι) → LinearOrder (β a)] :
      LinearOrder (Lex ((i : ι) → β i))

      Πₗ i, α i is a linear order if the original order is well-founded.

      Equations
      • Pi.instLinearOrderLexForallOfIsWellOrderLt = linearOrderOfSTO fun (x x_1 : Lex ((i : ι) → (fun (i : ι) => β i) i)) => x < x_1
      theorem Pi.toLex_monotone {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [(i : ι) → PartialOrder (β i)] :
      Monotone toLex
      theorem Pi.toLex_strictMono {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [(i : ι) → PartialOrder (β i)] :
      StrictMono toLex
      @[simp]
      theorem Pi.lt_toLex_update_self_iff {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [(i : ι) → PartialOrder (β i)] {x : (i : ι) → β i} {i : ι} {a : β i} :
      toLex x < toLex (Function.update x i a) x i < a
      @[simp]
      theorem Pi.toLex_update_lt_self_iff {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [(i : ι) → PartialOrder (β i)] {x : (i : ι) → β i} {i : ι} {a : β i} :
      toLex (Function.update x i a) < toLex x a < x i
      @[simp]
      theorem Pi.le_toLex_update_self_iff {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [(i : ι) → PartialOrder (β i)] {x : (i : ι) → β i} {i : ι} {a : β i} :
      toLex x toLex (Function.update x i a) x i a
      @[simp]
      theorem Pi.toLex_update_le_self_iff {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [(i : ι) → PartialOrder (β i)] {x : (i : ι) → β i} {i : ι} {a : β i} :
      toLex (Function.update x i a) toLex x a x i
      instance Pi.instOrderBotLexForallOfIsWellOrderLt {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [(a : ι) → PartialOrder (β a)] [(a : ι) → OrderBot (β a)] :
      OrderBot (Lex ((a : ι) → β a))
      Equations
      instance Pi.instOrderTopLexForallOfIsWellOrderLt {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [(a : ι) → PartialOrder (β a)] [(a : ι) → OrderTop (β a)] :
      OrderTop (Lex ((a : ι) → β a))
      Equations
      instance Pi.instBoundedOrderLexForallOfIsWellOrderLt {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [(a : ι) → PartialOrder (β a)] [(a : ι) → BoundedOrder (β a)] :
      BoundedOrder (Lex ((a : ι) → β a))
      Equations
      • Pi.instBoundedOrderLexForallOfIsWellOrderLt = BoundedOrder.mk
      instance Pi.instDenselyOrderedLexForall {ι : Type u_1} {β : ιType u_2} [Preorder ι] [(i : ι) → LT (β i)] [∀ (i : ι), DenselyOrdered (β i)] :
      DenselyOrdered (Lex ((i : ι) → β i))
      Equations
      • =
      theorem Pi.Lex.noMaxOrder' {ι : Type u_1} {β : ιType u_2} [Preorder ι] [(i : ι) → LT (β i)] (i : ι) [NoMaxOrder (β i)] :
      NoMaxOrder (Lex ((i : ι) → β i))
      instance Pi.instNoMaxOrderLexForallOfIsWellOrderLtOfNonempty {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [Nonempty ι] [(i : ι) → PartialOrder (β i)] [∀ (i : ι), NoMaxOrder (β i)] :
      NoMaxOrder (Lex ((i : ι) → β i))
      Equations
      • =
      instance Pi.instNoMinOrderLexForallOfIsWellOrderLtOfNonempty {ι : Type u_1} {β : ιType u_2} [LinearOrder ι] [IsWellOrder ι fun (x x_1 : ι) => x < x_1] [Nonempty ι] [(i : ι) → PartialOrder (β i)] [∀ (i : ι), NoMinOrder (β i)] :
      NoMinOrder (Lex ((i : ι) → β i))
      Equations
      • =
      theorem Pi.lex_desc {ι : Type u_1} {α : Type u_3} [Preorder ι] [DecidableEq ι] [Preorder α] {f : ια} {i : ι} {j : ι} (h₁ : i j) (h₂ : f j < f i) :
      toLex (f (Equiv.swap i j)) < toLex f

      If we swap two strictly decreasing values in a function, then the result is lexicographically smaller than the original function.