Skip to content

Commit

Permalink
Update Basic.lean
Browse files Browse the repository at this point in the history
  • Loading branch information
vihdzp committed Dec 4, 2024
1 parent a93cfd3 commit b7cb653
Showing 1 changed file with 7 additions and 4 deletions.
11 changes: 7 additions & 4 deletions Mathlib/Order/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -529,16 +529,19 @@ theorem PartialOrder.ext {A B : PartialOrder α}
ext x y
exact H x y

theorem PartialOrder.ext_lt {A B : PartialOrder α}
(H : ∀ x y : α, (haveI := A; x < y) ↔ x < y) : A = B := by
ext x y
rw [le_iff_lt_or_eq, @le_iff_lt_or_eq _ A, H]

theorem LinearOrder.ext {A B : LinearOrder α}
(H : ∀ x y : α, (haveI := A; x ≤ y) ↔ x ≤ y) : A = B := by
ext x y
exact H x y

theorem LinearOrder.ext_lt {A B : LinearOrder α}
(H : ∀ x y : α, (haveI := A; x < y) ↔ x < y) : A = B := by
ext x y
rw [← not_lt, ← @not_lt _ A, not_iff_not]
exact H y x
(H : ∀ x y : α, (haveI := A; x < y) ↔ x < y) : A = B :=
LinearOrder.toPartialOrder_injective (PartialOrder.ext_lt H)

/-- Given a relation `R` on `β` and a function `f : α → β`, the preimage relation on `α` is defined
by `x ≤ y ↔ f x ≤ f y`. It is the unique relation on `α` making `f` a `RelEmbedding` (assuming `f`
Expand Down

0 comments on commit b7cb653

Please sign in to comment.