MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  df-subrg Structured version   Visualization version   GIF version

Definition df-subrg 18601
Description: Define a subring of a ring as a set of elements that is a ring in its own right and contains the multiplicative identity.

The additional constraint is necessary because the multiplicative identity of a ring, unlike the additive identity of a ring/group or the multiplicative identity of a field, cannot be identified by a local property. Thus, it is possible for a subset of a ring to be a ring while not containing the true identity if it contains a false identity. For instance, the subset (ℤ × {0}) of (ℤ × ℤ) (where multiplication is component-wise) contains the false identity ⟨1, 0⟩ which preserves every element of the subset and thus appears to be the identity of the subset, but is not the identity of the larger ring. (Contributed by Stefan O'Rear, 27-Nov-2014.)

Assertion
Ref Expression
df-subrg SubRing = (𝑤 ∈ Ring ↦ {𝑠 ∈ 𝒫 (Base‘𝑤) ∣ ((𝑤s 𝑠) ∈ Ring ∧ (1r𝑤) ∈ 𝑠)})
Distinct variable group:   𝑤,𝑠

Detailed syntax breakdown of Definition df-subrg
StepHypRef Expression
1 csubrg 18599 . 2 class SubRing
2 vw . . 3 setvar 𝑤
3 crg 18370 . . 3 class Ring
42cv 1474 . . . . . . 7 class 𝑤
5 vs . . . . . . . 8 setvar 𝑠
65cv 1474 . . . . . . 7 class 𝑠
7 cress 15696 . . . . . . 7 class s
84, 6, 7co 6549 . . . . . 6 class (𝑤s 𝑠)
98, 3wcel 1977 . . . . 5 wff (𝑤s 𝑠) ∈ Ring
10 cur 18324 . . . . . . 7 class 1r
114, 10cfv 5804 . . . . . 6 class (1r𝑤)
1211, 6wcel 1977 . . . . 5 wff (1r𝑤) ∈ 𝑠
139, 12wa 383 . . . 4 wff ((𝑤s 𝑠) ∈ Ring ∧ (1r𝑤) ∈ 𝑠)
14 cbs 15695 . . . . . 6 class Base
154, 14cfv 5804 . . . . 5 class (Base‘𝑤)
1615cpw 4108 . . . 4 class 𝒫 (Base‘𝑤)
1713, 5, 16crab 2900 . . 3 class {𝑠 ∈ 𝒫 (Base‘𝑤) ∣ ((𝑤s 𝑠) ∈ Ring ∧ (1r𝑤) ∈ 𝑠)}
182, 3, 17cmpt 4643 . 2 class (𝑤 ∈ Ring ↦ {𝑠 ∈ 𝒫 (Base‘𝑤) ∣ ((𝑤s 𝑠) ∈ Ring ∧ (1r𝑤) ∈ 𝑠)})
191, 18wceq 1475 1 wff SubRing = (𝑤 ∈ Ring ↦ {𝑠 ∈ 𝒫 (Base‘𝑤) ∣ ((𝑤s 𝑠) ∈ Ring ∧ (1r𝑤) ∈ 𝑠)})
Colors of variables: wff setvar class
This definition is referenced by:  issubrg  18603
  Copyright terms: Public domain W3C validator