Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit ceb13ba

Browse files
urkudbryangingechen
andcommitted
chore(order/basic): add monotone.order_dual, strict_mono.order_dual (#2778)
Also split long lines and lint. Co-authored-by: Bryan Gin-ge Chen <bryangingechen@gmail.com>
1 parent 90abd3b commit ceb13ba

File tree

2 files changed

+70
-38
lines changed

2 files changed

+70
-38
lines changed

scripts/nolints.txt

Lines changed: 1 addition & 19 deletions
Original file line numberDiff line numberDiff line change
@@ -1999,25 +1999,7 @@ apply_nolint pell.xz doc_blame
19991999
apply_nolint pell.yz doc_blame
20002000

20012001
-- order/basic.lean
2002-
apply_nolint classical.DLO doc_blame
20032002
apply_nolint decidable_linear_order_of_is_well_order doc_blame
2004-
apply_nolint dense_or_discrete ge_or_gt
2005-
apply_nolint directed_order doc_blame
2006-
apply_nolint ge.is_antisymm ge_or_gt
2007-
apply_nolint ge.is_linear_order ge_or_gt
2008-
apply_nolint ge.is_partial_order ge_or_gt
2009-
apply_nolint ge.is_preorder ge_or_gt
2010-
apply_nolint ge.is_refl ge_or_gt
2011-
apply_nolint ge.is_total ge_or_gt
2012-
apply_nolint ge.is_total_preorder ge_or_gt
2013-
apply_nolint ge.is_trans ge_or_gt
2014-
apply_nolint ge_of_eq ge_or_gt
2015-
apply_nolint gt.is_antisymm ge_or_gt
2016-
apply_nolint gt.is_asymm ge_or_gt
2017-
apply_nolint gt.is_irrefl ge_or_gt
2018-
apply_nolint gt.is_strict_order ge_or_gt
2019-
apply_nolint gt.is_trans ge_or_gt
2020-
apply_nolint gt.is_trichotomous ge_or_gt
20212003
apply_nolint well_founded.succ doc_blame
20222004
apply_nolint well_founded.sup doc_blame
20232005

@@ -3094,4 +3076,4 @@ apply_nolint uniform_space.separation_setoid doc_blame
30943076

30953077
-- topology/uniform_space/uniform_embedding.lean
30963078
apply_nolint uniform_embedding doc_blame
3097-
apply_nolint uniform_inducing doc_blame
3079+
apply_nolint uniform_inducing doc_blame

src/order/basic.lean

Lines changed: 69 additions & 19 deletions
Original file line numberDiff line numberDiff line change
@@ -52,6 +52,7 @@ preorder, order, partial order, linear order, monotone, strictly monotone
5252
universes u v w
5353
variables {α : Type u} {β : Type v} {γ : Type w} {r : α → α → Prop}
5454

55+
@[nolint ge_or_gt] -- see Note [nolint_ge]
5556
theorem ge_of_eq [preorder α] {a b : α} : a = b → a ≥ b :=
5657
λ h, h ▸ le_refl a
5758

@@ -83,33 +84,48 @@ lemma antisymm_of_asymm (r) [is_asymm α r] : is_antisymm α r :=
8384

8485
/- Convert algebraic structure style to explicit relation style typeclasses -/
8586
instance [preorder α] : is_refl α (≤) := ⟨le_refl⟩
87+
@[nolint ge_or_gt] -- see Note [nolint_ge]
8688
instance [preorder α] : is_refl α (≥) := is_refl.swap _
8789
instance [preorder α] : is_trans α (≤) := ⟨@le_trans _ _⟩
90+
@[nolint ge_or_gt] -- see Note [nolint_ge]
8891
instance [preorder α] : is_trans α (≥) := is_trans.swap _
8992
instance [preorder α] : is_preorder α (≤) := {}
93+
@[nolint ge_or_gt] -- see Note [nolint_ge]
9094
instance [preorder α] : is_preorder α (≥) := {}
9195
instance [preorder α] : is_irrefl α (<) := ⟨lt_irrefl⟩
96+
@[nolint ge_or_gt] -- see Note [nolint_ge]
9297
instance [preorder α] : is_irrefl α (>) := is_irrefl.swap _
9398
instance [preorder α] : is_trans α (<) := ⟨@lt_trans _ _⟩
99+
@[nolint ge_or_gt] -- see Note [nolint_ge]
94100
instance [preorder α] : is_trans α (>) := is_trans.swap _
95101
instance [preorder α] : is_asymm α (<) := ⟨@lt_asymm _ _⟩
102+
@[nolint ge_or_gt] -- see Note [nolint_ge]
96103
instance [preorder α] : is_asymm α (>) := is_asymm.swap _
97104
instance [preorder α] : is_antisymm α (<) := antisymm_of_asymm _
105+
@[nolint ge_or_gt] -- see Note [nolint_ge]
98106
instance [preorder α] : is_antisymm α (>) := antisymm_of_asymm _
99107
instance [preorder α] : is_strict_order α (<) := {}
108+
@[nolint ge_or_gt] -- see Note [nolint_ge]
100109
instance [preorder α] : is_strict_order α (>) := {}
101110
instance preorder.is_total_preorder [preorder α] [is_total α (≤)] : is_total_preorder α (≤) := {}
102111
instance [partial_order α] : is_antisymm α (≤) := ⟨@le_antisymm _ _⟩
112+
@[nolint ge_or_gt] -- see Note [nolint_ge]
103113
instance [partial_order α] : is_antisymm α (≥) := is_antisymm.swap _
104114
instance [partial_order α] : is_partial_order α (≤) := {}
115+
@[nolint ge_or_gt] -- see Note [nolint_ge]
105116
instance [partial_order α] : is_partial_order α (≥) := {}
106117
instance [linear_order α] : is_total α (≤) := ⟨le_total⟩
118+
@[nolint ge_or_gt] -- see Note [nolint_ge]
107119
instance [linear_order α] : is_total α (≥) := is_total.swap _
108-
instance linear_order.is_total_preorder [linear_order α] : is_total_preorder α (≤) := by apply_instance
120+
instance linear_order.is_total_preorder [linear_order α] : is_total_preorder α (≤) :=
121+
by apply_instance
122+
@[nolint ge_or_gt] -- see Note [nolint_ge]
109123
instance [linear_order α] : is_total_preorder α (≥) := {}
110124
instance [linear_order α] : is_linear_order α (≤) := {}
125+
@[nolint ge_or_gt] -- see Note [nolint_ge]
111126
instance [linear_order α] : is_linear_order α (≥) := {}
112127
instance [linear_order α] : is_trichotomous α (<) := ⟨lt_trichotomy⟩
128+
@[nolint ge_or_gt] -- see Note [nolint_ge]
113129
instance [linear_order α] : is_trichotomous α (>) := is_trichotomous.swap _
114130

115131
theorem preorder.ext {α} {A B : preorder α}
@@ -291,17 +307,23 @@ instance pi.preorder {ι : Type u} {α : ι → Type v} [∀i, preorder (α i)]
291307
le_refl := assume a i, le_refl (a i),
292308
le_trans := assume a b c h₁ h₂ i, le_trans (h₁ i) (h₂ i) }
293309

