Skip to content

Commit 2367072

Browse files
arienmalecLukasMias
andcommitted
Port/SetTheory.Ordinal.NaturalOps (#2353)
Co-authored-by: Lukas Miaskiwskyi <lukas.mias@gmail.com>
1 parent 7e1713a commit 2367072

File tree

4 files changed

+477
-25
lines changed

4 files changed

+477
-25
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -975,6 +975,7 @@ import Mathlib.SetTheory.Cardinal.SchroederBernstein
975975
import Mathlib.SetTheory.Lists
976976
import Mathlib.SetTheory.Ordinal.Arithmetic
977977
import Mathlib.SetTheory.Ordinal.Basic
978+
import Mathlib.SetTheory.Ordinal.NaturalOps
978979
import Mathlib.SetTheory.Ordinal.Topology
979980
import Mathlib.Tactic.Abel
980981
import Mathlib.Tactic.Alias

Mathlib/SetTheory/Ordinal/Arithmetic.lean

Lines changed: 9 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -537,7 +537,7 @@ theorem sub_nonempty {a b : Ordinal} : { o | a ≤ b + o }.Nonempty :=
537537
#align ordinal.sub_nonempty Ordinal.sub_nonempty
538538

539539
/-- `a - b` is the unique ordinal satisfying `b + (a - b) = a` when `b ≤ a`. -/
540-
instance : Sub Ordinal :=
540+
instance hasSub : Sub Ordinal :=
541541
fun a b => infₛ { o | a ≤ b + o }⟩
542542

543543
theorem le_add_sub (a b : Ordinal) : a ≤ b + (a - b) :=
@@ -582,7 +582,7 @@ theorem sub_lt_of_le {a b c : Ordinal} (h : b ≤ a) : a - b < c ↔ a < b + c :
582582
lt_iff_lt_of_le_iff_le (le_sub_of_le h)
583583
#align ordinal.sub_lt_of_le Ordinal.sub_lt_of_le
584584

585-
instance : ExistsAddOfLE Ordinal :=
585+
instance existsAddOfLE : ExistsAddOfLE Ordinal :=
586586
fun h => ⟨_, (Ordinal.add_sub_cancel_of_le h).symm⟩⟩
587587

588588
@[simp]
@@ -685,13 +685,13 @@ private theorem mul_eq_zero' {a b : Ordinal} : a * b = 0 ↔ a = 0 ∨ b = 0 :=
685685
rw [or_comm]
686686
exact isEmpty_prod
687687

688-
instance : MonoidWithZero Ordinal :=
688+
instance monoidWithZero : MonoidWithZero Ordinal :=
689689
{ Ordinal.monoid with
690690
zero := 0
691691
mul_zero := fun _a => mul_eq_zero'.2 <| Or.inr rfl
692692
zero_mul := fun _a => mul_eq_zero'.2 <| Or.inl rfl }
693693

694-
instance : NoZeroDivisors Ordinal :=
694+
instance noZeroDivisors : NoZeroDivisors Ordinal :=
695695
fun {_ _} => mul_eq_zero'.1
696696

697697
@[simp]
@@ -708,7 +708,7 @@ theorem card_mul (a b) : card (a * b) = card a * card b :=
708708
Quotient.inductionOn₂ a b fun ⟨α, _r, _⟩ ⟨β, _s, _⟩ => mul_comm (#β) (#α)
709709
#align ordinal.card_mul Ordinal.card_mul
710710

711-
instance : LeftDistribClass Ordinal.{u} :=
711+
instance leftDistribClass : LeftDistribClass Ordinal.{u} :=
712712
fun a b c =>
713713
Quotient.inductionOn₃ a b c fun ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨γ, t, _⟩ =>
714714
Quotient.sound
@@ -875,7 +875,7 @@ theorem div_nonempty {a b : Ordinal} (h : b ≠ 0) : { o | a < b * succ o }.None
875875
#align ordinal.div_nonempty Ordinal.div_nonempty
876876

877877
/-- `a / b` is the unique ordinal `o` satisfying `a = b * o + o'` with `o' < b`. -/
878-
instance : Div Ordinal :=
878+
instance hasDiv : Div Ordinal :=
879879
fun a b => if _h : b = 0 then 0 else infₛ { o | a < b * succ o }⟩
880880

881881
@[simp]
@@ -1019,12 +1019,12 @@ theorem dvd_antisymm {a b : Ordinal} (h₁ : a ∣ b) (h₂ : b ∣ a) : a = b :
10191019
else (le_of_dvd b0 h₁).antisymm (le_of_dvd a0 h₂)
10201020
#align ordinal.dvd_antisymm Ordinal.dvd_antisymm
10211021

1022-
instance : IsAntisymm Ordinal (· ∣ ·) :=
1022+
instance isAntisymm : IsAntisymm Ordinal (· ∣ ·) :=
10231023
⟨@dvd_antisymm⟩
10241024

10251025
/-- `a % b` is the unique ordinal `o'` satisfying
10261026
`a = b * o + o'` with `o' < b`. -/
1027-
instance : Mod Ordinal :=
1027+
instance hasMod : Mod Ordinal :=
10281028
fun a b => a - b * (a / b)⟩
10291029

10301030
theorem mod_def (a b : Ordinal) : a % b = a - b * (a / b) :=
@@ -2268,7 +2268,7 @@ end
22682268

22692269

22702270
/-- The ordinal exponential, defined by transfinite recursion. -/
2271-
instance : Pow Ordinal Ordinal :=
2271+
instance hasPow: Pow Ordinal Ordinal :=
22722272
fun a b => if a = 0 then 1 - b else limitRecOn b 1 (fun _ IH => IH * a) fun b _ => bsup.{u, u} b⟩
22732273

22742274
-- Porting note: Ambiguous notations.

Mathlib/SetTheory/Ordinal/Basic.lean

Lines changed: 16 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -128,7 +128,7 @@ attribute [instance] WellOrder.wo
128128

129129
namespace WellOrder
130130

131-
instance : Inhabited WellOrder :=
131+
instance inhabited : Inhabited WellOrder :=
132132
⟨⟨PEmpty, _, inferInstanceAs (IsWellOrder PEmpty EmptyRelation)⟩⟩
133133

134134
@[simp]
@@ -174,13 +174,13 @@ def type (r : α → α → Prop) [wo : IsWellOrder α r] : Ordinal :=
174174
⟦⟨α, r, wo⟩⟧
175175
#align ordinal.type Ordinal.type
176176

177-
instance : Zero Ordinal :=
177+
instance hasZero : Zero Ordinal :=
178178
⟨type <| @EmptyRelation PEmpty⟩
179179

180-
instance : Inhabited Ordinal :=
180+
instance inhabited : Inhabited Ordinal :=
181181
0
182182

183-
instance : One Ordinal :=
183+
instance hasOne : One Ordinal :=
184184
⟨type <| @EmptyRelation PUnit⟩
185185

186186
/-- The order type of an element inside a well order. For the embedding as a principal segment, see
@@ -293,7 +293,7 @@ protected theorem one_ne_zero : (1 : Ordinal) ≠ 0 :=
293293
type_ne_zero_of_nonempty _
294294
#align ordinal.one_ne_zero Ordinal.one_ne_zero
295295

296-
instance : Nontrivial Ordinal.{u} :=
296+
instance nontrivial : Nontrivial Ordinal.{u} :=
297297
⟨⟨1, 0, Ordinal.one_ne_zero⟩⟩
298298

299299
--@[simp] -- Porting note: not in simp nf, added aux lemma below
@@ -316,7 +316,7 @@ theorem inductionOn {C : Ordinal → Prop} (o : Ordinal)
316316
/-! ### The order on ordinals -/
317317

318318

319-
instance : PartialOrder Ordinal
319+
instance partialOrder : PartialOrder Ordinal
320320
where
321321
le a b :=
322322
Quotient.liftOn₂ a b (fun ⟨_, r, _⟩ ⟨_, s, _⟩ => Nonempty (r ≼i s))
@@ -387,7 +387,7 @@ protected theorem zero_le (o : Ordinal) : 0 ≤ o :=
387387
inductionOn o fun _ r _ => (InitialSeg.ofIsEmpty _ r).ordinal_type_le
388388
#align ordinal.zero_le Ordinal.zero_le
389389

390-
instance : OrderBot Ordinal where
390+
instance orderBot : OrderBot Ordinal where
391391
bot := 0
392392
bot_le := Ordinal.zero_le
393393

@@ -413,7 +413,7 @@ theorem eq_zero_or_pos : ∀ a : Ordinal, a = 0 ∨ 0 < a :=
413413
eq_bot_or_bot_lt
414414
#align ordinal.eq_zero_or_pos Ordinal.eq_zero_or_pos
415415

416-
instance : ZeroLEOneClass Ordinal :=
416+
instance zeroLEOneClass : ZeroLEOneClass Ordinal :=
417417
⟨Ordinal.zero_le _⟩
418418

419419
instance NeZero.one : NeZero (1 : Ordinal) :=
@@ -574,7 +574,7 @@ theorem lt_wf : @WellFounded Ordinal (· < ·) :=
574574
exact IH _ ((typein_lt_typein r).1 h)⟩⟩
575575
#align ordinal.lt_wf Ordinal.lt_wf
576576

577-
instance : WellFoundedRelation Ordinal :=
577+
instance wellFoundedRelation : WellFoundedRelation Ordinal :=
578578
⟨(· < ·), lt_wf⟩
579579

580580
/-- Reformulation of well founded induction on ordinals as a lemma that works with the
@@ -885,12 +885,12 @@ the addition, together with properties of the other operations, are proved in
885885

886886
/-- `o₁ + o₂` is the order on the disjoint union of `o₁` and `o₂` obtained by declaring that
887887
every element of `o₁` is smaller than every element of `o₂`. -/
888-
instance : Add Ordinal.{u} :=
888+
instance hasAdd : Add Ordinal.{u} :=
889889
fun o₁ o₂ =>
890890
Quotient.liftOn₂ o₁ o₂ (fun ⟨_, r, _⟩ ⟨_, s, _⟩ => type (Sum.Lex r s))
891891
fun _ _ _ _ ⟨f⟩ ⟨g⟩ => Quot.sound ⟨RelIso.sumLexCongr f g⟩⟩
892892

893-
instance : AddMonoidWithOne Ordinal.{u}
893+
instance addMonoidWithOne : AddMonoidWithOne Ordinal.{u}
894894
where
895895
add := (· + ·)
896896
zero := 0
@@ -977,7 +977,7 @@ theorem le_add_left (a b : Ordinal) : a ≤ b + a := by
977977
simpa only [zero_add] using add_le_add_right (Ordinal.zero_le b) a
978978
#align ordinal.le_add_left Ordinal.le_add_left
979979

980-
instance : LinearOrder Ordinal :=
980+
instance linearOrder : LinearOrder Ordinal :=
981981
{inferInstanceAs (PartialOrder Ordinal) with
982982
le_total := fun a b =>
983983
match lt_or_eq_of_le (le_add_left b a), lt_or_eq_of_le (le_add_right a b) with
@@ -999,10 +999,10 @@ instance : LinearOrder Ordinal :=
999999
exact Or.inr (Or.inl h)]
10001000
decidable_le := Classical.decRel _ }
10011001

1002-
instance : WellFoundedLT Ordinal :=
1002+
instance wellFoundedLT : WellFoundedLT Ordinal :=
10031003
⟨lt_wf⟩
10041004

1005-
instance : IsWellOrder Ordinal (· < ·) where
1005+
instance isWellOrder : IsWellOrder Ordinal (· < ·) where
10061006

10071007
instance : ConditionallyCompleteLinearOrderBot Ordinal :=
10081008
IsWellOrder.conditionallyCompleteLinearOrderBot _
@@ -1055,10 +1055,10 @@ private theorem succ_le_iff' {a b : Ordinal} : a + 1 ≤ b ↔ a < b :=
10551055
cases' (hf b).1 h with w h
10561056
exact ⟨Sum.inl w, h⟩⟩
10571057

1058-
instance : NoMaxOrder Ordinal :=
1058+
instance noMaxOrder : NoMaxOrder Ordinal :=
10591059
fun _ => ⟨_, succ_le_iff'.1 le_rfl⟩⟩
10601060

1061-
instance : SuccOrder Ordinal.{u} :=
1061+
instance succOrder : SuccOrder Ordinal.{u} :=
10621062
SuccOrder.ofSuccLeIff (fun o => o + 1) succ_le_iff'
10631063

10641064
@[simp]

0 commit comments

Comments
 (0)