Skip to content

Commit

Permalink
chore(Order): move more defs to Defs (#17677)
Browse files Browse the repository at this point in the history
  • Loading branch information
urkud committed Oct 12, 2024
1 parent a02b71d commit 90f1b4d
Show file tree
Hide file tree
Showing 5 changed files with 31 additions and 26 deletions.
1 change: 1 addition & 0 deletions Mathlib/Data/Fintype/Order.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,7 @@ Authors: Peter Nelson, Yaël Dillies
import Mathlib.Data.Finset.Order
import Mathlib.Order.Atoms
import Mathlib.Data.Set.Finite
import Mathlib.Order.Minimal

/-!
# Order structures on finite types
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Order/CompleteLattice/Finset.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro
-/
import Mathlib.Data.Finset.Option
import Mathlib.Order.Minimal
import Mathlib.Data.Set.Lattice

/-!
# Lattice operations on finsets
Expand Down
26 changes: 26 additions & 0 deletions Mathlib/Order/Defs.lean
Original file line number Diff line number Diff line change
Expand Up @@ -184,6 +184,32 @@ end

end

/-! ### Minimal and maximal -/

section LE

variable {α : Type*} [LE α] {P : α → Prop} {x y : α}

/-- `Minimal P x` means that `x` is a minimal element satisfying `P`. -/
def Minimal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → y ≤ x → x ≤ y

/-- `Maximal P x` means that `x` is a maximal element satisfying `P`. -/
def Maximal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → x ≤ y → y ≤ x

lemma Minimal.prop (h : Minimal P x) : P x :=
h.1

lemma Maximal.prop (h : Maximal P x) : P x :=
h.1

lemma Minimal.le_of_le (h : Minimal P x) (hy : P y) (hle : y ≤ x) : x ≤ y :=
h.2 hy hle

lemma Maximal.le_of_ge (h : Maximal P x) (hy : P y) (hge : x ≤ y) : y ≤ x :=
h.2 hy hge

end LE

/-! ### Bundled classes -/

variable {α : Type*}
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Order/Filter/AtTopBot.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,6 +8,7 @@ import Mathlib.Order.ConditionallyCompleteLattice.Basic
import Mathlib.Order.Filter.Bases
import Mathlib.Order.Filter.Prod
import Mathlib.Order.Interval.Set.Disjoint
import Mathlib.Order.Interval.Set.OrderIso

/-!
# `Filter.atTop` and `Filter.atBot` filters on preorders, monoids and groups.
Expand Down
27 changes: 2 additions & 25 deletions Mathlib/Order/Minimal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -10,13 +10,8 @@ import Mathlib.Order.Interval.Set.Basic
/-!
# Minimality and Maximality
This file defines minimality and maximality of an element with respect to a predicate `P` on
an ordered type `α`.
## Main declarations
* `Minimal P x`: `x` is minimal satisfying `P`.
* `Maximal P x`: `x` is maximal satisfying `P`.
This file proves basic facts about minimality and maximality
of an element with respect to a predicate `P` on an ordered type `α`.
## Implementation Details
Expand Down Expand Up @@ -52,24 +47,6 @@ section LE

variable [LE α]

/-- `Minimal P x` means that `x` is a minimal element satisfying `P`. -/
def Minimal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → y ≤ x → x ≤ y

/-- `Maximal P x` means that `x` is a maximal element satisfying `P`. -/
def Maximal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → x ≤ y → y ≤ x

lemma Minimal.prop (h : Minimal P x) : P x :=
h.1

lemma Maximal.prop (h : Maximal P x) : P x :=
h.1

lemma Minimal.le_of_le (h : Minimal P x) (hy : P y) (hle : y ≤ x) : x ≤ y :=
h.2 hy hle

lemma Maximal.le_of_ge (h : Maximal P x) (hy : P y) (hge : x ≤ y) : y ≤ x :=
h.2 hy hge

@[simp] theorem minimal_toDual : Minimal (fun x ↦ P (ofDual x)) (toDual x) ↔ Maximal P x :=
Iff.rfl

Expand Down

0 comments on commit 90f1b4d

Please sign in to comment.