Home | Intuitionistic Logic Explorer Theorem List (p. 34 of 102) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > ILE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | inundifss 3301 | The intersection and class difference of a class with another class are contained in the original class. In classical logic we'd be able to make a stronger statement: that everything in the original class is in the intersection or the difference (that is, this theorem would be equality rather than subset). (Contributed by Jim Kingdon, 4-Aug-2018.) |
⊢ ((𝐴 ∩ 𝐵) ∪ (𝐴 ∖ 𝐵)) ⊆ 𝐴 | ||
Theorem | difun2 3302 | Absorption of union by difference. Theorem 36 of [Suppes] p. 29. (Contributed by NM, 19-May-1998.) |
⊢ ((𝐴 ∪ 𝐵) ∖ 𝐵) = (𝐴 ∖ 𝐵) | ||
Theorem | undifss 3303 | Union of complementary parts into whole. (Contributed by Jim Kingdon, 4-Aug-2018.) |
⊢ (𝐴 ⊆ 𝐵 ↔ (𝐴 ∪ (𝐵 ∖ 𝐴)) ⊆ 𝐵) | ||
Theorem | ssdifin0 3304 | A subset of a difference does not intersect the subtrahend. (Contributed by Jeff Hankins, 1-Sep-2013.) (Proof shortened by Mario Carneiro, 24-Aug-2015.) |
⊢ (𝐴 ⊆ (𝐵 ∖ 𝐶) → (𝐴 ∩ 𝐶) = ∅) | ||
Theorem | ssdifeq0 3305 | A class is a subclass of itself subtracted from another iff it is the empty set. (Contributed by Steve Rodriguez, 20-Nov-2015.) |
⊢ (𝐴 ⊆ (𝐵 ∖ 𝐴) ↔ 𝐴 = ∅) | ||
Theorem | ssundifim 3306 | A consequence of inclusion in the union of two classes. In classical logic this would be a biconditional. (Contributed by Jim Kingdon, 4-Aug-2018.) |
⊢ (𝐴 ⊆ (𝐵 ∪ 𝐶) → (𝐴 ∖ 𝐵) ⊆ 𝐶) | ||
Theorem | difdifdirss 3307 | Distributive law for class difference. In classical logic, as in Exercise 4.8 of [Stoll] p. 16, this would be equality rather than subset. (Contributed by Jim Kingdon, 4-Aug-2018.) |
⊢ ((𝐴 ∖ 𝐵) ∖ 𝐶) ⊆ ((𝐴 ∖ 𝐶) ∖ (𝐵 ∖ 𝐶)) | ||
Theorem | uneqdifeqim 3308 | Two ways that 𝐴 and 𝐵 can "partition" 𝐶 (when 𝐴 and 𝐵 don't overlap and 𝐴 is a part of 𝐶). In classical logic, the second implication would be a biconditional. (Contributed by Jim Kingdon, 4-Aug-2018.) |
⊢ ((𝐴 ⊆ 𝐶 ∧ (𝐴 ∩ 𝐵) = ∅) → ((𝐴 ∪ 𝐵) = 𝐶 → (𝐶 ∖ 𝐴) = 𝐵)) | ||
Theorem | r19.2m 3309* | Theorem 19.2 of [Margaris] p. 89 with restricted quantifiers (compare 19.2 1529). The restricted version is valid only when the domain of quantification is inhabited. (Contributed by Jim Kingdon, 5-Aug-2018.) |
⊢ ((∃𝑥 𝑥 ∈ 𝐴 ∧ ∀𝑥 ∈ 𝐴 𝜑) → ∃𝑥 ∈ 𝐴 𝜑) | ||
Theorem | r19.3rm 3310* | Restricted quantification of wff not containing quantified variable. (Contributed by Jim Kingdon, 19-Dec-2018.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (∃𝑦 𝑦 ∈ 𝐴 → (𝜑 ↔ ∀𝑥 ∈ 𝐴 𝜑)) | ||
Theorem | r19.28m 3311* | Restricted quantifier version of Theorem 19.28 of [Margaris] p. 90. It is valid only when the domain of quantification is inhabited. (Contributed by Jim Kingdon, 5-Aug-2018.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (∃𝑥 𝑥 ∈ 𝐴 → (∀𝑥 ∈ 𝐴 (𝜑 ∧ 𝜓) ↔ (𝜑 ∧ ∀𝑥 ∈ 𝐴 𝜓))) | ||
Theorem | r19.3rmv 3312* | Restricted quantification of wff not containing quantified variable. (Contributed by Jim Kingdon, 6-Aug-2018.) |
⊢ (∃𝑦 𝑦 ∈ 𝐴 → (𝜑 ↔ ∀𝑥 ∈ 𝐴 𝜑)) | ||
Theorem | r19.9rmv 3313* | Restricted quantification of wff not containing quantified variable. (Contributed by Jim Kingdon, 5-Aug-2018.) |
⊢ (∃𝑦 𝑦 ∈ 𝐴 → (𝜑 ↔ ∃𝑥 ∈ 𝐴 𝜑)) | ||
Theorem | r19.28mv 3314* | Restricted quantifier version of Theorem 19.28 of [Margaris] p. 90. It is valid only when the domain of quantification is inhabited. (Contributed by Jim Kingdon, 6-Aug-2018.) |
⊢ (∃𝑥 𝑥 ∈ 𝐴 → (∀𝑥 ∈ 𝐴 (𝜑 ∧ 𝜓) ↔ (𝜑 ∧ ∀𝑥 ∈ 𝐴 𝜓))) | ||
Theorem | r19.45mv 3315* | Restricted version of Theorem 19.45 of [Margaris] p. 90. (Contributed by NM, 27-May-1998.) |
⊢ (∃𝑥 𝑥 ∈ 𝐴 → (∃𝑥 ∈ 𝐴 (𝜑 ∨ 𝜓) ↔ (𝜑 ∨ ∃𝑥 ∈ 𝐴 𝜓))) | ||
Theorem | r19.27m 3316* | Restricted quantifier version of Theorem 19.27 of [Margaris] p. 90. It is valid only when the domain of quantification is inhabited. (Contributed by Jim Kingdon, 5-Aug-2018.) |
⊢ Ⅎ𝑥𝜓 ⇒ ⊢ (∃𝑥 𝑥 ∈ 𝐴 → (∀𝑥 ∈ 𝐴 (𝜑 ∧ 𝜓) ↔ (∀𝑥 ∈ 𝐴 𝜑 ∧ 𝜓))) | ||
Theorem | r19.27mv 3317* | Restricted quantifier version of Theorem 19.27 of [Margaris] p. 90. It is valid only when the domain of quantification is inhabited. (Contributed by Jim Kingdon, 5-Aug-2018.) |
⊢ (∃𝑥 𝑥 ∈ 𝐴 → (∀𝑥 ∈ 𝐴 (𝜑 ∧ 𝜓) ↔ (∀𝑥 ∈ 𝐴 𝜑 ∧ 𝜓))) | ||
Theorem | rzal 3318* | Vacuous quantification is always true. (Contributed by NM, 11-Mar-1997.) (Proof shortened by Andrew Salmon, 26-Jun-2011.) |
⊢ (𝐴 = ∅ → ∀𝑥 ∈ 𝐴 𝜑) | ||
Theorem | rexn0 3319* | Restricted existential quantification implies its restriction is nonempty (it is also inhabited as shown in rexm 3320). (Contributed by Szymon Jaroszewicz, 3-Apr-2007.) |
⊢ (∃𝑥 ∈ 𝐴 𝜑 → 𝐴 ≠ ∅) | ||
Theorem | rexm 3320* | Restricted existential quantification implies its restriction is inhabited. (Contributed by Jim Kingdon, 16-Oct-2018.) |
⊢ (∃𝑥 ∈ 𝐴 𝜑 → ∃𝑥 𝑥 ∈ 𝐴) | ||
Theorem | ralidm 3321* | Idempotent law for restricted quantifier. (Contributed by NM, 28-Mar-1997.) |
⊢ (∀𝑥 ∈ 𝐴 ∀𝑥 ∈ 𝐴 𝜑 ↔ ∀𝑥 ∈ 𝐴 𝜑) | ||
Theorem | ral0 3322 | Vacuous universal quantification is always true. (Contributed by NM, 20-Oct-2005.) |
⊢ ∀𝑥 ∈ ∅ 𝜑 | ||
Theorem | rgenm 3323* | Generalization rule that eliminates an inhabited class requirement. (Contributed by Jim Kingdon, 5-Aug-2018.) |
⊢ ((∃𝑥 𝑥 ∈ 𝐴 ∧ 𝑥 ∈ 𝐴) → 𝜑) ⇒ ⊢ ∀𝑥 ∈ 𝐴 𝜑 | ||
Theorem | ralf0 3324* | The quantification of a falsehood is vacuous when true. (Contributed by NM, 26-Nov-2005.) |
⊢ ¬ 𝜑 ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝜑 ↔ 𝐴 = ∅) | ||
Theorem | ralm 3325 | Inhabited classes and restricted quantification. (Contributed by Jim Kingdon, 6-Aug-2018.) |
⊢ ((∃𝑥 𝑥 ∈ 𝐴 → ∀𝑥 ∈ 𝐴 𝜑) ↔ ∀𝑥 ∈ 𝐴 𝜑) | ||
Theorem | raaanlem 3326* | Special case of raaan 3327 where 𝐴 is inhabited. (Contributed by Jim Kingdon, 6-Aug-2018.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 ⇒ ⊢ (∃𝑥 𝑥 ∈ 𝐴 → (∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 (𝜑 ∧ 𝜓) ↔ (∀𝑥 ∈ 𝐴 𝜑 ∧ ∀𝑦 ∈ 𝐴 𝜓))) | ||
Theorem | raaan 3327* | Rearrange restricted quantifiers. (Contributed by NM, 26-Oct-2010.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 ⇒ ⊢ (∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 (𝜑 ∧ 𝜓) ↔ (∀𝑥 ∈ 𝐴 𝜑 ∧ ∀𝑦 ∈ 𝐴 𝜓)) | ||
Theorem | raaanv 3328* | Rearrange restricted quantifiers. (Contributed by NM, 11-Mar-1997.) |
⊢ (∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 (𝜑 ∧ 𝜓) ↔ (∀𝑥 ∈ 𝐴 𝜑 ∧ ∀𝑦 ∈ 𝐴 𝜓)) | ||
Theorem | sbss 3329* | Set substitution into the first argument of a subset relation. (Contributed by Rodolfo Medina, 7-Jul-2010.) (Proof shortened by Mario Carneiro, 14-Nov-2016.) |
⊢ ([𝑦 / 𝑥]𝑥 ⊆ 𝐴 ↔ 𝑦 ⊆ 𝐴) | ||
Theorem | sbcssg 3330 | Distribute proper substitution through a subclass relation. (Contributed by Alan Sare, 22-Jul-2012.) (Proof shortened by Alexander van der Vekens, 23-Jul-2017.) |
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥]𝐵 ⊆ 𝐶 ↔ ⦋𝐴 / 𝑥⦌𝐵 ⊆ ⦋𝐴 / 𝑥⦌𝐶)) | ||
Syntax | cif 3331 | Extend class notation to include the conditional operator. See df-if 3332 for a description. (In older databases this was denoted "ded".) |
class if(𝜑, 𝐴, 𝐵) | ||
Definition | df-if 3332* |
Define the conditional operator. Read if(𝜑, 𝐴, 𝐵) as "if
𝜑 then 𝐴 else 𝐵."
See iftrue 3336 and iffalse 3339 for its
values. In mathematical literature, this operator is rarely defined
formally but is implicit in informal definitions such as "let
f(x)=0 if
x=0 and 1/x otherwise."
In the absence of excluded middle, this will tend to be useful where 𝜑 is decidable (in the sense of df-dc 743). (Contributed by NM, 15-May-1999.) |
⊢ if(𝜑, 𝐴, 𝐵) = {𝑥 ∣ ((𝑥 ∈ 𝐴 ∧ 𝜑) ∨ (𝑥 ∈ 𝐵 ∧ ¬ 𝜑))} | ||
Theorem | dfif6 3333* | An alternate definition of the conditional operator df-if 3332 as a simple class abstraction. (Contributed by Mario Carneiro, 8-Sep-2013.) |
⊢ if(𝜑, 𝐴, 𝐵) = ({𝑥 ∈ 𝐴 ∣ 𝜑} ∪ {𝑥 ∈ 𝐵 ∣ ¬ 𝜑}) | ||
Theorem | ifeq1 3334 | Equality theorem for conditional operator. (Contributed by NM, 1-Sep-2004.) (Revised by Mario Carneiro, 8-Sep-2013.) |
⊢ (𝐴 = 𝐵 → if(𝜑, 𝐴, 𝐶) = if(𝜑, 𝐵, 𝐶)) | ||
Theorem | ifeq2 3335 | Equality theorem for conditional operator. (Contributed by NM, 1-Sep-2004.) (Revised by Mario Carneiro, 8-Sep-2013.) |
⊢ (𝐴 = 𝐵 → if(𝜑, 𝐶, 𝐴) = if(𝜑, 𝐶, 𝐵)) | ||
Theorem | iftrue 3336 | Value of the conditional operator when its first argument is true. (Contributed by NM, 15-May-1999.) (Proof shortened by Andrew Salmon, 26-Jun-2011.) |
⊢ (𝜑 → if(𝜑, 𝐴, 𝐵) = 𝐴) | ||
Theorem | iftruei 3337 | Inference associated with iftrue 3336. (Contributed by BJ, 7-Oct-2018.) |
⊢ 𝜑 ⇒ ⊢ if(𝜑, 𝐴, 𝐵) = 𝐴 | ||
Theorem | iftrued 3338 | Value of the conditional operator when its first argument is true. (Contributed by Glauco Siliprandi, 11-Dec-2019.) |
⊢ (𝜑 → 𝜒) ⇒ ⊢ (𝜑 → if(𝜒, 𝐴, 𝐵) = 𝐴) | ||
Theorem | iffalse 3339 | Value of the conditional operator when its first argument is false. (Contributed by NM, 14-Aug-1999.) |
⊢ (¬ 𝜑 → if(𝜑, 𝐴, 𝐵) = 𝐵) | ||
Theorem | iffalsei 3340 | Inference associated with iffalse 3339. (Contributed by BJ, 7-Oct-2018.) |
⊢ ¬ 𝜑 ⇒ ⊢ if(𝜑, 𝐴, 𝐵) = 𝐵 | ||
Theorem | iffalsed 3341 | Value of the conditional operator when its first argument is false. (Contributed by Glauco Siliprandi, 11-Dec-2019.) |
⊢ (𝜑 → ¬ 𝜒) ⇒ ⊢ (𝜑 → if(𝜒, 𝐴, 𝐵) = 𝐵) | ||
Theorem | ifnefalse 3342 | When values are unequal, but an "if" condition checks if they are equal, then the "false" branch results. This is a simple utility to provide a slight shortening and simplification of proofs vs. applying iffalse 3339 directly in this case. (Contributed by David A. Wheeler, 15-May-2015.) |
⊢ (𝐴 ≠ 𝐵 → if(𝐴 = 𝐵, 𝐶, 𝐷) = 𝐷) | ||
Theorem | dfif3 3343* | Alternate definition of the conditional operator df-if 3332. Note that 𝜑 is independent of 𝑥 i.e. a constant true or false. (Contributed by NM, 25-Aug-2013.) (Revised by Mario Carneiro, 8-Sep-2013.) |
⊢ 𝐶 = {𝑥 ∣ 𝜑} ⇒ ⊢ if(𝜑, 𝐴, 𝐵) = ((𝐴 ∩ 𝐶) ∪ (𝐵 ∩ (V ∖ 𝐶))) | ||
Theorem | ifeq12 3344 | Equality theorem for conditional operators. (Contributed by NM, 1-Sep-2004.) |
⊢ ((𝐴 = 𝐵 ∧ 𝐶 = 𝐷) → if(𝜑, 𝐴, 𝐶) = if(𝜑, 𝐵, 𝐷)) | ||
Theorem | ifeq1d 3345 | Equality deduction for conditional operator. (Contributed by NM, 16-Feb-2005.) |
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → if(𝜓, 𝐴, 𝐶) = if(𝜓, 𝐵, 𝐶)) | ||
Theorem | ifeq2d 3346 | Equality deduction for conditional operator. (Contributed by NM, 16-Feb-2005.) |
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → if(𝜓, 𝐶, 𝐴) = if(𝜓, 𝐶, 𝐵)) | ||
Theorem | ifeq12d 3347 | Equality deduction for conditional operator. (Contributed by NM, 24-Mar-2015.) |
⊢ (𝜑 → 𝐴 = 𝐵) & ⊢ (𝜑 → 𝐶 = 𝐷) ⇒ ⊢ (𝜑 → if(𝜓, 𝐴, 𝐶) = if(𝜓, 𝐵, 𝐷)) | ||
Theorem | ifbi 3348 | Equivalence theorem for conditional operators. (Contributed by Raph Levien, 15-Jan-2004.) |
⊢ ((𝜑 ↔ 𝜓) → if(𝜑, 𝐴, 𝐵) = if(𝜓, 𝐴, 𝐵)) | ||
Theorem | ifbid 3349 | Equivalence deduction for conditional operators. (Contributed by NM, 18-Apr-2005.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → if(𝜓, 𝐴, 𝐵) = if(𝜒, 𝐴, 𝐵)) | ||
Theorem | ifbieq1d 3350 | Equivalence/equality deduction for conditional operators. (Contributed by JJ, 25-Sep-2018.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → if(𝜓, 𝐴, 𝐶) = if(𝜒, 𝐵, 𝐶)) | ||
Theorem | ifbieq2i 3351 | Equivalence/equality inference for conditional operators. (Contributed by Paul Chapman, 22-Jun-2011.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ 𝐴 = 𝐵 ⇒ ⊢ if(𝜑, 𝐶, 𝐴) = if(𝜓, 𝐶, 𝐵) | ||
Theorem | ifbieq2d 3352 | Equivalence/equality deduction for conditional operators. (Contributed by Paul Chapman, 22-Jun-2011.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → if(𝜓, 𝐶, 𝐴) = if(𝜒, 𝐶, 𝐵)) | ||
Theorem | ifbieq12i 3353 | Equivalence deduction for conditional operators. (Contributed by NM, 18-Mar-2013.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ 𝐴 = 𝐶 & ⊢ 𝐵 = 𝐷 ⇒ ⊢ if(𝜑, 𝐴, 𝐵) = if(𝜓, 𝐶, 𝐷) | ||
Theorem | ifbieq12d 3354 | Equivalence deduction for conditional operators. (Contributed by Jeff Madsen, 2-Sep-2009.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → 𝐴 = 𝐶) & ⊢ (𝜑 → 𝐵 = 𝐷) ⇒ ⊢ (𝜑 → if(𝜓, 𝐴, 𝐵) = if(𝜒, 𝐶, 𝐷)) | ||
Theorem | nfifd 3355 | Deduction version of nfif 3356. (Contributed by NM, 15-Feb-2013.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ (𝜑 → Ⅎ𝑥𝜓) & ⊢ (𝜑 → Ⅎ𝑥𝐴) & ⊢ (𝜑 → Ⅎ𝑥𝐵) ⇒ ⊢ (𝜑 → Ⅎ𝑥if(𝜓, 𝐴, 𝐵)) | ||
Theorem | nfif 3356 | Bound-variable hypothesis builder for a conditional operator. (Contributed by NM, 16-Feb-2005.) (Proof shortened by Andrew Salmon, 26-Jun-2011.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑥𝐴 & ⊢ Ⅎ𝑥𝐵 ⇒ ⊢ Ⅎ𝑥if(𝜑, 𝐴, 𝐵) | ||
Theorem | ifbothdc 3357 | A wff 𝜃 containing a conditional operator is true when both of its cases are true. (Contributed by Jim Kingdon, 8-Aug-2021.) |
⊢ (𝐴 = if(𝜑, 𝐴, 𝐵) → (𝜓 ↔ 𝜃)) & ⊢ (𝐵 = if(𝜑, 𝐴, 𝐵) → (𝜒 ↔ 𝜃)) ⇒ ⊢ ((𝜓 ∧ 𝜒 ∧ DECID 𝜑) → 𝜃) | ||
Theorem | ifcldcd 3358 | Membership (closure) of a conditional operator, deduction form. (Contributed by Jim Kingdon, 8-Aug-2021.) |
⊢ (𝜑 → 𝐴 ∈ 𝐶) & ⊢ (𝜑 → 𝐵 ∈ 𝐶) & ⊢ (𝜑 → DECID 𝜓) ⇒ ⊢ (𝜑 → if(𝜓, 𝐴, 𝐵) ∈ 𝐶) | ||
Syntax | cpw 3359 | Extend class notation to include power class. (The tilde in the Metamath token is meant to suggest the calligraphic font of the P.) |
class 𝒫 𝐴 | ||
Theorem | pwjust 3360* | Soundness justification theorem for df-pw 3361. (Contributed by Rodolfo Medina, 28-Apr-2010.) (Proof shortened by Andrew Salmon, 29-Jun-2011.) |
⊢ {𝑥 ∣ 𝑥 ⊆ 𝐴} = {𝑦 ∣ 𝑦 ⊆ 𝐴} | ||
Definition | df-pw 3361* | Define power class. Definition 5.10 of [TakeutiZaring] p. 17, but we also let it apply to proper classes, i.e. those that are not members of V. When applied to a set, this produces its power set. A power set of S is the set of all subsets of S, including the empty set and S itself. For example, if 𝐴 is { 3 , 5 , 7 }, then 𝒫 𝐴 is { (/) , { 3 } , { 5 } , { 7 } , { 3 , 5 } , { 3 , 7 } , { 5 , 7 } , { 3 , 5 , 7 } }. We will later introduce the Axiom of Power Sets. Still later we will prove that the size of the power set of a finite set is 2 raised to the power of the size of the set. (Contributed by NM, 5-Aug-1993.) |
⊢ 𝒫 𝐴 = {𝑥 ∣ 𝑥 ⊆ 𝐴} | ||
Theorem | pweq 3362 | Equality theorem for power class. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝐴 = 𝐵 → 𝒫 𝐴 = 𝒫 𝐵) | ||
Theorem | pweqi 3363 | Equality inference for power class. (Contributed by NM, 27-Nov-2013.) |
⊢ 𝐴 = 𝐵 ⇒ ⊢ 𝒫 𝐴 = 𝒫 𝐵 | ||
Theorem | pweqd 3364 | Equality deduction for power class. (Contributed by NM, 27-Nov-2013.) |
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → 𝒫 𝐴 = 𝒫 𝐵) | ||
Theorem | elpw 3365 | Membership in a power class. Theorem 86 of [Suppes] p. 47. (Contributed by NM, 31-Dec-1993.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (𝐴 ∈ 𝒫 𝐵 ↔ 𝐴 ⊆ 𝐵) | ||
Theorem | selpw 3366* | Setvar variable membership in a power class (common case). See elpw 3365. (Contributed by David A. Wheeler, 8-Dec-2018.) |
⊢ (𝑥 ∈ 𝒫 𝐴 ↔ 𝑥 ⊆ 𝐴) | ||
Theorem | elpwg 3367 | Membership in a power class. Theorem 86 of [Suppes] p. 47. (Contributed by NM, 6-Aug-2000.) |
⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ 𝒫 𝐵 ↔ 𝐴 ⊆ 𝐵)) | ||
Theorem | elpwi 3368 | Subset relation implied by membership in a power class. (Contributed by NM, 17-Feb-2007.) |
⊢ (𝐴 ∈ 𝒫 𝐵 → 𝐴 ⊆ 𝐵) | ||
Theorem | elpwid 3369 | An element of a power class is a subclass. Deduction form of elpwi 3368. (Contributed by David Moews, 1-May-2017.) |
⊢ (𝜑 → 𝐴 ∈ 𝒫 𝐵) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐵) | ||
Theorem | elelpwi 3370 | If 𝐴 belongs to a part of 𝐶 then 𝐴 belongs to 𝐶. (Contributed by FL, 3-Aug-2009.) |
⊢ ((𝐴 ∈ 𝐵 ∧ 𝐵 ∈ 𝒫 𝐶) → 𝐴 ∈ 𝐶) | ||
Theorem | nfpw 3371 | Bound-variable hypothesis builder for power class. (Contributed by NM, 28-Oct-2003.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ Ⅎ𝑥𝐴 ⇒ ⊢ Ⅎ𝑥𝒫 𝐴 | ||
Theorem | pwidg 3372 | Membership of the original in a power set. (Contributed by Stefan O'Rear, 1-Feb-2015.) |
⊢ (𝐴 ∈ 𝑉 → 𝐴 ∈ 𝒫 𝐴) | ||
Theorem | pwid 3373 | A set is a member of its power class. Theorem 87 of [Suppes] p. 47. (Contributed by NM, 5-Aug-1993.) |
⊢ 𝐴 ∈ V ⇒ ⊢ 𝐴 ∈ 𝒫 𝐴 | ||
Theorem | pwss 3374* | Subclass relationship for power class. (Contributed by NM, 21-Jun-2009.) |
⊢ (𝒫 𝐴 ⊆ 𝐵 ↔ ∀𝑥(𝑥 ⊆ 𝐴 → 𝑥 ∈ 𝐵)) | ||
Syntax | csn 3375 | Extend class notation to include singleton. |
class {𝐴} | ||
Syntax | cpr 3376 | Extend class notation to include unordered pair. |
class {𝐴, 𝐵} | ||
Syntax | ctp 3377 | Extend class notation to include unordered triplet. |
class {𝐴, 𝐵, 𝐶} | ||
Syntax | cop 3378 | Extend class notation to include ordered pair. |
class 〈𝐴, 𝐵〉 | ||
Syntax | cotp 3379 | Extend class notation to include ordered triple. |
class 〈𝐴, 𝐵, 𝐶〉 | ||
Theorem | snjust 3380* | Soundness justification theorem for df-sn 3381. (Contributed by Rodolfo Medina, 28-Apr-2010.) (Proof shortened by Andrew Salmon, 29-Jun-2011.) |
⊢ {𝑥 ∣ 𝑥 = 𝐴} = {𝑦 ∣ 𝑦 = 𝐴} | ||
Definition | df-sn 3381* | Define the singleton of a class. Definition 7.1 of [Quine] p. 48. For convenience, it is well-defined for proper classes, i.e., those that are not elements of V, although it is not very meaningful in this case. For an alternate definition see dfsn2 3389. (Contributed by NM, 5-Aug-1993.) |
⊢ {𝐴} = {𝑥 ∣ 𝑥 = 𝐴} | ||
Definition | df-pr 3382 | Define unordered pair of classes. Definition 7.1 of [Quine] p. 48. They are unordered, so {𝐴, 𝐵} = {𝐵, 𝐴} as proven by prcom 3446. For a more traditional definition, but requiring a dummy variable, see dfpr2 3394. (Contributed by NM, 5-Aug-1993.) |
⊢ {𝐴, 𝐵} = ({𝐴} ∪ {𝐵}) | ||
Definition | df-tp 3383 | Define unordered triple of classes. Definition of [Enderton] p. 19. (Contributed by NM, 9-Apr-1994.) |
⊢ {𝐴, 𝐵, 𝐶} = ({𝐴, 𝐵} ∪ {𝐶}) | ||
Definition | df-op 3384* |
Definition of an ordered pair, equivalent to Kuratowski's definition
{{𝐴}, {𝐴, 𝐵}} when the arguments are sets.
Since the
behavior of Kuratowski definition is not very useful for proper classes,
we define it to be empty in this case (see opprc1 3571 and opprc2 3572). For
Kuratowski's actual definition when the arguments are sets, see dfop 3548.
Definition 9.1 of [Quine] p. 58 defines an ordered pair unconditionally as 〈𝐴, 𝐵〉 = {{𝐴}, {𝐴, 𝐵}}, which has different behavior from our df-op 3384 when the arguments are proper classes. Ordinarily this difference is not important, since neither definition is meaningful in that case. Our df-op 3384 was chosen because it often makes proofs shorter by eliminating unnecessary sethood hypotheses. There are other ways to define ordered pairs. The basic requirement is that two ordered pairs are equal iff their respective members are equal. In 1914 Norbert Wiener gave the first successful definition 〈𝐴, 𝐵〉_2 = {{{𝐴}, ∅}, {{𝐵}}}. This was simplified by Kazimierz Kuratowski in 1921 to our present definition. An even simpler definition is 〈𝐴, 𝐵〉_3 = {𝐴, {𝐴, 𝐵}}, but it requires the Axiom of Regularity for its justification and is not commonly used. Finally, an ordered pair of real numbers can be represented by a complex number. (Contributed by NM, 28-May-1995.) (Revised by Mario Carneiro, 26-Apr-2015.) |
⊢ 〈𝐴, 𝐵〉 = {𝑥 ∣ (𝐴 ∈ V ∧ 𝐵 ∈ V ∧ 𝑥 ∈ {{𝐴}, {𝐴, 𝐵}})} | ||
Definition | df-ot 3385 | Define ordered triple of classes. Definition of ordered triple in [Stoll] p. 25. (Contributed by NM, 3-Apr-2015.) |
⊢ 〈𝐴, 𝐵, 𝐶〉 = 〈〈𝐴, 𝐵〉, 𝐶〉 | ||
Theorem | sneq 3386 | Equality theorem for singletons. Part of Exercise 4 of [TakeutiZaring] p. 15. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝐴 = 𝐵 → {𝐴} = {𝐵}) | ||
Theorem | sneqi 3387 | Equality inference for singletons. (Contributed by NM, 22-Jan-2004.) |
⊢ 𝐴 = 𝐵 ⇒ ⊢ {𝐴} = {𝐵} | ||
Theorem | sneqd 3388 | Equality deduction for singletons. (Contributed by NM, 22-Jan-2004.) |
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → {𝐴} = {𝐵}) | ||
Theorem | dfsn2 3389 | Alternate definition of singleton. Definition 5.1 of [TakeutiZaring] p. 15. (Contributed by NM, 24-Apr-1994.) |
⊢ {𝐴} = {𝐴, 𝐴} | ||
Theorem | elsng 3390 | There is exactly one element in a singleton. Exercise 2 of [TakeutiZaring] p. 15 (generalized). (Contributed by NM, 13-Sep-1995.) (Proof shortened by Andrew Salmon, 29-Jun-2011.) |
⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ {𝐵} ↔ 𝐴 = 𝐵)) | ||
Theorem | elsn 3391 | There is exactly one element in a singleton. Exercise 2 of [TakeutiZaring] p. 15. (Contributed by NM, 13-Sep-1995.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (𝐴 ∈ {𝐵} ↔ 𝐴 = 𝐵) | ||
Theorem | velsn 3392 | There is only one element in a singleton. Exercise 2 of [TakeutiZaring] p. 15. (Contributed by NM, 21-Jun-1993.) |
⊢ (𝑥 ∈ {𝐴} ↔ 𝑥 = 𝐴) | ||
Theorem | elsni 3393 | There is only one element in a singleton. (Contributed by NM, 5-Jun-1994.) |
⊢ (𝐴 ∈ {𝐵} → 𝐴 = 𝐵) | ||
Theorem | dfpr2 3394* | Alternate definition of unordered pair. Definition 5.1 of [TakeutiZaring] p. 15. (Contributed by NM, 24-Apr-1994.) |
⊢ {𝐴, 𝐵} = {𝑥 ∣ (𝑥 = 𝐴 ∨ 𝑥 = 𝐵)} | ||
Theorem | elprg 3395 | A member of an unordered pair of classes is one or the other of them. Exercise 1 of [TakeutiZaring] p. 15, generalized. (Contributed by NM, 13-Sep-1995.) |
⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ {𝐵, 𝐶} ↔ (𝐴 = 𝐵 ∨ 𝐴 = 𝐶))) | ||
Theorem | elpr 3396 | A member of an unordered pair of classes is one or the other of them. Exercise 1 of [TakeutiZaring] p. 15. (Contributed by NM, 13-Sep-1995.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (𝐴 ∈ {𝐵, 𝐶} ↔ (𝐴 = 𝐵 ∨ 𝐴 = 𝐶)) | ||
Theorem | elpr2 3397 | A member of an unordered pair of classes is one or the other of them. Exercise 1 of [TakeutiZaring] p. 15. (Contributed by NM, 14-Oct-2005.) |
⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V ⇒ ⊢ (𝐴 ∈ {𝐵, 𝐶} ↔ (𝐴 = 𝐵 ∨ 𝐴 = 𝐶)) | ||
Theorem | elpri 3398 | If a class is an element of a pair, then it is one of the two paired elements. (Contributed by Scott Fenton, 1-Apr-2011.) |
⊢ (𝐴 ∈ {𝐵, 𝐶} → (𝐴 = 𝐵 ∨ 𝐴 = 𝐶)) | ||
Theorem | nelpri 3399 | If an element doesn't match the items in an unordered pair, it is not in the unordered pair. (Contributed by David A. Wheeler, 10-May-2015.) |
⊢ 𝐴 ≠ 𝐵 & ⊢ 𝐴 ≠ 𝐶 ⇒ ⊢ ¬ 𝐴 ∈ {𝐵, 𝐶} | ||
Theorem | snidg 3400 | A set is a member of its singleton. Part of Theorem 7.6 of [Quine] p. 49. (Contributed by NM, 28-Oct-2003.) |
⊢ (𝐴 ∈ 𝑉 → 𝐴 ∈ {𝐴}) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |