Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

feat: linear order is isomorphic to lexicographic sum of two intervals #20409

Open
wants to merge 11 commits into
base: master
Choose a base branch
from
Open
Show file tree
Hide file tree
Changes from 6 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
1 change: 1 addition & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4134,6 +4134,7 @@ import Mathlib.Order.Hom.Basic
import Mathlib.Order.Hom.Bounded
import Mathlib.Order.Hom.CompleteLattice
import Mathlib.Order.Hom.Lattice
import Mathlib.Order.Hom.Lex
import Mathlib.Order.Hom.Order
import Mathlib.Order.Hom.Set
import Mathlib.Order.Ideal
Expand Down
26 changes: 25 additions & 1 deletion Mathlib/Data/Sum/Order.lean
Original file line number Diff line number Diff line change
Expand Up @@ -483,6 +483,18 @@ namespace OrderIso

variable [LE α] [LE β] [LE γ] (a : α) (b : β) (c : γ)

/-- `Equiv.sumCongr` promoted to an order isomorphism. -/
@[simps! apply]
def sumCongr {α₁ α₂ β₁ β₂ : Type*} [LE α₁] [LE α₂] [LE β₁] [LE β₂]
(ea : α₁ ≃o α₂) (eb : β₁ ≃o β₂) : α₁ ⊕ β₁ ≃o α₂ ⊕ β₂ where
toEquiv := .sumCongr ea eb
map_rel_iff' := by rintro (a | a) (b | b) <;> simp

@[simp]
theorem sumCongr_symm {α₁ α₂ β₁ β₂ : Type*} [LE α₁] [LE α₂] [LE β₁] [LE β₂]
(ea : α₁ ≃o α₂) (eb : β₁ ≃o β₂) : (sumCongr ea eb).symm = sumCongr ea.symm eb.symm :=
rfl

/-- `Equiv.sumComm` promoted to an order isomorphism. -/
@[simps! apply]
def sumComm (α β : Type*) [LE α] [LE β] : α ⊕ β ≃o β ⊕ α :=
Expand Down Expand Up @@ -552,7 +564,19 @@ theorem sumDualDistrib_symm_inl : (sumDualDistrib α β).symm (inl (toDual a)) =
theorem sumDualDistrib_symm_inr : (sumDualDistrib α β).symm (inr (toDual b)) = toDual (inr b) :=
rfl

/-- `Equiv.SumAssoc` promoted to an order isomorphism. -/
/-- `Equiv.sumCongr` promoted to an order isomorphism. -/
@[simps! apply]
def sumLexCongr {α₁ α₂ β₁ β₂ : Type*} [LE α₁] [LE α₂] [LE β₁] [LE β₂]
(ea : α₁ ≃o α₂) (eb : β₁ ≃o β₂) : α₁ ⊕ₗ β₁ ≃o α₂ ⊕ₗ β₂ where
toEquiv := ofLex.trans ((Equiv.sumCongr ea eb).trans toLex)
map_rel_iff' := by simp_rw [Lex.forall]; rintro (a | a) (b | b) <;> simp

@[simp]
theorem sumLexCongr_symm {α₁ α₂ β₁ β₂ : Type*} [LE α₁] [LE α₂] [LE β₁] [LE β₂]
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

It would be nice to have the trans and refl versions of this too.

(ea : α₁ ≃o α₂) (eb : β₁ ≃o β₂) : (sumLexCongr ea eb).symm = sumLexCongr ea.symm eb.symm :=
rfl