294-
instance pi.partial_order {ι : Type u} {α : ι → Type v} [∀i, partial_order (α i)] : partial_order (Πi, α i) :=
310+
instance pi.partial_order {ι : Type u} {α : ι → Type v} [∀i, partial_order (α i)] :
311+
partial_order (Πi, α i) :=
295312
{ le_antisymm := λf g h1 h2, funext (λb, le_antisymm (h1 b) (h2 b)),
296313
..pi.preorder }
297314

298315
theorem comp_le_comp_left_of_monotone [preorder α] [preorder β]
299-
{f : β → α} {g h : γ → β} (m_f : monotone f) (le_gh : g ≤ h) : has_le.le.{max w u} (f ∘ g) (f ∘ h) :=
316+
{f : β → α} {g h : γ → β} (m_f : monotone f) (le_gh : g ≤ h) :
317+
has_le.le.{max w u} (f ∘ g) (f ∘ h) :=
300318
assume x, m_f (le_gh x)
301319

302320
section monotone
303321
variables [preorder α] [preorder γ]
304322

323+
theorem monotone.order_dual {f : α → γ} (hf : monotone f) :
324+
@monotone (order_dual α) (order_dual γ) _ _ f :=
325+
λ x y hxy, hf hxy
326+
305327
theorem monotone_lam {f : α → β → γ} (m : ∀b, monotone (λa, f a b)) : monotone f :=
306328
assume a a' h b, m b h
307329

