Skip to content

Commit

Permalink
chore(set_theory/ordinal/basic): Small style tweaks (#13561)
Browse files Browse the repository at this point in the history
  • Loading branch information
vihdzp committed Apr 21, 2022
1 parent efab188 commit aeef727
Showing 1 changed file with 4 additions and 4 deletions.
8 changes: 4 additions & 4 deletions src/set_theory/ordinal/basic.lean
Expand Up @@ -467,9 +467,9 @@ end Well_order
isomorphism. -/
instance ordinal.is_equivalent : setoid Well_order :=
{ r := λ ⟨α, r, wo⟩ ⟨β, s, wo'⟩, nonempty (r ≃r s),
iseqv := ⟨λ⟨α, r, _⟩, ⟨rel_iso.refl _⟩,
λ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨e⟩, ⟨e.symm⟩,
λ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨γ, t, _⟩ ⟨e₁⟩ ⟨e₂⟩, ⟨e₁.trans e₂⟩⟩ }
iseqv := ⟨λ ⟨α, r, _⟩, ⟨rel_iso.refl _⟩,
λ ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨e⟩, ⟨e.symm⟩,
λ ⟨α, r, _⟩ ⟨β, s, _⟩ ⟨γ, t, _⟩ ⟨e₁⟩ ⟨e₂⟩, ⟨e₁.trans e₂⟩⟩ }

/-- `ordinal.{u}` is the type of well orders in `Type u`, up to order isomorphism. -/
def ordinal : Type (u + 1) := quotient ordinal.is_equivalent
Expand Down Expand Up @@ -503,7 +503,7 @@ theorem type_eq {α β} {r : α → α → Prop} {s : β → β → Prop}
[is_well_order α r] [is_well_order β s] :
type r = type s ↔ nonempty (r ≃r s) := quotient.eq

@[simp] lemma type_lt (o : ordinal) : type ((<) : o.out.α → o.out.α → Prop) = o :=
@[simp] theorem type_lt (o : ordinal) : type ((<) : o.out.α → o.out.α → Prop) = o :=
begin
change type o.out.r = _,
refine eq.trans _ (quotient.out_eq o),
Expand Down

0 comments on commit aeef727

Please sign in to comment.