| Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > MPE Home > Th. List > df-subrg | Structured version Visualization version GIF version | ||
| 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 componentwise) 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.) |
| Ref | Expression |
|---|---|
| df-subrg | ⊢ SubRing = (𝑤 ∈ Ring ↦ {𝑠 ∈ 𝒫 (Base‘𝑤) ∣ ((𝑤 ↾s 𝑠) ∈ Ring ∧ (1r‘𝑤) ∈ 𝑠)}) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | csubrg 20569 | . 2 class SubRing | |
| 2 | vw | . . 3 setvar 𝑤 | |
| 3 | crg 20230 | . . 3 class Ring | |
| 4 | 2 | cv 1539 | . . . . . . 7 class 𝑤 |
| 5 | vs | . . . . . . . 8 setvar 𝑠 | |
| 6 | 5 | cv 1539 | . . . . . . 7 class 𝑠 |
| 7 | cress 17274 | . . . . . . 7 class ↾s | |
| 8 | 4, 6, 7 | co 7431 | . . . . . 6 class (𝑤 ↾s 𝑠) |
| 9 | 8, 3 | wcel 2108 | . . . . 5 wff (𝑤 ↾s 𝑠) ∈ Ring |
| 10 | cur 20178 | . . . . . . 7 class 1r | |
| 11 | 4, 10 | cfv 6561 | . . . . . 6 class (1r‘𝑤) |
| 12 | 11, 6 | wcel 2108 | . . . . 5 wff (1r‘𝑤) ∈ 𝑠 |
| 13 | 9, 12 | wa 395 | . . . 4 wff ((𝑤 ↾s 𝑠) ∈ Ring ∧ (1r‘𝑤) ∈ 𝑠) |
| 14 | cbs 17247 | . . . . . 6 class Base | |
| 15 | 4, 14 | cfv 6561 | . . . . 5 class (Base‘𝑤) |
| 16 | 15 | cpw 4600 | . . . 4 class 𝒫 (Base‘𝑤) |
| 17 | 13, 5, 16 | crab 3436 | . . 3 class {𝑠 ∈ 𝒫 (Base‘𝑤) ∣ ((𝑤 ↾s 𝑠) ∈ Ring ∧ (1r‘𝑤) ∈ 𝑠)} |
| 18 | 2, 3, 17 | cmpt 5225 | . 2 class (𝑤 ∈ Ring ↦ {𝑠 ∈ 𝒫 (Base‘𝑤) ∣ ((𝑤 ↾s 𝑠) ∈ Ring ∧ (1r‘𝑤) ∈ 𝑠)}) |
| 19 | 1, 18 | wceq 1540 | 1 wff SubRing = (𝑤 ∈ Ring ↦ {𝑠 ∈ 𝒫 (Base‘𝑤) ∣ ((𝑤 ↾s 𝑠) ∈ Ring ∧ (1r‘𝑤) ∈ 𝑠)}) |
| Colors of variables: wff setvar class |
| This definition is referenced by: issubrg 20571 |
| Copyright terms: Public domain | W3C validator |