@@ -310,6 +332,10 @@ assume a a' h, m h b
310332

311333
end monotone
312334

335+
theorem strict_mono.order_dual [has_lt α] [has_lt β] {f : α → β} (hf : strict_mono f) :
336+
@strict_mono (order_dual α) (order_dual β) _ _ f :=
337+
λ x y hxy, hf hxy
338+
313339
/-- Transfer a `preorder` on `β` to a `preorder` on `α` using a function `f : α → β`. -/
314340
def preorder.lift {α β} (f : α → β) (i : preorder β) : preorder α :=
315341
by exactI
@@ -346,7 +372,8 @@ by exactI
346372
instance subtype.preorder {α} [i : preorder α] (p : α → Prop) : preorder (subtype p) :=
347373
preorder.lift subtype.val i
348374

349-
instance subtype.partial_order {α} [i : partial_order α] (p : α → Prop) : partial_order (subtype p) :=
375+
instance subtype.partial_order {α} [i : partial_order α] (p : α → Prop) :
376+
partial_order (subtype p) :=
350377
partial_order.lift subtype.val subtype.val_injective i
351378

352379
instance subtype.linear_order {α} [i : linear_order α] (p : α → Prop) : linear_order (subtype p) :=
@@ -411,7 +438,8 @@ instance order_dual.densely_ordered (α : Type u) [preorder α] [densely_ordered
411438
densely_ordered (order_dual α) :=
412439
⟨λ a₁ a₂ ha, (@dense α _ _ _ _ ha).imp $ λ a, and.symm⟩
413440

414-
lemma le_of_forall_le_of_dense [linear_order α] [densely_ordered α] {a₁ a₂ : α} (h : ∀a₃>a₂, a₁ ≤ a₃) :
441+
lemma le_of_forall_le_of_dense [linear_order α] [densely_ordered α] {a₁ a₂ : α}
442+
(h : ∀a₃>a₂, a₁ ≤ a₃) :
415443
a₁ ≤ a₂ :=
416444
le_of_not_gt $ assume ha,
417445
let ⟨a, ha₁, ha₂⟩ := dense ha in
@@ -421,7 +449,8 @@ lemma eq_of_le_of_forall_le_of_dense [linear_order α] [densely_ordered α] {a
421449
(h₁ : a₂ ≤ a₁) (h₂ : ∀a₃>a₂, a₁ ≤ a₃) : a₁ = a₂ :=
422450
le_antisymm (le_of_forall_le_of_dense h₂) h₁
423451

424-
lemma le_of_forall_ge_of_dense [linear_order α] [densely_ordered α] {a₁ a₂ : α}(h : ∀a₃<a₁, a₂ ≥ a₃) :
452+
lemma le_of_forall_ge_of_dense [linear_order α] [densely_ordered α] {a₁ a₂ : α}
453+
(h : ∀a₃<a₁, a₂ ≥ a₃) :
425454
a₁ ≤ a₂ :=
426455
le_of_not_gt $ assume ha,
427456
let ⟨a, ha₁, ha₂⟩ := dense ha in
@@ -431,6 +460,7 @@ lemma eq_of_le_of_forall_ge_of_dense [linear_order α] [densely_ordered α] {a
431460
(h₁ : a₂ ≤ a₁) (h₂ : ∀a₃<a₁, a₂ ≥ a₃) : a₁ = a₂ :=
432461
le_antisymm (le_of_forall_ge_of_dense h₂) h₁
433462

463+
@[nolint ge_or_gt] -- see Note [nolint_ge]
434464
lemma dense_or_discrete [linear_order α] (a₁ a₂ : α) :
435465
(∃a, a₁ < a ∧ a < a₂) ∨ ((∀a>a₁, a ≥ a₂) ∧ (∀a<a₂, a ≤ a₁)) :=
436466
classical.or_iff_not_imp_left.2 $ assume h,
@@ -483,7 +513,8 @@ section prio
483513
set_option default_priority 100 -- see Note [default priority]
484514
/-- This is basically the same as `is_strict_total_order`, but that definition is
485515
in Type (probably by mistake) and also has redundant assumptions. -/
486-
@[algebra] class is_strict_total_order' (α : Type u) (lt : α → α → Prop) extends is_trichotomous α lt, is_strict_order α lt : Prop.
516+
@[algebra] class is_strict_total_order' (α : Type u) (lt : α → α → Prop)
517+
extends is_trichotomous α lt, is_strict_order α lt : Prop.
487518
end prio
488519

489520
/-- Construct a linear order from a `is_strict_total_order'` relation -/
@@ -497,15 +528,19 @@ def linear_order_of_STO' (r) [is_strict_total_order' α r] : linear_order α :=
497528
..partial_order_of_SO r }
498529

499530
/-- Construct a decidable linear order from a `is_strict_total_order'` relation -/
500-
def decidable_linear_order_of_STO' (r) [is_strict_total_order' α r] [decidable_rel r] : decidable_linear_order α :=
531+
def decidable_linear_order_of_STO' (r) [is_strict_total_order' α r] [decidable_rel r] :
532+
decidable_linear_order α :=
501533
by letI LO := linear_order_of_STO' r; exact
502534
{ decidable_le := λ x y, decidable_of_iff (¬ r y x) (@not_lt _ _ y x),
503535
..LO }
504536

