Theorem List for New Foundations Explorer - 3701-3800 *Has distinct variable
group(s)
Type | Label | Description |
Statement |
|
Theorem | ifan 3701 |
Rewrite a conjunction in an if statement as two nested conditionals.
(Contributed by Mario Carneiro, 28-Jul-2014.)
|
⊢ if((φ
∧ ψ),
A, B)
= if(φ, if(ψ, A,
B), B) |
|
Theorem | ifor 3702 |
Rewrite a disjunction in an if statement as two nested conditionals.
(Contributed by Mario Carneiro, 28-Jul-2014.)
|
⊢ if((φ
∨ ψ),
A, B)
= if(φ, A, if(ψ,
A, B)) |
|
Theorem | dedth 3703 |
Weak deduction theorem that eliminates a hypothesis φ, making it
become an antecedent. We assume that a proof exists for φ when the
class variable A is
replaced with a specific class B. The
hypothesis χ
should be assigned to the inference, and the
inference's hypothesis eliminated with elimhyp 3710. If the inference has
other hypotheses with class variable A, these can be kept by
assigning keephyp 3716 to them. For more information, see the
Deduction
Theorem https://us.metamath.org/mpeuni/mmdeduction.html 3716.
(Contributed by NM, 15-May-1999.)
|
⊢ (A =
if(φ, A, B) →
(ψ ↔ χ))
& ⊢ χ ⇒ ⊢ (φ
→ ψ) |
|
Theorem | dedth2h 3704 |
Weak deduction theorem eliminating two hypotheses. This theorem is
simpler to use than dedth2v 3707 but requires that each hypothesis has
exactly one class variable. See also comments in dedth 3703.
(Contributed by NM, 15-May-1999.)
|
⊢ (A =
if(φ, A, C) →
(χ ↔ θ))
& ⊢ (B =
if(ψ, B, D) →
(θ ↔ τ))
& ⊢ τ ⇒ ⊢ ((φ
∧ ψ)
→ χ) |
|
Theorem | dedth3h 3705 |
Weak deduction theorem eliminating three hypotheses. See comments in
dedth2h 3704. (Contributed by NM, 15-May-1999.)
|
⊢ (A =
if(φ, A, D) →
(θ ↔ τ))
& ⊢ (B =
if(ψ, B, R) →
(τ ↔ η))
& ⊢ (C =
if(χ, C, S) →
(η ↔ ζ))
& ⊢ ζ ⇒ ⊢ ((φ
∧ ψ
∧ χ)
→ θ) |
|
Theorem | dedth4h 3706 |
Weak deduction theorem eliminating four hypotheses. See comments in
dedth2h 3704. (Contributed by NM, 16-May-1999.)
|
⊢ (A =
if(φ, A, R) →
(τ ↔ η))
& ⊢ (B =
if(ψ, B, S) →
(η ↔ ζ))
& ⊢ (C =
if(χ, C, F) →
(ζ ↔ σ))
& ⊢ (D =
if(θ, D, G) →
(σ ↔ ρ))
& ⊢ ρ ⇒ ⊢ (((φ
∧ ψ)
∧ (χ
∧ θ)) → τ) |
|
Theorem | dedth2v 3707 |
Weak deduction theorem for eliminating a hypothesis with 2 class
variables. Note: if the hypothesis can be separated into two
hypotheses, each with one class variable, then dedth2h 3704 is simpler to
use. See also comments in dedth 3703. (Contributed by NM, 13-Aug-1999.)
(Proof shortened by Eric Schmidt, 28-Jul-2009.)
|
⊢ (A =
if(φ, A, C) →
(ψ ↔ χ))
& ⊢ (B =
if(φ, B, D) →
(χ ↔ θ))
& ⊢ θ ⇒ ⊢ (φ
→ ψ) |
|
Theorem | dedth3v 3708 |
Weak deduction theorem for eliminating a hypothesis with 3 class
variables. See comments in dedth2v 3707. (Contributed by NM,
13-Aug-1999.) (Proof shortened by Eric Schmidt, 28-Jul-2009.)
|
⊢ (A =
if(φ, A, D) →
(ψ ↔ χ))
& ⊢ (B =
if(φ, B, R) →
(χ ↔ θ))
& ⊢ (C =
if(φ, C, S) →
(θ ↔ τ))
& ⊢ τ ⇒ ⊢ (φ
→ ψ) |
|
Theorem | dedth4v 3709 |
Weak deduction theorem for eliminating a hypothesis with 4 class
variables. See comments in dedth2v 3707. (Contributed by NM,
21-Apr-2007.) (Proof shortened by Eric Schmidt, 28-Jul-2009.)
|
⊢ (A =
if(φ, A, R) →
(ψ ↔ χ))
& ⊢ (B =
if(φ, B, S) →
(χ ↔ θ))
& ⊢ (C =
if(φ, C, T) →
(θ ↔ τ))
& ⊢ (D =
if(φ, D, U) →
(τ ↔ η))
& ⊢ η ⇒ ⊢ (φ
→ ψ) |
|
Theorem | elimhyp 3710 |
Eliminate a hypothesis containing class variable A when it is known
for a specific class B. For more information, see comments in
dedth 3703. (Contributed by NM, 15-May-1999.)
|
⊢ (A =
if(φ, A, B) →
(φ ↔ ψ))
& ⊢ (B =
if(φ, A, B) →
(χ ↔ ψ))
& ⊢ χ ⇒ ⊢ ψ |
|
Theorem | elimhyp2v 3711 |
Eliminate a hypothesis containing 2 class variables. (Contributed by
NM, 14-Aug-1999.)
|
⊢ (A =
if(φ, A, C) →
(φ ↔ χ))
& ⊢ (B =
if(φ, B, D) →
(χ ↔ θ))
& ⊢ (C =
if(φ, A, C) →
(τ ↔ η))
& ⊢ (D =
if(φ, B, D) →
(η ↔ θ))
& ⊢ τ ⇒ ⊢ θ |
|
Theorem | elimhyp3v 3712 |
Eliminate a hypothesis containing 3 class variables. (Contributed by
NM, 14-Aug-1999.)
|
⊢ (A =
if(φ, A, D) →
(φ ↔ χ))
& ⊢ (B =
if(φ, B, R) →
(χ ↔ θ))
& ⊢ (C =
if(φ, C, S) →
(θ ↔ τ))
& ⊢ (D =
if(φ, A, D) →
(η ↔ ζ))
& ⊢ (R =
if(φ, B, R) →
(ζ ↔ σ))
& ⊢ (S =
if(φ, C, S) →
(σ ↔ τ))
& ⊢ η ⇒ ⊢ τ |
|
Theorem | elimhyp4v 3713 |
Eliminate a hypothesis containing 4 class variables (for use with the
weak deduction theorem dedth 3703). (Contributed by NM, 16-Apr-2005.)
|
⊢ (A =
if(φ, A, D) →
(φ ↔ χ))
& ⊢ (B =
if(φ, B, R) →
(χ ↔ θ))
& ⊢ (C =
if(φ, C, S) →
(θ ↔ τ))
& ⊢ (F =
if(φ, F, G) →
(τ ↔ ψ))
& ⊢ (D =
if(φ, A, D) →
(η ↔ ζ))
& ⊢ (R =
if(φ, B, R) →
(ζ ↔ σ))
& ⊢ (S =
if(φ, C, S) →
(σ ↔ ρ))
& ⊢ (G =
if(φ, F, G) →
(ρ ↔ ψ))
& ⊢ η ⇒ ⊢ ψ |
|
Theorem | elimel 3714 |
Eliminate a membership hypothesis for weak deduction theorem, when
special case B ∈ C is
provable. (Contributed by NM,
15-May-1999.)
|
⊢ B ∈ C ⇒ ⊢ if(A ∈ C, A, B) ∈ C |
|
Theorem | elimdhyp 3715 |
Version of elimhyp 3710 where the hypothesis is deduced from the
final
antecedent. See ghomgrplem in set.mm for an example of its use.
(Contributed by Paul Chapman, 25-Mar-2008.)
|
⊢ (φ
→ ψ) & ⊢ (A =
if(φ, A, B) →
(ψ ↔ χ))
& ⊢ (B =
if(φ, A, B) →
(θ ↔ χ))
& ⊢ θ ⇒ ⊢ χ |
|
Theorem | keephyp 3716 |
Transform a hypothesis ψ
that we want to keep (but contains the
same class variable A
used in the eliminated hypothesis) for use
with the weak deduction theorem. (Contributed by NM, 15-May-1999.)
|
⊢ (A =
if(φ, A, B) →
(ψ ↔ θ))
& ⊢ (B =
if(φ, A, B) →
(χ ↔ θ))
& ⊢ ψ
& ⊢ χ ⇒ ⊢ θ |
|
Theorem | keephyp2v 3717 |
Keep a hypothesis containing 2 class variables (for use with the weak
deduction theorem dedth 3703). (Contributed by NM, 16-Apr-2005.)
|
⊢ (A =
if(φ, A, C) →
(ψ ↔ χ))
& ⊢ (B =
if(φ, B, D) →
(χ ↔ θ))
& ⊢ (C =
if(φ, A, C) →
(τ ↔ η))
& ⊢ (D =
if(φ, B, D) →
(η ↔ θ))
& ⊢ ψ
& ⊢ τ ⇒ ⊢ θ |
|
Theorem | keephyp3v 3718 |
Keep a hypothesis containing 3 class variables. (Contributed by NM,
27-Sep-1999.)
|
⊢ (A =
if(φ, A, D) →
(ρ ↔ χ))
& ⊢ (B =
if(φ, B, R) →
(χ ↔ θ))
& ⊢ (C =
if(φ, C, S) →
(θ ↔ τ))
& ⊢ (D =
if(φ, A, D) →
(η ↔ ζ))
& ⊢ (R =
if(φ, B, R) →
(ζ ↔ σ))
& ⊢ (S =
if(φ, C, S) →
(σ ↔ τ))
& ⊢ ρ
& ⊢ η ⇒ ⊢ τ |
|
Theorem | keepel 3719 |
Keep a membership hypothesis for weak deduction theorem, when special
case B ∈ C is
provable. (Contributed by NM, 14-Aug-1999.)
|
⊢ A ∈ C
& ⊢ B ∈ C ⇒ ⊢ if(φ,
A, B)
∈ C |
|
Theorem | ifex 3720 |
Conditional operator existence. (Contributed by NM, 2-Sep-2004.)
|
⊢ A ∈ V
& ⊢ B ∈ V ⇒ ⊢ if(φ,
A, B)
∈ V |
|
Theorem | ifexg 3721 |
Conditional operator existence. (Contributed by NM, 21-Mar-2011.)
|
⊢ ((A ∈ V ∧ B ∈ W) →
if(φ, A, B) ∈ V) |
|
2.1.15 Power classes
|
|
Syntax | cpw 3722 |
Extend class notation to include power class. (The tilde in the Metamath
token is meant to suggest the calligraphic font of the P.)
|
class
℘A |
|
Theorem | pwjust 3723* |
Soundness justification theorem for df-pw 3724. (Contributed by Rodolfo
Medina, 28-Apr-2010.) (Proof shortened by Andrew Salmon,
29-Jun-2011.)
|
⊢ {x ∣ x ⊆ A} =
{y ∣
y ⊆
A} |
|
Definition | df-pw 3724* |
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 A =
{ 3 , 5 , 7 }, then
℘A = {∅, {
3 }, { 5 }, { 7 }, { 3 , 5 },
{ 3 , 7 }, { 5 , 7 }, { 3 , 5
, 7 }} (ex-pw in
set.mm). We will later introduce the Axiom of Power Sets ax-pow in
set.mm, which can be expressed in class notation per pwexg 4328. Still
later we will prove, in hashpw in set.mm, 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.)
|
⊢ ℘A = {x ∣ x ⊆ A} |
|
Theorem | pweq 3725 |
Equality theorem for power class. (Contributed by NM, 5-Aug-1993.)
|
⊢ (A =
B → ℘A = ℘B) |
|
Theorem | pweqi 3726 |
Equality inference for power class. (Contributed by NM,
27-Nov-2013.)
|
⊢ A =
B ⇒ ⊢ ℘A = ℘B |
|
Theorem | pweqd 3727 |
Equality deduction for power class. (Contributed by NM,
27-Nov-2013.)
|
⊢ (φ
→ A = B) ⇒ ⊢ (φ
→ ℘A = ℘B) |
|
Theorem | elpw 3728 |
Membership in a power class. Theorem 86 of [Suppes] p. 47.
(Contributed by NM, 31-Dec-1993.)
|
⊢ A ∈ V ⇒ ⊢ (A ∈ ℘B ↔ A
⊆ B) |
|
Theorem | elpwg 3729 |
Membership in a power class. Theorem 86 of [Suppes] p. 47. See also
elpw2g in set.mm. (Contributed by NM, 6-Aug-2000.)
|
⊢ (A ∈ V →
(A ∈
℘B
↔ A ⊆ B)) |
|
Theorem | elpwi 3730 |
Subset relation implied by membership in a power class. (Contributed by
NM, 17-Feb-2007.)
|
⊢ (A ∈ ℘B → A
⊆ B) |
|
Theorem | elpwid 3731 |
An element of a power class is a subclass. Deduction form of elpwi 3730.
(Contributed by David Moews, 1-May-2017.)
|
⊢ (φ
→ A ∈ ℘B) ⇒ ⊢ (φ
→ A ⊆ B) |
|
Theorem | elelpwi 3732 |
If A belongs to a part of
C then A belongs to C.
(Contributed by FL, 3-Aug-2009.)
|
⊢ ((A ∈ B ∧ B ∈ ℘C) → A
∈ C) |
|
Theorem | nfpw 3733 |
Bound-variable hypothesis builder for power class. (Contributed by NM,
28-Oct-2003.) (Revised by Mario Carneiro, 13-Oct-2016.)
|
⊢ ℲxA ⇒ ⊢ Ⅎx℘A |
|
Theorem | pwidg 3734 |
Membership of the original in a power set. (Contributed by Stefan O'Rear,
1-Feb-2015.)
|
⊢ (A ∈ V →
A ∈
℘A) |
|
Theorem | pwid 3735 |
A set is a member of its power class. Theorem 87 of [Suppes] p. 47.
(Contributed by NM, 5-Aug-1993.)
|
⊢ A ∈ V ⇒ ⊢ A ∈ ℘A |
|
Theorem | pwss 3736* |
Subclass relationship for power class. (Contributed by NM,
21-Jun-2009.)
|
⊢ (℘A ⊆ B ↔ ∀x(x ⊆ A → x
∈ B)) |
|
2.1.16 Unordered and ordered pairs
|
|
Syntax | csn 3737 |
Extend class notation to include singleton.
|
class
{A} |
|
Syntax | cpr 3738 |
Extend class notation to include unordered pair.
|
class
{A, B} |
|
Syntax | ctp 3739 |
Extend class notation to include unordered triplet.
|
class
{A, B, C} |
|
Theorem | snjust 3740* |
Soundness justification theorem for df-sn 3741. (Contributed by Rodolfo
Medina, 28-Apr-2010.) (Proof shortened by Andrew Salmon,
29-Jun-2011.)
|
⊢ {x ∣ x =
A} = {y ∣ y = A} |
|
Definition | df-sn 3741* |
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 3747. (Contributed by NM,
5-Aug-1993.)
|
⊢ {A} =
{x ∣
x = A} |
|
Definition | df-pr 3742 |
Define unordered pair of classes. Definition 7.1 of [Quine] p. 48. For
example, A ∈ { 1 , -u 1 } → (A ^ 2 ) = 1 (ex-pr in set.mm).
They are unordered, so {A,
B} = {B, A} as
proven by prcom 3798.
For a more traditional definition, but requiring a dummy variable, see
dfpr2 3749. (Contributed by NM, 5-Aug-1993.)
|
⊢ {A,
B} = ({A} ∪ {B}) |
|
Definition | df-tp 3743 |
Define unordered triple of classes. Definition of [Enderton] p. 19.
(Contributed by NM, 9-Apr-1994.)
|
⊢ {A,
B, C}
= ({A, B} ∪ {C}) |
|
Theorem | sneq 3744 |
Equality theorem for singletons. Part of Exercise 4 of [TakeutiZaring]
p. 15. (Contributed by NM, 5-Aug-1993.)
|
⊢ (A =
B → {A} = {B}) |
|
Theorem | sneqi 3745 |
Equality inference for singletons. (Contributed by NM, 22-Jan-2004.)
|
⊢ A =
B ⇒ ⊢ {A} =
{B} |
|
Theorem | sneqd 3746 |
Equality deduction for singletons. (Contributed by NM, 22-Jan-2004.)
|
⊢ (φ
→ A = B) ⇒ ⊢ (φ
→ {A} = {B}) |
|
Theorem | dfsn2 3747 |
Alternate definition of singleton. Definition 5.1 of [TakeutiZaring]
p. 15. (Contributed by NM, 24-Apr-1994.)
|
⊢ {A} =
{A, A} |
|
Theorem | elsn 3748* |
There is only one element in a singleton. Exercise 2 of [TakeutiZaring]
p. 15. (Contributed by NM, 5-Aug-1993.)
|
⊢ (x ∈ {A} ↔
x = A) |
|
Theorem | dfpr2 3749* |
Alternate definition of unordered pair. Definition 5.1 of
[TakeutiZaring] p. 15.
(Contributed by NM, 24-Apr-1994.)
|
⊢ {A,
B} = {x ∣ (x = A ∨ x = B)} |
|
Theorem | elprg 3750 |
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.)
|
⊢ (A ∈ V →
(A ∈
{B, C}
↔ (A = B ∨ A = C))) |
|
Theorem | elpr 3751 |
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.)
|
⊢ A ∈ V ⇒ ⊢ (A ∈ {B, C} ↔ (A =
B ∨
A = C)) |
|
Theorem | elpr2 3752 |
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.)
|
⊢ B ∈ V
& ⊢ C ∈ V ⇒ ⊢ (A ∈ {B, C} ↔ (A =
B ∨
A = C)) |
|
Theorem | elpri 3753 |
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.)
|
⊢ (A ∈ {B, C} → (A =
B ∨
A = C)) |
|
Theorem | nelpri 3754 |
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.)
|
⊢ A ≠
B
& ⊢ A ≠
C ⇒ ⊢ ¬ A
∈ {B,
C} |
|
Theorem | elsncg 3755 |
There is only 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.)
|
⊢ (A ∈ V →
(A ∈
{B} ↔ A = B)) |
|
Theorem | elsnc 3756 |
There is only one element in a singleton. Exercise 2 of [TakeutiZaring]
p. 15. (Contributed by NM, 13-Sep-1995.)
|
⊢ A ∈ V ⇒ ⊢ (A ∈ {B} ↔
A = B) |
|
Theorem | elsni 3757 |
There is only one element in a singleton. (Contributed by NM,
5-Jun-1994.)
|
⊢ (A ∈ {B} →
A = B) |
|
Theorem | snidg 3758 |
A set is a member of its singleton. Part of Theorem 7.6 of [Quine] p. 49.
(Contributed by NM, 28-Oct-2003.)
|
⊢ (A ∈ V →
A ∈
{A}) |
|
Theorem | snidb 3759 |
A class is a set iff it is a member of its singleton. (Contributed by NM,
5-Apr-2004.)
|
⊢ (A ∈ V ↔ A
∈ {A}) |
|
Theorem | snid 3760 |
A set is a member of its singleton. Part of Theorem 7.6 of [Quine]
p. 49. (Contributed by NM, 31-Dec-1993.)
|
⊢ A ∈ V ⇒ ⊢ A ∈ {A} |
|
Theorem | elsnc2g 3761 |
There is only one element in a singleton. Exercise 2 of [TakeutiZaring]
p. 15. This variation requires only that B, rather than A, be a
set. (Contributed by NM, 28-Oct-2003.)
|
⊢ (B ∈ V →
(A ∈
{B} ↔ A = B)) |
|
Theorem | elsnc2 3762 |
There is only one element in a singleton. Exercise 2 of [TakeutiZaring]
p. 15. This variation requires only that B, rather than A, be
a set. (Contributed by NM, 12-Jun-1994.)
|
⊢ B ∈ V ⇒ ⊢ (A ∈ {B} ↔
A = B) |
|
Theorem | ralsns 3763* |
Substitution expressed in terms of quantification over a singleton.
(Contributed by Mario Carneiro, 23-Apr-2015.)
|
⊢ (A ∈ V →
(∀x
∈ {A}φ ↔
[̣A / x]̣φ)) |
|
Theorem | rexsns 3764* |
Restricted existential quantification over a singleton. (Contributed by
Mario Carneiro, 23-Apr-2015.)
|
⊢ (A ∈ V →
(∃x
∈ {A}φ ↔
[̣A / x]̣φ)) |
|
Theorem | ralsng 3765* |
Substitution expressed in terms of quantification over a singleton.
(Contributed by NM, 14-Dec-2005.) (Revised by Mario Carneiro,
23-Apr-2015.)
|
⊢ (x =
A → (φ ↔ ψ)) ⇒ ⊢ (A ∈ V →
(∀x
∈ {A}φ ↔
ψ)) |
|
Theorem | rexsng 3766* |
Restricted existential quantification over a singleton. (Contributed by
NM, 29-Jan-2012.)
|
⊢ (x =
A → (φ ↔ ψ)) ⇒ ⊢ (A ∈ V →
(∃x
∈ {A}φ ↔
ψ)) |
|
Theorem | ralsn 3767* |
Convert a quantification over a singleton to a substitution.
(Contributed by NM, 27-Apr-2009.)
|
⊢ A ∈ V
& ⊢ (x =
A → (φ ↔ ψ)) ⇒ ⊢ (∀x ∈ {A}φ ↔
ψ) |
|
Theorem | rexsn 3768* |
Restricted existential quantification over a singleton. (Contributed by
Jeff Madsen, 5-Jan-2011.)
|
⊢ A ∈ V
& ⊢ (x =
A → (φ ↔ ψ)) ⇒ ⊢ (∃x ∈ {A}φ ↔
ψ) |
|
Theorem | eltpg 3769 |
Members of an unordered triple of classes. (Contributed by FL,
2-Feb-2014.) (Proof shortened by Mario Carneiro, 11-Feb-2015.)
|
⊢ (A ∈ V →
(A ∈
{B, C,
D} ↔ (A = B ∨ A = C ∨ A = D))) |
|
Theorem | eltpi 3770 |
A member of an unordered triple of classes is one of them. (Contributed
by Mario Carneiro, 11-Feb-2015.)
|
⊢ (A ∈ {B, C, D} →
(A = B
∨ A =
C ∨
A = D)) |
|
Theorem | eltp 3771 |
A member of an unordered triple of classes is one of them. Special case
of Exercise 1 of [TakeutiZaring]
p. 17. (Contributed by NM,
8-Apr-1994.) (Revised by Mario Carneiro, 11-Feb-2015.)
|
⊢ A ∈ V ⇒ ⊢ (A ∈ {B, C, D} ↔
(A = B
∨ A =
C ∨
A = D)) |
|
Theorem | dftp2 3772* |
Alternate definition of unordered triple of classes. Special case of
Definition 5.3 of [TakeutiZaring]
p. 16. (Contributed by NM,
8-Apr-1994.)
|
⊢ {A,
B, C}
= {x ∣
(x = A
∨ x =
B ∨
x = C)} |
|
Theorem | nfpr 3773 |
Bound-variable hypothesis builder for unordered pairs. (Contributed by
NM, 14-Nov-1995.)
|
⊢ ℲxA & ⊢ ℲxB ⇒ ⊢ Ⅎx{A, B} |
|
Theorem | ifpr 3774 |
Membership of a conditional operator in an unordered pair. (Contributed
by NM, 17-Jun-2007.)
|
⊢ ((A ∈ C ∧ B ∈ D) →
if(φ, A, B) ∈ {A, B}) |
|
Theorem | ralprg 3775* |
Convert a quantification over a pair to a conjunction. (Contributed by
NM, 17-Sep-2011.) (Revised by Mario Carneiro, 23-Apr-2015.)
|
⊢ (x =
A → (φ ↔ ψ))
& ⊢ (x =
B → (φ ↔ χ)) ⇒ ⊢ ((A ∈ V ∧ B ∈ W) →
(∀x
∈ {A,
B}φ ↔ (ψ ∧ χ))) |
|
Theorem | rexprg 3776* |
Convert a quantification over a pair to a disjunction. (Contributed by
NM, 17-Sep-2011.) (Revised by Mario Carneiro, 23-Apr-2015.)
|
⊢ (x =
A → (φ ↔ ψ))
& ⊢ (x =
B → (φ ↔ χ)) ⇒ ⊢ ((A ∈ V ∧ B ∈ W) →
(∃x
∈ {A,
B}φ ↔ (ψ ∨ χ))) |
|
Theorem | raltpg 3777* |
Convert a quantification over a triple to a conjunction. (Contributed
by NM, 17-Sep-2011.) (Revised by Mario Carneiro, 23-Apr-2015.)
|
⊢ (x =
A → (φ ↔ ψ))
& ⊢ (x =
B → (φ ↔ χ))
& ⊢ (x =
C → (φ ↔ θ)) ⇒ ⊢ ((A ∈ V ∧ B ∈ W ∧ C ∈ X) →
(∀x
∈ {A,
B, C}φ ↔
(ψ ∧
χ ∧
θ))) |
|
Theorem | rextpg 3778* |
Convert a quantification over a triple to a disjunction. (Contributed
by Mario Carneiro, 23-Apr-2015.)
|
⊢ (x =
A → (φ ↔ ψ))
& ⊢ (x =
B → (φ ↔ χ))
& ⊢ (x =
C → (φ ↔ θ)) ⇒ ⊢ ((A ∈ V ∧ B ∈ W ∧ C ∈ X) →
(∃x
∈ {A,
B, C}φ ↔
(ψ ∨
χ ∨
θ))) |
|
Theorem | ralpr 3779* |
Convert a quantification over a pair to a conjunction. (Contributed by
NM, 3-Jun-2007.) (Revised by Mario Carneiro, 23-Apr-2015.)
|
⊢ A ∈ V
& ⊢ B ∈ V
& ⊢ (x =
A → (φ ↔ ψ))
& ⊢ (x =
B → (φ ↔ χ)) ⇒ ⊢ (∀x ∈ {A, B}φ ↔ (ψ ∧ χ)) |
|
Theorem | rexpr 3780* |
Convert an existential quantification over a pair to a disjunction.
(Contributed by NM, 3-Jun-2007.) (Revised by Mario Carneiro,
23-Apr-2015.)
|
⊢ A ∈ V
& ⊢ B ∈ V
& ⊢ (x =
A → (φ ↔ ψ))
& ⊢ (x =
B → (φ ↔ χ)) ⇒ ⊢ (∃x ∈ {A, B}φ ↔ (ψ ∨ χ)) |
|
Theorem | raltp 3781* |
Convert a quantification over a triple to a conjunction. (Contributed
by NM, 13-Sep-2011.) (Revised by Mario Carneiro, 23-Apr-2015.)
|
⊢ A ∈ V
& ⊢ B ∈ V
& ⊢ C ∈ V
& ⊢ (x =
A → (φ ↔ ψ))
& ⊢ (x =
B → (φ ↔ χ))
& ⊢ (x =
C → (φ ↔ θ)) ⇒ ⊢ (∀x ∈ {A, B, C}φ ↔
(ψ ∧
χ ∧
θ)) |
|
Theorem | rextp 3782* |
Convert a quantification over a triple to a disjunction. (Contributed
by Mario Carneiro, 23-Apr-2015.)
|
⊢ A ∈ V
& ⊢ B ∈ V
& ⊢ C ∈ V
& ⊢ (x =
A → (φ ↔ ψ))
& ⊢ (x =
B → (φ ↔ χ))
& ⊢ (x =
C → (φ ↔ θ)) ⇒ ⊢ (∃x ∈ {A, B, C}φ ↔
(ψ ∨
χ ∨
θ)) |
|
Theorem | sbcsng 3783* |
Substitution expressed in terms of quantification over a singleton.
(Contributed by NM, 14-Dec-2005.) (Revised by Mario Carneiro,
23-Apr-2015.)
|
⊢ (A ∈ V →
([̣A / x]̣φ
↔ ∀x ∈ {A}φ)) |
|
Theorem | nfsn 3784 |
Bound-variable hypothesis builder for singletons. (Contributed by NM,
14-Nov-1995.)
|
⊢ ℲxA ⇒ ⊢ Ⅎx{A} |
|
Theorem | csbsng 3785 |
Distribute proper substitution through the singleton of a class.
csbsng 3785 is derived from the virtual deduction proof
csbsngVD in set.mm.
(Contributed by Alan Sare, 10-Nov-2012.)
|
⊢ (A ∈ V →
[A / x]{B} =
{[A / x]B}) |
|
Theorem | disjsn 3786 |
Intersection with the singleton of a non-member is disjoint.
(Contributed by NM, 22-May-1998.) (Proof shortened by Andrew Salmon,
29-Jun-2011.) (Proof shortened by Wolf Lammen, 30-Sep-2014.)
|
⊢ ((A ∩
{B}) = ∅ ↔ ¬ B ∈ A) |
|
Theorem | disjsn2 3787 |
Intersection of distinct singletons is disjoint. (Contributed by NM,
25-May-1998.)
|
⊢ (A ≠
B → ({A} ∩ {B}) =
∅) |
|
Theorem | snprc 3788 |
The singleton of a proper class (one that doesn't exist) is the empty
set. Theorem 7.2 of [Quine] p. 48.
(Contributed by NM, 5-Aug-1993.)
|
⊢ (¬ A
∈ V ↔ {A} = ∅) |
|
Theorem | r19.12sn 3789* |
Special case of r19.12 2727 where its converse holds. (Contributed by
NM,
19-May-2008.) (Revised by Mario Carneiro, 23-Apr-2015.)
|
⊢ A ∈ V ⇒ ⊢ (∃x ∈ {A}∀y ∈ B φ ↔
∀y
∈ B
∃x
∈ {A}φ) |
|
Theorem | rabsn 3790* |
Condition where a restricted class abstraction is a singleton.
(Contributed by NM, 28-May-2006.)
|
⊢ (B ∈ A →
{x ∈
A ∣
x = B}
= {B}) |
|
Theorem | euabsn2 3791* |
Another way to express existential uniqueness of a wff: its class
abstraction is a singleton. (Contributed by Mario Carneiro,
14-Nov-2016.)
|
⊢ (∃!xφ ↔
∃y{x ∣ φ} =
{y}) |
|
Theorem | euabsn 3792 |
Another way to express existential uniqueness of a wff: its class
abstraction is a singleton. (Contributed by NM, 22-Feb-2004.)
|
⊢ (∃!xφ ↔
∃x{x ∣ φ} =
{x}) |
|
Theorem | reusn 3793* |
A way to express restricted existential uniqueness of a wff: its
restricted class abstraction is a singleton. (Contributed by NM,
30-May-2006.) (Proof shortened by Mario Carneiro, 14-Nov-2016.)
|
⊢ (∃!x ∈ A φ ↔
∃y{x ∈ A ∣ φ} =
{y}) |
|
Theorem | absneu 3794 |
Restricted existential uniqueness determined by a singleton.
(Contributed by NM, 29-May-2006.)
|
⊢ ((A ∈ V ∧ {x ∣ φ} =
{A}) → ∃!xφ) |
|
Theorem | rabsneu 3795 |
Restricted existential uniqueness determined by a singleton.
(Contributed by NM, 29-May-2006.) (Revised by Mario Carneiro,
23-Dec-2016.)
|
⊢ ((A ∈ V ∧ {x ∈ B ∣ φ} =
{A}) → ∃!x ∈ B φ) |
|
Theorem | eusn 3796* |
Two ways to express "A
is a singleton." (Contributed by NM,
30-Oct-2010.)
|
⊢ (∃!x x ∈ A ↔
∃x
A = {x}) |
|
Theorem | rabsnt 3797* |
Truth implied by equality of a restricted class abstraction and a
singleton. (Contributed by NM, 29-May-2006.) (Proof shortened by Mario
Carneiro, 23-Dec-2016.)
|
⊢ B ∈ V
& ⊢ (x =
B → (φ ↔ ψ)) ⇒ ⊢ ({x ∈ A ∣ φ} =
{B} → ψ) |
|
Theorem | prcom 3798 |
Commutative law for unordered pairs. (Contributed by NM, 5-Aug-1993.)
|
⊢ {A,
B} = {B, A} |
|
Theorem | preq1 3799 |
Equality theorem for unordered pairs. (Contributed by NM,
29-Mar-1998.)
|
⊢ (A =
B → {A, C} =
{B, C}) |
|
Theorem | preq2 3800 |
Equality theorem for unordered pairs. (Contributed by NM, 5-Aug-1993.)
|
⊢ (A =
B → {C, A} =
{C, B}) |