Skip to content

Commit 970d241

Browse files
refactor: delete IsWellOrderLimitElement.lean (#15904)
I understand this is quite a bold move, and I hope not to come off as confrontational. I argue that `Order/IsWellOrderLimitElement.lean` duplicates existing API, and that all the theorems proven within it already exist in some other form. First, note that the typeclass assumptions `LinearOrder α` + `IsWellOrder α (· < ·)` **very nearly** imply `ConditionallyCompleteLinearOrderBot α` through [`IsWellOrder.conditionallyCompleteLinearOrderBot`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/ConditionallyCompleteLattice/Basic.html#IsWellOrder.conditionallyCompleteLinearOrderBot). The only missing assumption is `OrderBot α`, which actually follows from `Nonempty α` through [`WellFoundedLT.toOrderBot`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/WellFounded.html#WellFoundedLT.toOrderBot). All theorems in this file assume the existence of some `a : α`, so that's covered. Most of the only other file importing this one assumes `OrderBot α` directly. `ConditionallyCompleteLinearOrder α` in turn implies `SuccOrder α` through `ConditionallyCompleteLinearOrder.toSuccOrder` (#15863), so assuming it yields no loss of generality. In fact, doing things this way means we can set nicer def-eqs for `@succ α`, such as `succ o = o + 1` on ordinals. As for `IsWellOrderLimitElement`, we already have [`Order.IsSuccLimit`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit) with only trivial typeclass assumptions. To summarize, here are the correspondences between definitions and theorems in this file and other definitions and theorems in Mathlib, applicable under equivalent typeclasss assumptions: - `wellOrderSucc` → [`Order.succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.succ) - `self_le_wellOrderSucc` → [`Order.le_succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.le_succ) - `wellOrderSucc_le` → [`Order.succ_le_of_lt`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.succ_le_of_lt) - `self_lt_wellOrderSucc` → [`Order.lt_succ_of_not_isMax`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.lt_succ_of_not_isMax) - `le_of_lt_wellOrderSucc` → [`Order.le_of_lt_succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Basic.html#Order.le_of_lt_succ). - `IsWellOrderLimitElement`→ [`Order.IsSuccLimit`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit) - `IsWellOrderLimitElement.neq_bot` → (already implied by `IsSuccLimit`) - `IsWellOrderLimitElement.bot_lt` → [`Ne.bot_lt`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/BoundedOrder.html#Ne.bot_lt) - `IsWellOrderLimitElement.wellOrderSucc_lt` → [`Order.IsSuccLimit.succ_lt`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit.succ_lt) - `eq_bot_or_eq_succ_or_isWellOrderLimitElement` → [`Order.isSuccLimitRecOn`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.isSuccLimitRecOn) - `Nat.wellOrderSucc_eq` → [`Nat.succ_eq_succ`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Data/Nat/SuccPred.html#Nat.succ_eq_succ) - `Nat.not_isWellOrderLimitElement` → [`Order.IsSuccLimit.isMin_of_noMax`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/Order/SuccPred/Limit.html#Order.IsSuccLimit.isMin_of_noMax) Co-authored-by: leanprover-community-mathlib4-bot <leanprover-community-mathlib4-bot@users.noreply.github.com>
1 parent 6ac1e12 commit 970d241

File tree

3 files changed

+35
-126
lines changed

3 files changed

+35
-126
lines changed

Mathlib.lean

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3801,7 +3801,6 @@ import Mathlib.Order.Interval.Set.SurjOn
38013801
import Mathlib.Order.Interval.Set.UnorderedInterval
38023802
import Mathlib.Order.Interval.Set.WithBotTop
38033803
import Mathlib.Order.Irreducible
3804-
import Mathlib.Order.IsWellOrderLimitElement
38053804
import Mathlib.Order.Iterate
38063805
import Mathlib.Order.JordanHolder
38073806
import Mathlib.Order.KonigLemma

Mathlib/CategoryTheory/SmallObject/Iteration.lean

Lines changed: 35 additions & 27 deletions
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,8 @@ Authors: Joël Riou
55
-/
66
import Mathlib.CategoryTheory.Category.Preorder
77
import Mathlib.CategoryTheory.Limits.IsLimit
8-
import Mathlib.Order.IsWellOrderLimitElement
8+
import Mathlib.Order.ConditionallyCompleteLattice.Basic
9+
import Mathlib.Order.SuccPred.Limit
910

1011
/-! # Transfinite iterations of a functor
1112
@@ -37,19 +38,19 @@ namespace CategoryTheory
3738
open Category Limits
3839

3940
variable {C : Type*} [Category C] {Φ : C ⥤ C} (ε : 𝟭 C ⟶ Φ)
40-
{J : Type u} [LinearOrder J]
41+
{J : Type u} [Preorder J]
4142

4243
namespace Functor
4344

4445
namespace Iteration
4546

4647
variable {j : J} (F : { i // i ≤ j } ⥤ C)
4748

48-
/-- The map `F.obj ⟨i, _⟩ ⟶ F.obj ⟨wellOrderSucc i, _⟩` when `F : { i // i ≤ j } ⥤ C`
49+
/-- The map `F.obj ⟨i, _⟩ ⟶ F.obj ⟨Order.succ i, _⟩` when `F : { i // i ≤ j } ⥤ C`
4950
and `i : J` is such that `i < j`. -/
50-
noncomputable abbrev mapSucc' [IsWellOrder J (· < ·)] (i : J) (hi : i < j) :
51-
F.obj ⟨i, hi.le⟩ ⟶ F.obj ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ :=
52-
F.map (homOfLE (by simpa only [Subtype.mk_le_mk] using self_le_wellOrderSucc i))
51+
noncomputable abbrev mapSucc' [SuccOrder J] (i : J) (hi : i < j) :
52+
F.obj ⟨i, hi.le⟩ ⟶ F.obj ⟨Order.succ i, Order.succ_le_of_lt hi⟩ :=
53+
F.map <| homOfLE <| Subtype.mk_le_mk.2 <| Order.le_succ i
5354

5455
variable {i : J} (hi : i ≤ j)
5556

@@ -78,29 +79,26 @@ def coconeOfLE : Cocone (restrictionLT F hi) where
7879

7980
end Iteration
8081

81-
variable [IsWellOrder J (· < ·)] [OrderBot J]
82-
8382
/-- The category of `j`th iterations of a functor `Φ` equipped with a natural
8483
transformation `ε : 𝟭 C ⟶ Φ`. An object consists of the data of all iterations
8584
of `Φ` for `i : J` such that `i ≤ j` (this is the field `F`). Such objects are
8685
equipped with data and properties which characterizes the iterations up to a unique
8786
isomorphism for the three types of elements: `⊥`, successors, limit elements. -/
88-
structure Iteration (j : J) where
87+
structure Iteration [Preorder J] [OrderBot J] [SuccOrder J] (j : J) where
8988
/-- The data of all `i`th iterations for `i : J` such that `i ≤ j`. -/
9089
F : { i // i ≤ j } ⥤ C ⥤ C
9190
/-- The zeroth iteration is the identity functor. -/
9291
isoZero : F.obj ⟨⊥, bot_le⟩ ≅ 𝟭 C
9392
/-- The iteration on a successor element is obtained by composition of
9493
the previous iteration with `Φ`. -/
95-
isoSucc (i : J) (hi : i < j) :
96-
F.obj ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ ≅ F.obj ⟨i, hi.le⟩ ⋙ Φ
94+
isoSucc (i : J) (hi : i < j) : F.obj ⟨Order.succ i, Order.succ_le_of_lt hi⟩ ≅ F.obj ⟨i, hi.le⟩ ⋙ Φ
9795
/-- The natural map from an iteration to its successor is induced by `ε`. -/
9896
mapSucc'_eq (i : J) (hi : i < j) :
9997
Iteration.mapSucc' F i hi = whiskerLeft _ ε ≫ (isoSucc i hi).inv
10098
/-- If `i` is a limit element, the `i`th iteration is the colimit
10199
of `k`th iterations for `k < i`. -/
102-
isColimit (i : J) [IsWellOrderLimitElement i] (hi : i ≤ j) :
103-
IsColimit (Iteration.coconeOfLE F hi)
100+
isColimit (i : J) (hi : Order.IsSuccLimit i) (hij : i ≤ j) :
101+
IsColimit (Iteration.coconeOfLE F hij)
104102

105103
namespace Iteration
106104

@@ -109,12 +107,12 @@ variable {j : J}
109107

110108
section
111109

112-
variable (iter : Φ.Iteration ε j)
110+
variable [OrderBot J] [SuccOrder J] (iter : Φ.Iteration ε j)
113111

114112
/-- For `iter : Φ.Iteration.ε j`, this is the map
115-
`iter.F.obj ⟨i, _⟩ ⟶ iter.F.obj ⟨wellOrderSucc i, _⟩` if `i : J` is such that `i < j`. -/
113+
`iter.F.obj ⟨i, _⟩ ⟶ iter.F.obj ⟨Order.succ i, _⟩` if `i : J` is such that `i < j`. -/
116114
noncomputable abbrev mapSucc (i : J) (hi : i < j) :
117-
iter.F.obj ⟨i, hi.le⟩ ⟶ iter.F.obj ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ :=
115+
iter.F.obj ⟨i, hi.le⟩ ⟶ iter.F.obj ⟨Order.succ i, Order.succ_le_of_lt hi⟩ :=
118116
mapSucc' iter.F i hi
119117

120118
lemma mapSucc_eq (i : J) (hi : i < j) :
@@ -123,7 +121,7 @@ lemma mapSucc_eq (i : J) (hi : i < j) :
123121

124122
end
125123

126-
variable (iter₁ iter₂ : Φ.Iteration ε j)
124+
variable [OrderBot J] [SuccOrder J] (iter₁ iter₂ : Φ.Iteration ε j)
127125

128126
/-- A morphism between two objects `iter₁` and `iter₂` in the
129127
category `Φ.Iteration ε j` of `j`th iterations of a functor `Φ`
@@ -136,7 +134,7 @@ structure Hom where
136134
natTrans_app_zero :
137135
natTrans.app ⟨⊥, bot_le⟩ = iter₁.isoZero.hom ≫ iter₂.isoZero.inv := by aesop_cat
138136
natTrans_app_succ (i : J) (hi : i < j) :
139-
natTrans.app ⟨wellOrderSucc i, wellOrderSucc_le hi⟩ = (iter₁.isoSucc i hi).hom ≫
137+
natTrans.app ⟨Order.succ i, Order.succ_le_of_lt hi⟩ = (iter₁.isoSucc i hi).hom ≫
140138
whiskerRight (natTrans.app ⟨i, hi.le⟩) _ ≫ (iter₂.isoSucc i hi).inv := by aesop_cat
141139

142140
namespace Hom
@@ -172,17 +170,27 @@ instance : Category (Iteration ε j) where
172170
id := id
173171
comp := comp
174172

175-
instance : Subsingleton (iter₁ ⟶ iter₂) where
176-
allEq f g := ext' (by
177-
let P := fun (i : J) => ∀ (hi : i ≤ j), f.natTrans.app ⟨i, hi⟩ = g.natTrans.app ⟨i, hi⟩
178-
suffices ∀ (i : J), P i by
173+
instance {J} {j : J} [PartialOrder J] [OrderBot J] [WellFoundedLT J] [SuccOrder J]
174+
{iter₁ iter₂ : Iteration ε j} :
175+
Subsingleton (iter₁ ⟶ iter₂) where
176+
allEq f g := by
177+
apply ext'
178+
suffices ∀ i hi, f.natTrans.app ⟨i, hi⟩ = g.natTrans.app ⟨i, hi⟩ by
179179
ext ⟨i, hi⟩ : 2
180180
apply this
181-
refine fun _ => WellFoundedLT.induction _ (fun i hi hi' => ?_)
182-
obtain rfl|⟨i, rfl, hi''⟩|_ := eq_bot_or_eq_succ_or_isWellOrderLimitElement i
183-
· simp only [natTrans_app_zero]
184-
· simp only [Hom.natTrans_app_succ _ i (lt_of_lt_of_le hi'' hi'), hi i hi'']
185-
· exact (iter₁.isColimit i hi').hom_ext (fun ⟨k, hk⟩ => by simp [hi k hk]))
181+
intro i
182+
induction i using SuccOrder.limitRecOn with
183+
| hm j H =>
184+
have := H.eq_bot
185+
subst this
186+
simp [natTrans_app_zero]
187+
| hs j H IH =>
188+
intro hj
189+
simp [Hom.natTrans_app_succ, IH, (Order.lt_succ_of_not_isMax H).trans_le hj]
190+
| hl j H IH =>
191+
apply fun hj ↦ (iter₁.isColimit j H hj).hom_ext ?_
192+
rintro ⟨k, hk⟩
193+
simp [IH k hk]
186194

187195
end Hom
188196

Mathlib/Order/IsWellOrderLimitElement.lean

Lines changed: 0 additions & 98 deletions
This file was deleted.

0 commit comments

Comments
 (0)