537+
/-- Any `linear_order` is a noncomputable `decidable_linear_order`. This is not marked
538+
as an instance to avoid a loop. -/
505539
noncomputable def classical.DLO (α) [LO : linear_order α] : decidable_linear_order α :=
506540
{ decidable_le := classical.dec_rel _, ..LO }
507541

508-
theorem is_strict_total_order'.swap (r) [is_strict_total_order' α r] : is_strict_total_order' α (swap r) :=
542+
theorem is_strict_total_order'.swap (r) [is_strict_total_order' α r] :
543+
is_strict_total_order' α (swap r) :=
509544
{..is_trichotomous.swap r, ..is_strict_order.swap r}
510545

511546
instance [linear_order α] : is_strict_total_order' α (<) := {}
@@ -560,23 +595,31 @@ instance is_extensional_of_is_strict_total_order'
560595
section prio
561596
set_option default_priority 100 -- see Note [default priority]
562597
/-- A well order is a well-founded linear order. -/
563-
@[algebra] class is_well_order (α : Type u) (r : α → α → Prop) extends is_strict_total_order' α r : Prop :=
598+
@[algebra] class is_well_order (α : Type u) (r : α → α → Prop)
599+
extends is_strict_total_order' α r : Prop :=
564600
(wf : well_founded r)
565601
end prio
566602

567603
@[priority 100] -- see Note [lower instance priority]
568-
instance is_well_order.is_strict_total_order {α} (r : α → α → Prop) [is_well_order α r] : is_strict_total_order α r := by apply_instance
604+
instance is_well_order.is_strict_total_order {α} (r : α → α → Prop) [is_well_order α r] :
605+
is_strict_total_order α r := by apply_instance
569606
@[priority 100] -- see Note [lower instance priority]
570-
instance is_well_order.is_extensional {α} (r : α → α → Prop) [is_well_order α r] : is_extensional α r := by apply_instance
607+
instance is_well_order.is_extensional {α} (r : α → α → Prop) [is_well_order α r] :
608+
is_extensional α r := by apply_instance
571609
@[priority 100] -- see Note [lower instance priority]
572-
instance is_well_order.is_trichotomous {α} (r : α → α → Prop) [is_well_order α r] : is_trichotomous α r := by apply_instance
610+
instance is_well_order.is_trichotomous {α} (r : α → α → Prop) [is_well_order α r] :
611+
is_trichotomous α r := by apply_instance
573612
@[priority 100] -- see Note [lower instance priority]
574-
instance is_well_order.is_trans {α} (r : α → α → Prop) [is_well_order α r] : is_trans α r := by apply_instance
613+
instance is_well_order.is_trans {α} (r : α → α → Prop) [is_well_order α r] :
614+
is_trans α r := by apply_instance
575615
@[priority 100] -- see Note [lower instance priority]
576-
instance is_well_order.is_irrefl {α} (r : α → α → Prop) [is_well_order α r] : is_irrefl α r := by apply_instance
616+
instance is_well_order.is_irrefl {α} (r : α → α → Prop) [is_well_order α r] :
617+
is_irrefl α r := by apply_instance
577618
@[priority 100] -- see Note [lower instance priority]
578-
instance is_well_order.is_asymm {α} (r : α → α → Prop) [is_well_order α r] : is_asymm α r := by apply_instance
619+
instance is_well_order.is_asymm {α} (r : α → α → Prop) [is_well_order α r] :
620+
is_asymm α r := by apply_instance
579621

