HomeHome New Foundations Explorer
Theorem List (p. 47 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 - 4601-4700   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremproj2op 4601 The second projection operator applied to an ordered pair yields its second member. Theorem X.2.8 of [Rosser] p. 283. (Contributed by SF, 3-Feb-2015.)
Proj2 A, B = B
 
Theoremopth 4602 The ordered pair theorem. Two ordered pairs are equal iff their components are equal. (Contributed by SF, 2-Jan-2015.)
(A, B = C, D ↔ (A = C B = D))
 
Theoremopexb 4603 An ordered pair is a set iff its components are sets. (Contributed by SF, 2-Jan-2015.)
(A, B V ↔ (A V B V))
 
Theoremnfop 4604 Bound-variable hypothesis builder for ordered pairs. (Contributed by SF, 2-Jan-2015.)
xA    &   xB       xA, B
 
Theoremnfopd 4605 Deduction version of bound-variable hypothesis builder nfop 4604. (Contributed by SF, 2-Jan-2015.)
(φxA)    &   (φxB)       (φxA, B)
 
Theoremeqvinop 4606* A variable introduction law for ordered pairs. Analog of Lemma 15 of [Monk2] p. 109. (Contributed by NM, 28-May-1995.)
B V    &   C V       (A = B, Cxy(A = x, y x, y = B, C))
 
Theoremcopsexg 4607* Substitution of class A for ordered pair x, y. (Contributed by NM, 27-Dec-1996.) (Revised by Andrew Salmon, 25-Jul-2011.)
(A = x, y → (φxy(A = x, y φ)))
 
Theoremcopsex2t 4608* Closed theorem form of copsex2g 4609. (Contributed by NM, 17-Feb-2013.)
((xy((x = A y = B) → (φψ)) (A V B W)) → (xy(A, B = x, y φ) ↔ ψ))
 
Theoremcopsex2g 4609* Implicit substitution inference for ordered pairs. (Contributed by NM, 28-May-1995.)
((x = A y = B) → (φψ))       ((A V B W) → (xy(A, B = x, y φ) ↔ ψ))
 
Theoremcopsex4g 4610* An implicit substitution inference for 2 ordered pairs. (Contributed by NM, 5-Aug-1995.)
(((x = A y = B) (z = C w = D)) → (φψ))       (((A R B S) (C R D S)) → (xyzw((A, B = x, y C, D = z, w) φ) ↔ ψ))
 
Theoremeqop 4611* Express equality to an ordered pair. (Contributed by SF, 6-Jan-2015.)
(A = B, Cz(z A ↔ (t B z = Phi t t C z = ( Phi t ∪ {0c}))))
 
Theoremmosubopt 4612* "At most one" remains true inside ordered pair quantification. (Contributed by NM, 28-Aug-2007.)
(yz∃*xφ∃*xyz(A = y, z φ))
 
Theoremmosubop 4613* "At most one" remains true inside ordered pair quantification. (Contributed by NM, 28-May-1995.)
∃*xφ       ∃*xyz(A = y, z φ)
 
Theoremphiun 4614 The phi operation distributes over union. (Contributed by SF, 20-Feb-2015.)
Phi (AB) = ( Phi A Phi B)
 
Theoremphidisjnn 4615 The phi operation applied to a set disjoint from the naturals has no effect. (Contributed by SF, 20-Feb-2015.)
((ANn ) = Phi A = A)
 
Theoremphialllem1 4616* Lemma for phiall 4618. Any set of numbers without zero is the Phi of a set. (Contributed by Scott Fenton, 14-Apr-2021.)
A V       ((A Nn ¬ 0c A) → x A = Phi x)
 
Theoremphialllem2 4617* Lemma for phiall 4618. Any set without 0c is equal to the Phi of a set. (Contributed by Scott Fenton, 8-Apr-2021.)
A V       (¬ 0c Ax A = Phi x)
 
Theoremphiall 4618* Any set is equal to either the Phi of another set or to a Phi with 0c adjoined. (Contributed by Scott Fenton, 8-Apr-2021.)
A V       x(A = Phi x A = ( Phi x ∪ {0c}))
 
Theoremopeq 4619 Any class is equal to an ordered pair. (Contributed by Scott Fenton, 8-Apr-2021.)
A = Proj1 A, Proj2 A
 
Theoremopeqexb 4620* A class is a set iff it is equal to an ordered pair. (Contributed by Scott Fenton, 19-Apr-2021.)
(A V ↔ xy A = x, y)
 
Theoremopeqex 4621* Any set is equal to some ordered pair. (Contributed by Scott Fenton, 16-Apr-2021.)
(A Vxy A = x, y)
 
2.3.2  Ordered-pair class abstractions (class builders)
 
Syntaxcopab 4622 Extend class notation to include ordered-pair class abstraction (class builder).
class {x, y φ}
 
Definitiondf-opab 4623* Define the class abstraction of a collection of ordered pairs. Definition 3.3 of [Monk1] p. 34. Usually x and y are distinct, although the definition doesn't strictly require it (see dfid2 4769 for a case where they are not distinct). The brace notation is called "class abstraction" by Quine; it is also (more commonly) called a "class builder" in the literature. (Contributed by SF, 12-Jan-2015.)
{x, y φ} = {z xy(z = x, y φ)}
 
Theoremopabbid 4624 Equivalent wff's yield equal ordered-pair class abstractions (deduction rule). (Contributed by NM, 21-Feb-2004.) (Proof shortened by Andrew Salmon, 9-Jul-2011.)
xφ    &   yφ    &   (φ → (ψχ))       (φ → {x, y ψ} = {x, y χ})
 
Theoremopabbidv 4625* Equivalent wff's yield equal ordered-pair class abstractions (deduction rule). (Contributed by NM, 15-May-1995.)
(φ → (ψχ))       (φ → {x, y ψ} = {x, y χ})
 
Theoremopabbii 4626 Equivalent wff's yield equal class abstractions. (Contributed by NM, 15-May-1995.)
(φψ)       {x, y φ} = {x, y ψ}
 
Theoremnfopab 4627* Bound-variable hypothesis builder for class abstraction. (Contributed by NM, 1-Sep-1999.) (Unnecessary distinct variable restrictions were removed by Andrew Salmon, 11-Jul-2011.)
zφ       z{x, y φ}
 
Theoremnfopab1 4628 The first abstraction variable in an ordered-pair class abstraction (class builder) is effectively not free. (Contributed by NM, 16-May-1995.) (Revised by Mario Carneiro, 14-Oct-2016.)
x{x, y φ}
 
Theoremnfopab2 4629 The second abstraction variable in an ordered-pair class abstraction (class builder) is effectively not free. (Contributed by NM, 16-May-1995.) (Revised by Mario Carneiro, 14-Oct-2016.)
y{x, y φ}
 
Theoremcbvopab 4630* Rule used to change bound variables in an ordered-pair class abstraction, using implicit substitution. (Contributed by NM, 14-Sep-2003.)
zφ    &   wφ    &   xψ    &   yψ    &   ((x = z y = w) → (φψ))       {x, y φ} = {z, w ψ}
 
Theoremcbvopabv 4631* Rule used to change bound variables in an ordered-pair class abstraction, using implicit substitution. (Contributed by NM, 15-Oct-1996.)
((x = z y = w) → (φψ))       {x, y φ} = {z, w ψ}
 
Theoremcbvopab1 4632* Change first bound variable in an ordered-pair class abstraction, using explicit substitution. (Contributed by NM, 6-Oct-2004.) (Revised by Mario Carneiro, 14-Oct-2016.)
zφ    &   xψ    &   (x = z → (φψ))       {x, y φ} = {z, y ψ}
 
Theoremcbvopab2 4633* Change second bound variable in an ordered-pair class abstraction, using explicit substitution. (Contributed by NM, 22-Aug-2013.)
zφ    &   yψ    &   (y = z → (φψ))       {x, y φ} = {x, z ψ}
 
Theoremcbvopab1s 4634* Change first bound variable in an ordered-pair class abstraction, using explicit substitution. (Contributed by NM, 31-Jul-2003.)
{x, y φ} = {z, y [z / x]φ}
 
Theoremcbvopab1v 4635* Rule used to change the first bound variable in an ordered pair abstraction, using implicit substitution. (Contributed by NM, 31-Jul-2003.) (Proof shortened by Eric Schmidt, 4-Apr-2007.)
(x = z → (φψ))       {x, y φ} = {z, y ψ}
 
Theoremcbvopab2v 4636* Rule used to change the second bound variable in an ordered pair abstraction, using implicit substitution. (Contributed by NM, 2-Sep-1999.)
(y = z → (φψ))       {x, y φ} = {x, z ψ}
 
Theoremcsbopabg 4637* Move substitution into a class abstraction. (Contributed by NM, 6-Aug-2007.) (Proof shortened by Mario Carneiro, 17-Nov-2016.)
(A V[A / x]{y, z φ} = {y, z A / xφ})
 
Theoremunopab 4638 Union of two ordered pair class abstractions. (Contributed by NM, 30-Sep-2002.)
({x, y φ} ∪ {x, y ψ}) = {x, y (φ ψ)}
 
2.3.3  Binary relations
 
Syntaxwbr 4639 Extend wff notation to include the general binary relation predicate. Note that the syntax is simply three class symbols in a row. Since binary relations are the only possible wff expressions consisting of three class expressions in a row, the syntax is unambiguous.
wff ARB
 
Definitiondf-br 4640 Define a general binary relation. Note that the syntax is simply three class symbols in a row. Definition 6.18 of [TakeutiZaring] p. 29 generalized to arbitrary classes. Class R normally denotes a relation that compares two classes A and B. This definition is well-defined, although not very meaningful, when classes A and/or B are proper classes (i.e. are not sets). On the other hand, we often find uses for this definition when R is a proper class. (Contributed by NM, 4-Jun-1995.)
(ARBA, B R)
 
Theorembreq 4641 Equality theorem for binary relations. (Contributed by NM, 4-Jun-1995.)
(R = S → (ARBASB))
 
Theorembreq1 4642 Equality theorem for a binary relation. (Contributed by NM, 31-Dec-1993.)
(A = B → (ARCBRC))
 
Theorembreq2 4643 Equality theorem for a binary relation. (Contributed by NM, 31-Dec-1993.)
(A = B → (CRACRB))
 
Theorembreq12 4644 Equality theorem for a binary relation. (Contributed by NM, 8-Feb-1996.)
((A = B C = D) → (ARCBRD))
 
Theorembreqi 4645 Equality inference for binary relations. (Contributed by NM, 19-Feb-2005.)
R = S       (ARBASB)
 
Theorembreq1i 4646 Equality inference for a binary relation. (Contributed by NM, 8-Feb-1996.)
A = B       (ARCBRC)
 
Theorembreq2i 4647 Equality inference for a binary relation. (Contributed by NM, 8-Feb-1996.)
A = B       (CRACRB)
 
Theorembreq12i 4648 Equality inference for a binary relation. (Contributed by NM, 8-Feb-1996.) (Revised by Eric Schmidt, 4-Apr-2007.)
A = B    &   C = D       (ARCBRD)
 
Theorembreq1d 4649 Equality deduction for a binary relation. (Contributed by NM, 8-Feb-1996.)
(φA = B)       (φ → (ARCBRC))
 
Theorembreqd 4650 Equality deduction for a binary relation. (Contributed by NM, 29-Oct-2011.)
(φA = B)       (φ → (CADCBD))
 
Theorembreq2d 4651 Equality deduction for a binary relation. (Contributed by NM, 8-Feb-1996.)
(φA = B)       (φ → (CRACRB))
 
Theorembreq12d 4652 Equality deduction for a binary relation. (The proof was shortened by Andrew Salmon, 9-Jul-2011.) (Contributed by NM, 8-Feb-1996.) (Revised by set.mm contributors, 9-Jul-2011.)
(φA = B)    &   (φC = D)       (φ → (ARCBRD))
 
Theorembreq123d 4653 Equality deduction for a binary relation. (Contributed by NM, 29-Oct-2011.)
(φA = B)    &   (φR = S)    &   (φC = D)       (φ → (ARCBSD))
 
Theorembreqan12d 4654 Equality deduction for a binary relation. (Contributed by NM, 8-Feb-1996.)
(φA = B)    &   (ψC = D)       ((φ ψ) → (ARCBRD))
 
Theorembreqan12rd 4655 Equality deduction for a binary relation. (Contributed by NM, 8-Feb-1996.)
(φA = B)    &   (ψC = D)       ((ψ φ) → (ARCBRD))
 
Theoremnbrne1 4656 Two classes are different if they don't have the same relationship to a third class. (Contributed by NM, 3-Jun-2012.)
((ARB ¬ ARC) → BC)
 
Theoremnbrne2 4657 Two classes are different if they don't have the same relationship to a third class. (Contributed by NM, 3-Jun-2012.)
((ARC ¬ BRC) → AB)
 
Theoremeqbrtri 4658 Substitution of equal classes into a binary relation. (Contributed by NM, 5-Aug-1993.)
A = B    &   BRC       ARC
 
Theoremeqbrtrd 4659 Substitution of equal classes into a binary relation. (Contributed by NM, 8-Oct-1999.)
(φA = B)    &   (φBRC)       (φARC)
 
Theoremeqbrtrri 4660 Substitution of equal classes into a binary relation. (Contributed by NM, 5-Aug-1993.)
A = B    &   ARC       BRC
 
Theoremeqbrtrrd 4661 Substitution of equal classes into a binary relation. (Contributed by NM, 24-Oct-1999.)
(φA = B)    &   (φARC)       (φBRC)
 
Theorembreqtri 4662 Substitution of equal classes into a binary relation. (Contributed by NM, 5-Aug-1993.)
ARB    &   B = C       ARC
 
Theorembreqtrd 4663 Substitution of equal classes into a binary relation. (Contributed by NM, 24-Oct-1999.)
(φARB)    &   (φB = C)       (φARC)
 
Theorembreqtrri 4664 Substitution of equal classes into a binary relation. (Contributed by NM, 5-Aug-1993.)
ARB    &   C = B       ARC
 
Theorembreqtrrd 4665 Substitution of equal classes into a binary relation. (Contributed by NM, 24-Oct-1999.)
(φARB)    &   (φC = B)       (φARC)
 
Theorem3brtr3i 4666 Substitution of equality into both sides of a binary relation. (Contributed by NM, 11-Aug-1999.)
ARB    &   A = C    &   B = D       CRD
 
Theorem3brtr4i 4667 Substitution of equality into both sides of a binary relation. (Contributed by NM, 11-Aug-1999.)
ARB    &   C = A    &   D = B       CRD
 
Theorem3brtr3d 4668 Substitution of equality into both sides of a binary relation. (Contributed by NM, 18-Oct-1999.)
(φARB)    &   (φA = C)    &   (φB = D)       (φCRD)
 
Theorem3brtr4d 4669 Substitution of equality into both sides of a binary relation. (Contributed by NM, 21-Feb-2005.)
(φARB)    &   (φC = A)    &   (φD = B)       (φCRD)
 
Theorem3brtr3g 4670 Substitution of equality into both sides of a binary relation. (Contributed by NM, 16-Jan-1997.)
(φARB)    &   A = C    &   B = D       (φCRD)
 
Theorem3brtr4g 4671 Substitution of equality into both sides of a binary relation. (Contributed by NM, 16-Jan-1997.)
(φARB)    &   C = A    &   D = B       (φCRD)
 
Theoremsyl5eqbr 4672 B chained equality inference for a binary relation. (Contributed by NM, 11-Oct-1999.)
A = B    &   (φBRC)       (φARC)
 
Theoremsyl5eqbrr 4673 B chained equality inference for a binary relation. (Contributed by NM, 17-Sep-2004.)
B = A    &   (φBRC)       (φARC)
 
Theoremsyl5breq 4674 B chained equality inference for a binary relation. (Contributed by NM, 11-Oct-1999.)
ARB    &   (φB = C)       (φARC)
 
Theoremsyl5breqr 4675 B chained equality inference for a binary relation. (Contributed by NM, 24-Apr-2005.)
ARB    &   (φC = B)       (φARC)
 
Theoremsyl6eqbr 4676 A chained equality inference for a binary relation. (Contributed by NM, 12-Oct-1999.)
(φA = B)    &   BRC       (φARC)
 
Theoremsyl6eqbrr 4677 A chained equality inference for a binary relation. (Contributed by NM, 4-Jan-2006.)
(φB = A)    &   BRC       (φARC)
 
Theoremsyl6breq 4678 A chained equality inference for a binary relation. (Contributed by NM, 11-Oct-1999.)
(φARB)    &   B = C       (φARC)
 
Theoremsyl6breqr 4679 A chained equality inference for a binary relation. (Contributed by NM, 24-Apr-2005.)
(φARB)    &   C = B       (φARC)
 
Theoremssbrd 4680 Deduction from a subclass relationship of binary relations. (Contributed by NM, 30-Apr-2004.)
(φA B)       (φ → (CADCBD))
 
Theoremssbri 4681 Inference from a subclass relationship of binary relations. (The proof was shortened by Andrew Salmon, 9-Jul-2011.) (Contributed by NM, 28-Mar-2007.) (Revised by set.mm contributors, 9-Jul-2011.)
A B       (CADCBD)
 
Theoremnfbrd 4682 Deduction version of bound-variable hypothesis builder nfbr 4683. (Contributed by NM, 13-Dec-2005.) (Revised by Mario Carneiro, 14-Oct-2016.)
(φxA)    &   (φxR)    &   (φxB)       (φ → Ⅎx ARB)
 
Theoremnfbr 4683 Bound-variable hypothesis builder for binary relation. (Contributed by NM, 1-Sep-1999.) (Revised by Mario Carneiro, 14-Oct-2016.)
xA    &   xR    &   xB       x ARB
 
Theorembrab1 4684* Relationship between a binary relation and a class abstraction. (Contributed by Andrew Salmon, 8-Jul-2011.)
(xRAx {z zRA})
 
Theoremsbcbrg 4685 Move substitution in and out of a binary relation. (Contributed by NM, 13-Dec-2005.) (Proof shortened by Andrew Salmon, 9-Jul-2011.)
(A D → ([̣A / xBRC[A / x]B[A / x]R[A / x]C))
 
Theoremsbcbr12g 4686* Move substitution in and out of a binary relation. (Contributed by NM, 13-Dec-2005.)
(A D → ([̣A / xBRC[A / x]BR[A / x]C))
 
Theoremsbcbr1g 4687* Move substitution in and out of a binary relation. (Contributed by NM, 13-Dec-2005.)
(A D → ([̣A / xBRC[A / x]BRC))
 
Theoremsbcbr2g 4688* Move substitution in and out of a binary relation. (Contributed by NM, 13-Dec-2005.)
(A D → ([̣A / xBRCBR[A / x]C))
 
Theorembrex 4689 Binary relationship implies sethood of both parts. (Contributed by SF, 7-Jan-2015.)
(ARB → (A V B V))
 
Theorembrreldmex 4690 Binary relationship implies sethood of domain. (Contributed by SF, 7-Jan-2018.)
(ARBA V)
 
Theorembrrelrnex 4691 Binary relationship implies sethood of range. (Contributed by SF, 7-Jan-2018.)
(ARBB V)
 
Theorembrun 4692 The union of two binary relations. (Contributed by NM, 21-Dec-2008.)
(A(RS)B ↔ (ARB ASB))
 
Theorembrin 4693 The intersection of two relations. (Contributed by FL, 7-Oct-2008.)
(A(RS)B ↔ (ARB ASB))
 
Theorembrdif 4694 The difference of two binary relations. (Contributed by Scott Fenton, 11-Apr-2011.)
(A(R S)B ↔ (ARB ¬ ASB))
 
2.3.4  Ordered-pair class abstractions (cont.)
 
Theoremopabid 4695 The law of concretion. Special case of Theorem 9.5 of [Quine] p. 61. (The proof was shortened by Andrew Salmon, 25-Jul-2011.) (Contributed by NM, 14-Apr-1995.) (Revised by set.mm contributors, 25-Jul-2011.)
(x, y {x, y φ} ↔ φ)
 
Theoremelopab 4696* Membership in a class abstraction of pairs. (Contributed by NM, 24-Mar-1998.)
(A {x, y φ} ↔ xy(A = x, y φ))
 
Theoremopelopabsb 4697* The law of concretion in terms of substitutions. (Contributed by NM, 30-Sep-2002.) (Revised by Mario Carneiro, 18-Nov-2016.)
(A, B {x, y φ} ↔ [̣A / x]̣[̣B / yφ)
 
Theorembrabsb 4698* The law of concretion in terms of substitutions. (Contributed by NM, 17-Mar-2008.)
R = {x, y φ}       (ARB ↔ [̣A / x]̣[̣B / yφ)
 
Theoremopelopabt 4699* Closed theorem form of opelopab 4708. (Contributed by NM, 19-Feb-2013.)
((xy(x = A → (φψ)) xy(y = B → (ψχ)) (A V B W)) → (A, B {x, y φ} ↔ χ))
 
Theoremopelopabga 4700* The law of concretion. Theorem 9.5 of [Quine] p. 61. (Contributed by Mario Carneiro, 19-Dec-2013.)
((x = A y = B) → (φψ))       ((A V B W) → (A, B {x, y φ} ↔ ψ))
    < 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 >