| Intuitionistic Logic Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > ILE Home > Th. List > df-iress | GIF version | ||
| Description: Define a multifunction
restriction operator for extensible structures,
which can be used to turn statements about rings into statements about
subrings, modules into submodules, etc. This definition knows nothing
about individual structures and merely truncates the Base set while
leaving operators alone; individual kinds of structures will need to
handle this behavior, by ignoring operators' values outside the range,
defining a function using the base set and applying that, or explicitly
truncating the slot before use.
(Credit for this operator, as well as the 2023 modification for iset.mm, goes to Mario Carneiro.) (Contributed by Stefan O'Rear, 29-Nov-2014.) (Revised by Jim Kingdon, 7-Oct-2023.) |
| Ref | Expression |
|---|---|
| df-iress | ⊢ ↾s = (𝑤 ∈ V, 𝑥 ∈ V ↦ (𝑤 sSet 〈(Base‘ndx), (𝑥 ∩ (Base‘𝑤))〉)) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | cress 12704 | . 2 class ↾s | |
| 2 | vw | . . 3 setvar 𝑤 | |
| 3 | vx | . . 3 setvar 𝑥 | |
| 4 | cvv 2763 | . . 3 class V | |
| 5 | 2 | cv 1363 | . . . 4 class 𝑤 |
| 6 | cnx 12700 | . . . . . 6 class ndx | |
| 7 | cbs 12703 | . . . . . 6 class Base | |
| 8 | 6, 7 | cfv 5259 | . . . . 5 class (Base‘ndx) |
| 9 | 3 | cv 1363 | . . . . . 6 class 𝑥 |
| 10 | 5, 7 | cfv 5259 | . . . . . 6 class (Base‘𝑤) |
| 11 | 9, 10 | cin 3156 | . . . . 5 class (𝑥 ∩ (Base‘𝑤)) |
| 12 | 8, 11 | cop 3626 | . . . 4 class 〈(Base‘ndx), (𝑥 ∩ (Base‘𝑤))〉 |
| 13 | csts 12701 | . . . 4 class sSet | |
| 14 | 5, 12, 13 | co 5925 | . . 3 class (𝑤 sSet 〈(Base‘ndx), (𝑥 ∩ (Base‘𝑤))〉) |
| 15 | 2, 3, 4, 4, 14 | cmpo 5927 | . 2 class (𝑤 ∈ V, 𝑥 ∈ V ↦ (𝑤 sSet 〈(Base‘ndx), (𝑥 ∩ (Base‘𝑤))〉)) |
| 16 | 1, 15 | wceq 1364 | 1 wff ↾s = (𝑤 ∈ V, 𝑥 ∈ V ↦ (𝑤 sSet 〈(Base‘ndx), (𝑥 ∩ (Base‘𝑤))〉)) |
| Colors of variables: wff set class |
| This definition is referenced by: reldmress 12766 ressvalsets 12767 |
| Copyright terms: Public domain | W3C validator |