@@ -128,7 +128,7 @@ attribute [instance] WellOrder.wo
128
128
129
129
namespace WellOrder
130
130
131
- instance : Inhabited WellOrder :=
131
+ instance inhabited : Inhabited WellOrder :=
132
132
⟨⟨PEmpty, _, inferInstanceAs (IsWellOrder PEmpty EmptyRelation)⟩⟩
133
133
134
134
@[simp]
@@ -174,13 +174,13 @@ def type (r : α → α → Prop) [wo : IsWellOrder α r] : Ordinal :=
174
174
⟦⟨α, r, wo⟩⟧
175
175
#align ordinal.type Ordinal.type
176
176
177
- instance : Zero Ordinal :=
177
+ instance hasZero : Zero Ordinal :=
178
178
⟨type <| @EmptyRelation PEmpty⟩
179
179
180
- instance : Inhabited Ordinal :=
180
+ instance inhabited : Inhabited Ordinal :=
181
181
⟨0 ⟩
182
182
183
- instance : One Ordinal :=
183
+ instance hasOne : One Ordinal :=
184
184
⟨type <| @EmptyRelation PUnit⟩
185
185
186
186
/-- 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 :=
293
293
type_ne_zero_of_nonempty _
294
294
#align ordinal.one_ne_zero Ordinal.one_ne_zero
295
295
296
- instance : Nontrivial Ordinal.{u} :=
296
+ instance nontrivial : Nontrivial Ordinal.{u} :=
297
297
⟨⟨1 , 0 , Ordinal.one_ne_zero⟩⟩
298
298
299
299
--@[ simp ] -- Porting note: not in simp nf, added aux lemma below
@@ -316,7 +316,7 @@ theorem inductionOn {C : Ordinal → Prop} (o : Ordinal)
316
316
/-! ### The order on ordinals -/
317
317
318
318
319
- instance : PartialOrder Ordinal
319
+ instance partialOrder : PartialOrder Ordinal
320
320
where
321
321
le a b :=
322
322
Quotient.liftOn₂ a b (fun ⟨_, r, _⟩ ⟨_, s, _⟩ => Nonempty (r ≼i s))
@@ -387,7 +387,7 @@ protected theorem zero_le (o : Ordinal) : 0 ≤ o :=
387
387
inductionOn o fun _ r _ => (InitialSeg.ofIsEmpty _ r).ordinal_type_le
388
388
#align ordinal.zero_le Ordinal.zero_le
389
389
390
- instance : OrderBot Ordinal where
390
+ instance orderBot : OrderBot Ordinal where
391
391
bot := 0
392
392
bot_le := Ordinal.zero_le
393
393
@@ -413,7 +413,7 @@ theorem eq_zero_or_pos : ∀ a : Ordinal, a = 0 ∨ 0 < a :=
413
413
eq_bot_or_bot_lt
414
414
#align ordinal.eq_zero_or_pos Ordinal.eq_zero_or_pos
415
415
416
- instance : ZeroLEOneClass Ordinal :=
416
+ instance zeroLEOneClass : ZeroLEOneClass Ordinal :=
417
417
⟨Ordinal.zero_le _⟩
418
418
419
419
instance NeZero.one : NeZero (1 : Ordinal) :=
@@ -574,7 +574,7 @@ theorem lt_wf : @WellFounded Ordinal (· < ·) :=
574
574
exact IH _ ((typein_lt_typein r).1 h)⟩⟩
575
575
#align ordinal.lt_wf Ordinal.lt_wf
576
576
577
- instance : WellFoundedRelation Ordinal :=
577
+ instance wellFoundedRelation : WellFoundedRelation Ordinal :=
578
578
⟨(· < ·), lt_wf⟩
579
579
580
580
/-- 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
885
885
886
886
/-- `o₁ + o₂` is the order on the disjoint union of `o₁` and `o₂` obtained by declaring that
887
887
every element of `o₁` is smaller than every element of `o₂`. -/
888
- instance : Add Ordinal.{u} :=
888
+ instance hasAdd : Add Ordinal.{u} :=
889
889
⟨fun o₁ o₂ =>
890
890
Quotient.liftOn₂ o₁ o₂ (fun ⟨_, r, _⟩ ⟨_, s, _⟩ => type (Sum.Lex r s))
891
891
fun _ _ _ _ ⟨f⟩ ⟨g⟩ => Quot.sound ⟨RelIso.sumLexCongr f g⟩⟩
892
892
893
- instance : AddMonoidWithOne Ordinal.{u}
893
+ instance addMonoidWithOne : AddMonoidWithOne Ordinal.{u}
894
894
where
895
895
add := (· + ·)
896
896
zero := 0
@@ -977,7 +977,7 @@ theorem le_add_left (a b : Ordinal) : a ≤ b + a := by
977
977
simpa only [zero_add] using add_le_add_right (Ordinal.zero_le b) a
978
978
#align ordinal.le_add_left Ordinal.le_add_left
979
979
980
- instance : LinearOrder Ordinal :=
980
+ instance linearOrder : LinearOrder Ordinal :=
981
981
{inferInstanceAs (PartialOrder Ordinal) with
982
982
le_total := fun a b =>
983
983
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 :=
999
999
exact Or.inr (Or.inl h)]
1000
1000
decidable_le := Classical.decRel _ }
1001
1001
1002
- instance : WellFoundedLT Ordinal :=
1002
+ instance wellFoundedLT : WellFoundedLT Ordinal :=
1003
1003
⟨lt_wf⟩
1004
1004
1005
- instance : IsWellOrder Ordinal (· < ·) where
1005
+ instance isWellOrder : IsWellOrder Ordinal (· < ·) where
1006
1006
1007
1007
instance : ConditionallyCompleteLinearOrderBot Ordinal :=
1008
1008
IsWellOrder.conditionallyCompleteLinearOrderBot _
@@ -1055,10 +1055,10 @@ private theorem succ_le_iff' {a b : Ordinal} : a + 1 ≤ b ↔ a < b :=
1055
1055
cases' (hf b).1 h with w h
1056
1056
exact ⟨Sum.inl w, h⟩⟩
1057
1057
1058
- instance : NoMaxOrder Ordinal :=
1058
+ instance noMaxOrder : NoMaxOrder Ordinal :=
1059
1059
⟨fun _ => ⟨_, succ_le_iff'.1 le_rfl⟩⟩
1060
1060
1061
- instance : SuccOrder Ordinal.{u} :=
1061
+ instance succOrder : SuccOrder Ordinal.{u} :=
1062
1062
SuccOrder.ofSuccLeIff (fun o => o + 1 ) succ_le_iff'
1063
1063
1064
1064
@[simp]
0 commit comments