/-- `Equiv.sumAssoc` promoted to an order isomorphism. -/
def sumLexAssoc (α β γ : Type*) [LE α] [LE β] [LE γ] : (α ⊕ₗ β) ⊕ₗ γ ≃o α ⊕ₗ β ⊕ₗ γ :=
{ Equiv.sumAssoc α β γ with
map_rel_iff' := fun {a b} =>
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Logic/Equiv/Set.lean
Original file line number Diff line number Diff line change
Expand Up @@ -148,7 +148,7 @@ def setProdEquivSigma {α β : Type*} (s : Set (α × β)) :
right_inv := fun ⟨_, _, _⟩ => rfl

/-- The subtypes corresponding to equal sets are equivalent. -/
@[simps! apply]
@[simps! apply symm_apply]
def setCongr {α : Type*} {s t : Set α} (h : s = t) : s ≃ t :=
subtypeEquivProp h

Expand Down
177 changes: 177 additions & 0 deletions Mathlib/Order/Hom/Lex.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,177 @@
/-
Copyright (c) 2025 Violeta Hernández Palacios. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Violeta Hernández Palacios
-/
import Mathlib.Data.Sum.Order
import Mathlib.Order.Hom.Set
import Mathlib.Order.RelIso.Set

/-!
# Lexicographic order and order isomorphisms

The main result in this file is the following: if `α` is a linear order and `x : α`, then `α` is
order isomorphic to either of `Iio x ⊕ₗ Ici x` or `Iic x ⊕ₗ Ioi x`.
-/

open Set

variable {α : Type*}

/-! ### Relation isomorphism -/

namespace RelIso

variable {r : α → α → Prop} {x y : α} [IsTrans α r] [IsTrichotomous α r] [DecidableRel r]

variable (r x) in
/-- A relation is isomorphic to the lexicographic sum of elements less than `x` and elements not
less than `x`. -/
-- The explicit type signature stops `simp` from complaining.
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Where does this complaining happen? If you're able to make a mwe, can you file an issue with it?

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

This complaining is addressed in #20479, I can make this a dependent PR so that this doesn't come up

Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Do you mean #20475?

def sumLexComplLeft :
@Sum.Lex {y // r y x} {y // ¬ r y x} (Subrel r {y | r y x}) (Subrel r {y | ¬ r y x}) ≃r r where
toEquiv := Equiv.sumCompl (r · x)
map_rel_iff' := by
rintro (⟨a, ha⟩ | ⟨a, ha⟩) (⟨b, hb⟩ | ⟨b, hb⟩)
· simp
· simpa using trans_trichotomous_right ha hb
· simpa using fun h ↦ ha <| trans h hb
· simp

@[simp]
theorem sumLexComplLeft_apply_inl (a : {y // r y x}) : sumLexComplLeft r x (Sum.inl a) = a :=
rfl

@[simp]
theorem sumLexComplLeft_apply_inr (a : {y // ¬ r y x}) : sumLexComplLeft r x (Sum.inr a) = a :=
rfl

theorem sumLexComplLeft_symm_apply_of_mem (h : r y x) :
(sumLexComplLeft r x).symm y = Sum.inl ⟨y, h⟩ :=
Equiv.sumCompl_apply_symm_of_pos (r · x) _ h

theorem sumLexComplLeft_symm_apply_of_not_mem (h : ¬ r y x) :
(sumLexComplLeft r x).symm y = Sum.inr ⟨y, h⟩ :=
Equiv.sumCompl_apply_symm_of_neg (r · x) _ h

@[simp]
theorem sumLexComplLeft_symm_apply (a : {y // r y x}) :
(sumLexComplLeft r x).symm a = Sum.inl a :=
sumLexComplLeft_symm_apply_of_mem a.2

@[simp]
theorem sumLexComplLeft_symm_apply_neg (a : {y // ¬ r y x}) :
(sumLexComplLeft r x).symm a = Sum.inr a :=
sumLexComplLeft_symm_apply_of_not_mem a.2
vihdzp marked this conversation as resolved.
Show resolved Hide resolved

variable (r x) in
/-- A relation is isomorphic to the lexicographic sum of elements not greater than `x` and elements
greater than `x`. -/
-- The explicit type signature stops `simp` from complaining.
def sumLexComplRight :
@Sum.Lex {y // ¬ r x y} {y // r x y} (Subrel r {y | ¬ r x y}) (Subrel r {y | r x y}) ≃r r where
toEquiv := (Equiv.sumComm _ _).trans <| Equiv.sumCompl (r x ·)
map_rel_iff' := by
rintro (⟨a, ha⟩ | ⟨a, ha⟩) (⟨b, hb⟩ | ⟨b, hb⟩)
· simp
· simpa using trans_trichotomous_left ha hb
· simpa using fun h ↦ hb <| trans ha h
· simp

@[simp]
theorem sumLexComplRight_apply_inl (a : {y // ¬ r x y}) : sumLexComplRight r x (Sum.inl a) = a :=
rfl

@[simp]
theorem sumLexComplRight_apply_inr (a : {y // r x y}) : sumLexComplRight r x (Sum.inr a) = a :=
rfl

theorem sumLexComplRight_symm_apply_of_mem (h : r x y) :
(sumLexComplRight r x).symm y = Sum.inr ⟨y, h⟩ := by
simp [sumLexComplRight, h]

theorem sumLexComplRight_symm_apply_of_not_mem (h : ¬ r x y) :
(sumLexComplRight r x).symm y = Sum.inl ⟨y, h⟩ := by
simp [sumLexComplRight, h]

@[simp]
theorem sumLexComplRight_symm_apply (a : {y // r x y}) :
(sumLexComplRight r x).symm a = Sum.inr a :=
sumLexComplRight_symm_apply_of_mem a.2

@[simp]
theorem sumLexComplRight_symm_apply_neg (a : {y // ¬ r x y}) :
(sumLexComplRight r x).symm a = Sum.inl a :=
sumLexComplRight_symm_apply_of_not_mem a.2

end RelIso

/-! ### Order isomorphism -/

namespace OrderIso

variable [LinearOrder α] {x y : α}

variable (x) in
/-- A linear order is isomorphic to the lexicographic sum of elements less than `x` and elements
greater or equal to `x`. -/
def sumLexIioIci : Iio x ⊕ₗ Ici x ≃o α :=
(sumLexCongr (refl _) (setCongr (Ici x) {y | ¬ y < x} (by ext; simp))).trans <|
ofRelIsoLT (RelIso.sumLexComplLeft (· < ·) x)
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I think you should be composing with ofLex here somehow. Do we have ofLex as a RelIso for Sum?

Copy link
Collaborator Author

@vihdzp vihdzp Jan 5, 2025

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

There's some pretty subtle def-eq finagling going on here: the reason ofRelIsoLT works is because Sum.Lex (· < ·) (· < ·) (Iio x) (Ici x) is def-eq to the < relation on Iio x ⊕ₗ Ici x. I'm not sure this can be made less abusive by just inserting ofLex.

In any case things should probably be fine as long as we keep the actual API free of def-eq abuse.

Copy link
Member

@eric-wieser eric-wieser Jan 5, 2025

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I think we should introduce a Sum.toLexRelIso : Sum.Lex (. < .) (. < .) ≃r (fun a b : Lex (Sum _ _) => a < b) or whatever the true version is, with simp lemmas reducing to toLex, and cast through that here


@[simp]
theorem sumLexIioIci_apply_inl (a : Iio x) : sumLexIioIci x (Sum.inl a) = a :=
rfl

@[simp]
theorem sumLexIioIci_apply_inr (a : Ici x) : sumLexIioIci x (Sum.inr a) = a :=
rfl
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Suggested change
theorem sumLexIioIci_apply_inl (a : Iio x) : sumLexIioIci x (Sum.inl a) = a :=
rfl
@[simp]
theorem sumLexIioIci_apply_inr (a : Ici x) : sumLexIioIci x (Sum.inr a) = a :=
rfl
theorem sumLexIioIci_apply_inl (a : Iio x) : sumLexIioIci x (ofLex <| Sum.inl a) = a :=
rfl
@[simp]
theorem sumLexIioIci_apply_inr (a : Ici x) : sumLexIioIci x (ofLex <| Sum.inr a) = a :=
rfl

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

It'd be better for simp to instead put toLex on the RHS.

Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

That suggestion doesn't make any sense to me, the type of this bundled map is Lex (Sum (Iio x) (Ici x)) ≃o α, so when applying the map to a Sum you have to call toLex first.

(I realize I made a type above too...)

Copy link
Collaborator Author

@vihdzp vihdzp Jan 5, 2025

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Oh nevermind, I see why what I was saying is nonsense. And in fact I'm missing some other ofLex elsewhere.


theorem sumLexIioIci_symm_apply_of_lt (h : y < x) :
(sumLexIioIci x).symm y = Sum.inl ⟨y, h⟩ := by
rw [symm_apply_eq, sumLexIioIci_apply_inl]

theorem sumLexIioIci_symm_apply_of_le {y : α} (h : x ≤ y) :
(sumLexIioIci x).symm y = Sum.inr ⟨y, h⟩ := by
rw [symm_apply_eq, sumLexIioIci_apply_inr]

@[simp]
theorem sumLexIioIci_symm_apply_Iio (a : Iio x) : (sumLexIioIci x).symm a = Sum.inl a :=
sumLexIioIci_symm_apply_of_lt a.2

@[simp]
theorem sumLexIioIci_symm_apply_Ici (a : Ici x) : (sumLexIioIci x).symm a = Sum.inr a :=
sumLexIioIci_symm_apply_of_le a.2

variable (x) in
/-- A linear order is isomorphic to the lexicographic sum of elements less or equal to `x` and
elements greater than `x`. -/
def sumLexIicIoi : Iic x ⊕ₗ Ioi x ≃o α :=
(sumLexCongr (setCongr (Iic x) {y | ¬ x < y} (by ext; simp)) (refl _)).trans <|
ofRelIsoLT (RelIso.sumLexComplRight (· < ·) x)

@[simp]
theorem sumLexIicIoi_apply_inl (a : Iic x) : sumLexIicIoi x (Sum.inl a) = a :=
rfl

@[simp]
theorem sumLexIicIoi_apply_inr (a : Ioi x) : sumLexIicIoi x (Sum.inr a) = a :=
rfl

theorem sumLexIicIoi_symm_apply_of_le (h : y ≤ x) :
(sumLexIicIoi x).symm y = Sum.inl ⟨y, h⟩ := by
rw [symm_apply_eq, sumLexIicIoi_apply_inl]

theorem sumLexIicIoi_symm_apply_of_lt {y : α} (h : x < y) :
(sumLexIicIoi x).symm y = Sum.inr ⟨y, h⟩ := by
rw [symm_apply_eq, sumLexIicIoi_apply_inr]

@[simp]
theorem sumLexIicIoi_symm_apply_Iic (a : Iic x) : (sumLexIicIoi x).symm a = Sum.inl a :=
sumLexIicIoi_symm_apply_of_le a.2

@[simp]
theorem sumLexIicIoi_symm_apply_Ioi (a : Ioi x) : (sumLexIicIoi x).symm a = Sum.inr a :=
sumLexIicIoi_symm_apply_of_lt a.2

end OrderIso
1 change: 1 addition & 0 deletions Mathlib/Order/Hom/Set.lean
Original file line number Diff line number Diff line change
Expand Up @@ -62,6 +62,7 @@ open Set
variable [Preorder α]

/-- Order isomorphism between two equal sets. -/
@[simps! apply symm_apply]
def setCongr (s t : Set α) (h : s = t) :
s ≃o t where
toEquiv := Equiv.setCongr h
Expand Down
Loading