Metamath Proof Explorer


Theorem ordeq

Description: Equality theorem for the ordinal predicate. (Contributed by NM, 17-Sep-1993)

Ref Expression
Assertion ordeq ( 𝐴 = 𝐵 → ( Ord 𝐴 ↔ Ord 𝐵 ) )

Proof

Step Hyp Ref Expression
1 treq ( 𝐴 = 𝐵 → ( Tr 𝐴 ↔ Tr 𝐵 ) )
2 weeq2 ( 𝐴 = 𝐵 → ( E We 𝐴 ↔ E We 𝐵 ) )
3 1 2 anbi12d ( 𝐴 = 𝐵 → ( ( Tr 𝐴 ∧ E We 𝐴 ) ↔ ( Tr 𝐵 ∧ E We 𝐵 ) ) )
4 df-ord ( Ord 𝐴 ↔ ( Tr 𝐴 ∧ E We 𝐴 ) )
5 df-ord ( Ord 𝐵 ↔ ( Tr 𝐵 ∧ E We 𝐵 ) )
6 3 4 5 3bitr4g ( 𝐴 = 𝐵 → ( Ord 𝐴 ↔ Ord 𝐵 ) )