Metamath Proof Explorer


Theorem isdomn2

Description: A ring is a domain iff all nonzero elements are regular elements. (Contributed by Mario Carneiro, 28-Mar-2015) (Proof shortened by SN, 21-Jun-2025)

Ref Expression
Hypotheses isdomn2.b 𝐵 = ( Base ‘ 𝑅 )
isdomn2.t 𝐸 = ( RLReg ‘ 𝑅 )
isdomn2.z 0 = ( 0g𝑅 )
Assertion isdomn2 ( 𝑅 ∈ Domn ↔ ( 𝑅 ∈ NzRing ∧ ( 𝐵 ∖ { 0 } ) ⊆ 𝐸 ) )

Proof

Step Hyp Ref Expression
1 isdomn2.b 𝐵 = ( Base ‘ 𝑅 )
2 isdomn2.t 𝐸 = ( RLReg ‘ 𝑅 )
3 isdomn2.z 0 = ( 0g𝑅 )
4 eqid ( .r𝑅 ) = ( .r𝑅 )
5 1 4 3 isdomn ( 𝑅 ∈ Domn ↔ ( 𝑅 ∈ NzRing ∧ ∀ 𝑥𝐵𝑦𝐵 ( ( 𝑥 ( .r𝑅 ) 𝑦 ) = 0 → ( 𝑥 = 0𝑦 = 0 ) ) ) )
6 eldifi ( 𝑥 ∈ ( 𝐵 ∖ { 0 } ) → 𝑥𝐵 )
7 2 1 4 3 isrrg ( 𝑥𝐸 ↔ ( 𝑥𝐵 ∧ ∀ 𝑦𝐵 ( ( 𝑥 ( .r𝑅 ) 𝑦 ) = 0𝑦 = 0 ) ) )
8 7 baib ( 𝑥𝐵 → ( 𝑥𝐸 ↔ ∀ 𝑦𝐵 ( ( 𝑥 ( .r𝑅 ) 𝑦 ) = 0𝑦 = 0 ) ) )
9 6 8 syl ( 𝑥 ∈ ( 𝐵 ∖ { 0 } ) → ( 𝑥𝐸 ↔ ∀ 𝑦𝐵 ( ( 𝑥 ( .r𝑅 ) 𝑦 ) = 0𝑦 = 0 ) ) )
10 9 ralbiia ( ∀ 𝑥 ∈ ( 𝐵 ∖ { 0 } ) 𝑥𝐸 ↔ ∀ 𝑥 ∈ ( 𝐵 ∖ { 0 } ) ∀ 𝑦𝐵 ( ( 𝑥 ( .r𝑅 ) 𝑦 ) = 0𝑦 = 0 ) )
11 dfss3 ( ( 𝐵 ∖ { 0 } ) ⊆ 𝐸 ↔ ∀ 𝑥 ∈ ( 𝐵 ∖ { 0 } ) 𝑥𝐸 )
12 isdomn5 ( ∀ 𝑥𝐵𝑦𝐵 ( ( 𝑥 ( .r𝑅 ) 𝑦 ) = 0 → ( 𝑥 = 0𝑦 = 0 ) ) ↔ ∀ 𝑥 ∈ ( 𝐵 ∖ { 0 } ) ∀ 𝑦𝐵 ( ( 𝑥 ( .r𝑅 ) 𝑦 ) = 0𝑦 = 0 ) )
13 10 11 12 3bitr4ri ( ∀ 𝑥𝐵𝑦𝐵 ( ( 𝑥 ( .r𝑅 ) 𝑦 ) = 0 → ( 𝑥 = 0𝑦 = 0 ) ) ↔ ( 𝐵 ∖ { 0 } ) ⊆ 𝐸 )
14 13 anbi2i ( ( 𝑅 ∈ NzRing ∧ ∀ 𝑥𝐵𝑦𝐵 ( ( 𝑥 ( .r𝑅 ) 𝑦 ) = 0 → ( 𝑥 = 0𝑦 = 0 ) ) ) ↔ ( 𝑅 ∈ NzRing ∧ ( 𝐵 ∖ { 0 } ) ⊆ 𝐸 ) )
15 5 14 bitri ( 𝑅 ∈ Domn ↔ ( 𝑅 ∈ NzRing ∧ ( 𝐵 ∖ { 0 } ) ⊆ 𝐸 ) )