622+
/-- Construct a decidable linear order from a well-founded linear order. -/
580623
noncomputable def decidable_linear_order_of_is_well_order (r : α → α → Prop) [is_well_order α r] :
581624
decidable_linear_order α :=
582625
by { haveI := linear_order_of_STO' r, exact classical.DLO α }
@@ -589,13 +632,15 @@ instance empty_relation.is_well_order [subsingleton α] : is_well_order α empty
589632

590633
instance nat.lt.is_well_order : is_well_order ℕ (<) := ⟨nat.lt_wf⟩
591634

592-
instance sum.lex.is_well_order [is_well_order α r] [is_well_order β s] : is_well_order (α ⊕ β) (sum.lex r s) :=
635+
instance sum.lex.is_well_order [is_well_order α r] [is_well_order β s] :
636+
is_well_order (α ⊕ β) (sum.lex r s) :=
593637
{ trichotomous := λ a b, by cases a; cases b; simp; apply trichotomous,
594638
irrefl := λ a, by cases a; simp; apply irrefl,
595639
trans := λ a b c, by cases a; cases b; simp; cases c; simp; apply trans,
596640
wf := sum.lex_wf is_well_order.wf is_well_order.wf }
597641

598-
instance prod.lex.is_well_order [is_well_order α r] [is_well_order β s] : is_well_order (α × β) (prod.lex r s) :=
642+
instance prod.lex.is_well_order [is_well_order α r] [is_well_order β s] :
643+
is_well_order (α × β) (prod.lex r s) :=
599644
{ trichotomous := λ ⟨a₁, a₂⟩ ⟨b₁, b₂⟩,
600645
match @trichotomous _ r _ a₁ b₁ with
601646
| or.inl h₁ := or.inl $ prod.lex.left _ _ h₁
@@ -633,7 +678,7 @@ end
633678
by { classical, rw [not_iff_comm, not_bounded_iff] }
634679

635680
namespace well_founded
636-
/-- If `r` is a well founded relation, then any nonempty set has a minimum element
681+
/-- If `r` is a well-founded relation, then any nonempty set has a minimum element
637682
with respect to `r`. -/
638683
theorem has_min {α} {r : α → α → Prop} (H : well_founded r)
639684
(s : set α) : s.nonempty → ∃ a ∈ s, ∀ x ∈ s, ¬ r x a
@@ -654,6 +699,7 @@ theorem not_lt_min {α} {r : α → α → Prop} (H : well_founded r)
654699
let ⟨_, h'⟩ := classical.some_spec (H.has_min p h) in h' _ xp
655700

656701
open set
702+
/-- The supremum of a bounded, well-founded order -/
657703
protected noncomputable def sup {α} {r : α → α → Prop} (wf : well_founded r) (s : set α)
658704
(h : bounded r s) : α :=
659705
wf.min { x | ∀a ∈ s, r a x } h
@@ -664,6 +710,8 @@ min_mem wf { x | ∀a ∈ s, r a x } h x hx
664710

665711
section
666712
open_locale classical
713+
/-- The successor of an element `x` in a well-founded order is the minimum element `y` such that
714+
`x < y` if one exists. Otherwise it is `x` itself. -/
667715
protected noncomputable def succ {α} {r : α → α → Prop} (wf : well_founded r) (x : α) : α :=
668716
if h : ∃y, r x y then wf.min { y | r x y } h else x
669717

@@ -727,6 +775,8 @@ theorem directed.mono_comp {ι} {rb : β → β → Prop} {g : α → β} {f :
727775

728776
section prio
729777
set_option default_priority 100 -- see Note [default priority]
778+
/-- A `preorder` is a `directed_order` if for any two elements `i`, `j`
779+
there is an element `k` such that `i ≤ k` and `j ≤ k`. -/
730780
class directed_order (α : Type u) extends preorder α :=
731781
(directed : ∀ i j : α, ∃ k, i ≤ k ∧ j ≤ k)
732782
end prio

0 commit comments

Comments
 (0)