HomeHome New Foundations Explorer
Theorem List (p. 38 of 64)
< Previous  Next >
Bad symbols? Try the
GIF version.

Mirrors  >  Metamath Home Page  >  NFE Home Page  >  Theorem List Contents       This page: Page List

Theorem List for New Foundations Explorer - 3701-3800   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremifan 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)
 
Theoremifor 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))
 
Theoremdedth 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 http://us.metamath.org/mpeuni/mmdeduction.html. (Contributed by NM, 15-May-1999.)
(A = if(φ, A, B) → (ψχ))    &   χ       (φψ)
 
Theoremdedth2h 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) → (θτ))    &   τ       ((φ ψ) → χ)
 
Theoremdedth3h 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) → (ηζ))    &   ζ       ((φ ψ χ) → θ)
 
Theoremdedth4h 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) → (σρ))    &   ρ       (((φ ψ) (χ θ)) → τ)
 
Theoremdedth2v 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) → (χθ))    &   θ       (φψ)
 
Theoremdedth3v 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) → (θτ))    &   τ       (φψ)
 
Theoremdedth4v 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) → (τη))    &   η       (φψ)
 
Theoremelimhyp 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) → (χψ))    &   χ       ψ
 
Theoremelimhyp2v 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) → (ηθ))    &   τ       θ
 
Theoremelimhyp3v 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) → (στ))    &   η       τ
 
Theoremelimhyp4v 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) → (ρψ))    &   η       ψ
 
Theoremelimel 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
 
Theoremelimdhyp 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) → (θχ))    &   θ       χ
 
Theoremkeephyp 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) → (χθ))    &   ψ    &   χ       θ
 
Theoremkeephyp2v 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) → (ηθ))    &   ψ    &   τ       θ
 
Theoremkeephyp3v 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) → (στ))    &   ρ    &   η       τ
 
Theoremkeepel 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
 
Theoremifex 3720 Conditional operator existence. (Contributed by NM, 2-Sep-2004.)
A V    &   B V        if(φ, A, B) V
 
Theoremifexg 3721 Conditional operator existence. (Contributed by NM, 21-Mar-2011.)
((A V B W) → if(φ, A, B) V)
 
2.1.15  Power classes
 
Syntaxcpw 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
 
Theorempwjust 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}
 
Definitiondf-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}
 
Theorempweq 3725 Equality theorem for power class. (Contributed by NM, 5-Aug-1993.)
(A = BA = B)
 
Theorempweqi 3726 Equality inference for power class. (Contributed by NM, 27-Nov-2013.)
A = B       A = B
 
Theorempweqd 3727 Equality deduction for power class. (Contributed by NM, 27-Nov-2013.)
(φA = B)       (φA = B)
 
Theoremelpw 3728 Membership in a power class. Theorem 86 of [Suppes] p. 47. (Contributed by NM, 31-Dec-1993.)
A V       (A BA B)
 
Theoremelpwg 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 BA B))
 
Theoremelpwi 3730 Subset relation implied by membership in a power class. (Contributed by NM, 17-Feb-2007.)
(A BA B)
 
Theoremelpwid 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)
 
Theoremelelpwi 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)
 
Theoremnfpw 3733 Bound-variable hypothesis builder for power class. (Contributed by NM, 28-Oct-2003.) (Revised by Mario Carneiro, 13-Oct-2016.)
xA       xA
 
Theorempwidg 3734 Membership of the original in a power set. (Contributed by Stefan O'Rear, 1-Feb-2015.)
(A VA A)
 
Theorempwid 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
 
Theorempwss 3736* Subclass relationship for power class. (Contributed by NM, 21-Jun-2009.)
(A Bx(x Ax B))
 
2.1.16  Unordered and ordered pairs
 
Syntaxcsn 3737 Extend class notation to include singleton.
class {A}
 
Syntaxcpr 3738 Extend class notation to include unordered pair.
class {A, B}
 
Syntaxctp 3739 Extend class notation to include unordered triplet.
class {A, B, C}
 
Theoremsnjust 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}
 
Definitiondf-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}
 
Definitiondf-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})
 
Definitiondf-tp 3743 Define unordered triple of classes. Definition of [Enderton] p. 19. (Contributed by NM, 9-Apr-1994.)
{A, B, C} = ({A, B} ∪ {C})
 
Theoremsneq 3744 Equality theorem for singletons. Part of Exercise 4 of [TakeutiZaring] p. 15. (Contributed by NM, 5-Aug-1993.)
(A = B → {A} = {B})
 
Theoremsneqi 3745 Equality inference for singletons. (Contributed by NM, 22-Jan-2004.)
A = B       {A} = {B}
 
Theoremsneqd 3746 Equality deduction for singletons. (Contributed by NM, 22-Jan-2004.)
(φA = B)       (φ → {A} = {B})
 
Theoremdfsn2 3747 Alternate definition of singleton. Definition 5.1 of [TakeutiZaring] p. 15. (Contributed by NM, 24-Apr-1994.)
{A} = {A, A}
 
Theoremelsn 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)
 
Theoremdfpr2 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)}
 
Theoremelprg 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)))
 
Theoremelpr 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))
 
Theoremelpr2 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))
 
Theoremelpri 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))
 
Theoremnelpri 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.)
AB    &   AC        ¬ A {B, C}
 
Theoremelsncg 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))
 
Theoremelsnc 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)
 
