Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
Mirrors > Home > MPE Home > Th. List > abnexg | Structured version Visualization version GIF version |
Description: Sufficient condition for a class abstraction to be a proper class. The class 𝐹 can be thought of as an expression in 𝑥 and the abstraction appearing in the statement as the class of values 𝐹 as 𝑥 varies through 𝐴. Assuming the antecedents, if that class is a set, then so is the "domain" 𝐴. The converse holds without antecedent, see abrexexg 7797. Note that the second antecedent ∀𝑥 ∈ 𝐴𝑥 ∈ 𝐹 cannot be translated to 𝐴 ⊆ 𝐹 since 𝐹 may depend on 𝑥. In applications, one may take 𝐹 = {𝑥} or 𝐹 = 𝒫 𝑥 (see snnex 7602 and pwnex 7603 respectively, proved from abnex 7601, which is a consequence of abnexg 7600 with 𝐴 = V). (Contributed by BJ, 2-Dec-2021.) |
Ref | Expression |
---|---|
abnexg | ⊢ (∀𝑥 ∈ 𝐴 (𝐹 ∈ 𝑉 ∧ 𝑥 ∈ 𝐹) → ({𝑦 ∣ ∃𝑥 ∈ 𝐴 𝑦 = 𝐹} ∈ 𝑊 → 𝐴 ∈ V)) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | uniexg 7587 | . 2 ⊢ ({𝑦 ∣ ∃𝑥 ∈ 𝐴 𝑦 = 𝐹} ∈ 𝑊 → ∪ {𝑦 ∣ ∃𝑥 ∈ 𝐴 𝑦 = 𝐹} ∈ V) | |
2 | simpl 483 | . . . . 5 ⊢ ((𝐹 ∈ 𝑉 ∧ 𝑥 ∈ 𝐹) → 𝐹 ∈ 𝑉) | |
3 | 2 | ralimi 3089 | . . . 4 ⊢ (∀𝑥 ∈ 𝐴 (𝐹 ∈ 𝑉 ∧ 𝑥 ∈ 𝐹) → ∀𝑥 ∈ 𝐴 𝐹 ∈ 𝑉) |
4 | dfiun2g 4966 | . . . . . 6 ⊢ (∀𝑥 ∈ 𝐴 𝐹 ∈ 𝑉 → ∪ 𝑥 ∈ 𝐴 𝐹 = ∪ {𝑦 ∣ ∃𝑥 ∈ 𝐴 𝑦 = 𝐹}) | |
5 | 4 | eleq1d 2825 | . . . . 5 ⊢ (∀𝑥 ∈ 𝐴 𝐹 ∈ 𝑉 → (∪ 𝑥 ∈ 𝐴 𝐹 ∈ V ↔ ∪ {𝑦 ∣ ∃𝑥 ∈ 𝐴 𝑦 = 𝐹} ∈ V)) |
6 | 5 | biimprd 247 | . . . 4 ⊢ (∀𝑥 ∈ 𝐴 𝐹 ∈ 𝑉 → (∪ {𝑦 ∣ ∃𝑥 ∈ 𝐴 𝑦 = 𝐹} ∈ V → ∪ 𝑥 ∈ 𝐴 𝐹 ∈ V)) |
7 | 3, 6 | syl 17 | . . 3 ⊢ (∀𝑥 ∈ 𝐴 (𝐹 ∈ 𝑉 ∧ 𝑥 ∈ 𝐹) → (∪ {𝑦 ∣ ∃𝑥 ∈ 𝐴 𝑦 = 𝐹} ∈ V → ∪ 𝑥 ∈ 𝐴 𝐹 ∈ V)) |
8 | simpr 485 | . . . . 5 ⊢ ((𝐹 ∈ 𝑉 ∧ 𝑥 ∈ 𝐹) → 𝑥 ∈ 𝐹) | |
9 | 8 | ralimi 3089 | . . . 4 ⊢ (∀𝑥 ∈ 𝐴 (𝐹 ∈ 𝑉 ∧ 𝑥 ∈ 𝐹) → ∀𝑥 ∈ 𝐴 𝑥 ∈ 𝐹) |
10 | iunid 4995 | . . . . 5 ⊢ ∪ 𝑥 ∈ 𝐴 {𝑥} = 𝐴 | |
11 | snssi 4747 | . . . . . . 7 ⊢ (𝑥 ∈ 𝐹 → {𝑥} ⊆ 𝐹) | |
12 | 11 | ralimi 3089 | . . . . . 6 ⊢ (∀𝑥 ∈ 𝐴 𝑥 ∈ 𝐹 → ∀𝑥 ∈ 𝐴 {𝑥} ⊆ 𝐹) |
13 | ss2iun 4948 | . . . . . 6 ⊢ (∀𝑥 ∈ 𝐴 {𝑥} ⊆ 𝐹 → ∪ 𝑥 ∈ 𝐴 {𝑥} ⊆ ∪ 𝑥 ∈ 𝐴 𝐹) | |
14 | 12, 13 | syl 17 | . . . . 5 ⊢ (∀𝑥 ∈ 𝐴 𝑥 ∈ 𝐹 → ∪ 𝑥 ∈ 𝐴 {𝑥} ⊆ ∪ 𝑥 ∈ 𝐴 𝐹) |
15 | 10, 14 | eqsstrrid 3975 | . . . 4 ⊢ (∀𝑥 ∈ 𝐴 𝑥 ∈ 𝐹 → 𝐴 ⊆ ∪ 𝑥 ∈ 𝐴 𝐹) |
16 | ssexg 5251 | . . . . 5 ⊢ ((𝐴 ⊆ ∪ 𝑥 ∈ 𝐴 𝐹 ∧ ∪ 𝑥 ∈ 𝐴 𝐹 ∈ V) → 𝐴 ∈ V) | |
17 | 16 | ex 413 | . . . 4 ⊢ (𝐴 ⊆ ∪ 𝑥 ∈ 𝐴 𝐹 → (∪ 𝑥 ∈ 𝐴 𝐹 ∈ V → 𝐴 ∈ V)) |
18 | 9, 15, 17 | 3syl 18 | . . 3 ⊢ (∀𝑥 ∈ 𝐴 (𝐹 ∈ 𝑉 ∧ 𝑥 ∈ 𝐹) → (∪ 𝑥 ∈ 𝐴 𝐹 ∈ V → 𝐴 ∈ V)) |
19 | 7, 18 | syld 47 | . 2 ⊢ (∀𝑥 ∈ 𝐴 (𝐹 ∈ 𝑉 ∧ 𝑥 ∈ 𝐹) → (∪ {𝑦 ∣ ∃𝑥 ∈ 𝐴 𝑦 = 𝐹} ∈ V → 𝐴 ∈ V)) |
20 | 1, 19 | syl5 34 | 1 ⊢ (∀𝑥 ∈ 𝐴 (𝐹 ∈ 𝑉 ∧ 𝑥 ∈ 𝐹) → ({𝑦 ∣ ∃𝑥 ∈ 𝐴 𝑦 = 𝐹} ∈ 𝑊 → 𝐴 ∈ V)) |
Colors of variables: wff setvar class |
Syntax hints: → wi 4 ∧ wa 396 = wceq 1542 ∈ wcel 2110 {cab 2717 ∀wral 3066 ∃wrex 3067 Vcvv 3431 ⊆ wss 3892 {csn 4567 ∪ cuni 4845 ∪ ciun 4930 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1802 ax-4 1816 ax-5 1917 ax-6 1975 ax-7 2015 ax-8 2112 ax-9 2120 ax-10 2141 ax-11 2158 ax-12 2175 ax-ext 2711 ax-sep 5227 ax-un 7582 |
This theorem depends on definitions: df-bi 206 df-an 397 df-or 845 df-tru 1545 df-ex 1787 df-nf 1791 df-sb 2072 df-clab 2718 df-cleq 2732 df-clel 2818 df-nfc 2891 df-ral 3071 df-rex 3072 df-rab 3075 df-v 3433 df-in 3899 df-ss 3909 df-sn 4568 df-uni 4846 df-iun 4932 |
This theorem is referenced by: abnex 7601 |
Copyright terms: Public domain | W3C validator |