Theoremelsni 3757 There is only one element in a singleton. (Contributed by NM, 5-Jun-1994.)
(A {B} → A = B)
 
Theoremsnidg 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 VA {A})
 
Theoremsnidb 3759 A class is a set iff it is a member of its singleton. (Contributed by NM, 5-Apr-2004.)
(A V ↔ A {A})
 
Theoremsnid 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}
 
Theoremelsnc2g 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))
 
Theoremelsnc2 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)
 
Theoremralsns 3763* Substitution expressed in terms of quantification over a singleton. (Contributed by Mario Carneiro, 23-Apr-2015.)
(A V → (x {A}φ ↔ [̣A / xφ))
 
Theoremrexsns 3764* Restricted existential quantification over a singleton. (Contributed by Mario Carneiro, 23-Apr-2015.)
(A V → (x {A}φ ↔ [̣A / xφ))
 
Theoremralsng 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}φψ))
 
Theoremrexsng 3766* Restricted existential quantification over a singleton. (Contributed by NM, 29-Jan-2012.)
(x = A → (φψ))       (A V → (x {A}φψ))
 
Theoremralsn 3767* Convert a quantification over a singleton to a substitution. (Contributed by NM, 27-Apr-2009.)
A V    &   (x = A → (φψ))       (x {A}φψ)
 
Theoremrexsn 3768* Restricted existential quantification over a singleton. (Contributed by Jeff Madsen, 5-Jan-2011.)
A V    &   (x = A → (φψ))       (x {A}φψ)
 
Theoremeltpg 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)))
 
Theoremeltpi 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))
 
Theoremeltp 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))
 
Theoremdftp2 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)}
 
Theoremnfpr 3773 Bound-variable hypothesis builder for unordered pairs. (Contributed by NM, 14-Nov-1995.)
xA    &   xB       x{A, B}
 
Theoremifpr 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})
 
Theoremralprg 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}φ ↔ (ψ χ)))
 
Theoremrexprg 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}φ ↔ (ψ χ)))
 
Theoremraltpg 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}φ ↔ (ψ χ θ)))
 
Theoremrextpg 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}φ ↔ (ψ χ θ)))
 
Theoremralpr 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}φ ↔ (ψ χ))
 
Theoremrexpr 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}φ ↔ (ψ χ))
 
Theoremraltp 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}φ ↔ (ψ χ θ))
 
Theoremrextp 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}φ ↔ (ψ χ θ))
 
Theoremsbcsng 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}φ))
 
Theoremnfsn 3784 Bound-variable hypothesis builder for singletons. (Contributed by NM, 14-Nov-1995.)
xA       x{A}
 
Theoremcsbsng 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})
 
Theoremdisjsn 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)
 
Theoremdisjsn2 3787 Intersection of distinct singletons is disjoint. (Contributed by NM, 25-May-1998.)
(AB → ({A} ∩ {B}) = )
 
Theoremsnprc 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} = )
 
Theoremr19.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}φ)
 
Theoremrabsn 3790* Condition where a restricted class abstraction is a singleton. (Contributed by NM, 28-May-2006.)
(B A → {x A x = B} = {B})
 
Theoremeuabsn2 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})
 
Theoremeuabsn 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})
 
Theoremreusn 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})
 
Theoremabsneu 3794 Restricted existential uniqueness determined by a singleton. (Contributed by NM, 29-May-2006.)
((A V {x φ} = {A}) → ∃!xφ)
 
Theoremrabsneu 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 φ)
 
Theoremeusn 3796* Two ways to express "A is a singleton." (Contributed by NM, 30-Oct-2010.)
(∃!x x Ax A = {x})
 
Theoremrabsnt 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} → ψ)
 
Theoremprcom 3798 Commutative law for unordered pairs. (Contributed by NM, 5-Aug-1993.)
{A, B} = {B, A}
 
Theorempreq1 3799 Equality theorem for unordered pairs. (Contributed by NM, 29-Mar-1998.)
(A = B → {A, C} = {B, C})
 
Theorempreq2 3800 Equality theorem for unordered pairs. (Contributed by NM, 5-Aug-1993.)
(A = B → {C, A} = {C, B})
    < Previous  Next >

Page List
Jump to page: Contents  1 1-100 2 101-200 3 201-300 4 301-400 5 401-500 6 501-600 7 601-700 8 701-800 9 801-900 10 901-1000 11 1001-1100 12 1101-1200 13 1201-1300 14 1301-1400 15 1401-1500 16 1501-1600 17 1601-1700 18 1701-1800 19 1801-1900 20 1901-2000 21 2001-2100 22 2101-2200 23 2201-2300 24 2301-2400 25 2401-2500 26 2501-2600 27 2601-2700 28 2701-2800 29 2801-2900 30 2901-3000 31 3001-3100 32 3101-3200 33 3201-3300 34 3301-3400 35 3401-3500 36 3501-3600 37 3601-3700 38 3701-3800 39 3801-3900 40 3901-4000 41 4001-4100 42 4101-4200 43 4201-4300 44 4301-4400 45 4401-4500 46 4501-4600 47 4601-4700 48 4701-4800 49 4801-4900 50 4901-5000 51 5001-5100 52 5101-5200 53 5201-5300 54 5301-5400 55 5401-5500 56 5501-5600 57 5601-5700 58 5701-5800 59 5801-5900 60 5901-6000 61 6001-6100 62 6101-6200 63 6201-6300 64 6301-6329
  Copyright terms: Public domain < Previous  Next >