HomeHome New Foundations Explorer
Theorem List (p. 60 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 - 5901-6000   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Definitiondf-antisym 5901* Define the set of all antisymmetric relationships over a base set. (Contributed by SF, 19-Feb-2015.)
Antisym = {r, a x a y a ((xry yrx) → x = y)}
 
Definitiondf-partial 5902 Define the set of all partial orderings over a base set. (Contributed by SF, 19-Feb-2015.)
Po = (( RefTrans ) ∩ Antisym )
 
Definitiondf-connex 5903* Define the set of all connected relationships over a base set. (Contributed by SF, 19-Feb-2015.)
Connex = {r, a x a y a (xry yrx)}
 
Definitiondf-strict 5904 Define the set of all strict orderings over a base set. (Contributed by SF, 19-Feb-2015.)
Or = ( PoConnex )
 
Definitiondf-found 5905* Define the set of all founded relationships over a base set. (Contributed by SF, 19-Feb-2015.)
Fr = {r, a x((x a x) → z x y x (yrzy = z))}
 
Definitiondf-we 5906 Define the set of all well orderings over a base set. (Contributed by SF, 19-Feb-2015.)
We = ( OrFr )
 
Definitiondf-ext 5907* Define the set of all extensional relationships over a base set. (Contributed by SF, 19-Feb-2015.)
Ext = {r, a x a y a (z a (zrxzry) → x = y)}
 
Definitiondf-sym 5908* Define the set of all symmetric relationships over a base set. (Contributed by SF, 19-Feb-2015.)
Sym = {r, a x a y a (xryyrx)}
 
Definitiondf-er 5909 Define the set of all equivalence relationships over a base set. (Contributed by SF, 19-Feb-2015.)
Er = ( SymTrans )
 
Theoremtransex 5910 The class of all transitive relationships is a set. (Contributed by SF, 19-Feb-2015.)
Trans V
 
Theoremrefex 5911 The class of all reflexive relationships is a set. (Contributed by SF, 11-Mar-2015.)
Ref V
 
Theoremantisymex 5912 The class of all antisymmetric relationships is a set. (Contributed by SF, 11-Mar-2015.)
Antisym V
 
Theoremconnexex 5913 The class of all connected relationships is a set. (Contributed by SF, 11-Mar-2015.)
Connex V
 
Theoremfoundex 5914 The class of all founded relationships is a set. (Contributed by SF, 19-Feb-2015.)
Fr V
 
Theoremextex 5915 The class of all extensional relationships is a set. (Contributed by SF, 19-Feb-2015.)
Ext V
 
Theoremsymex 5916 The class of all symmetric relationships is a set. (Contributed by SF, 20-Feb-2015.)
Sym V
 
Theorempartialex 5917 The class of all partial orderings is a set. (Contributed by SF, 11-Mar-2015.)
Po V
 
Theoremstrictex 5918 The class of all strict orderings is a set. (Contributed by SF, 19-Feb-2015.)
Or V
 
Theoremweex 5919 The class of all well orderings is a set. (Contributed by SF, 19-Feb-2015.)
We V
 
Theoremerex 5920 The class of all equivalence relationships is a set. (Contributed by SF, 20-Feb-2015.)
Er V
 
Theoremtrd 5921 Transitivity law in natural deduction form. (Contributed by SF, 20-Feb-2015.)
(φR Trans A)    &   (φX A)    &   (φY A)    &   (φZ A)    &   (φXRY)    &   (φYRZ)       (φXRZ)
 
Theoremfrd 5922* Founded relationship in natural deduction form. (Contributed by SF, 12-Mar-2015.)
(φR Fr A)    &   (φX V)    &   (φX A)    &   (φX)       (φy X z X (zRyz = y))
 
Theoremextd 5923* Extensional relationship in natural deduction form. (Contributed by SF, 20-Feb-2015.)
(φR Ext A)    &   (φX A)    &   (φY A)    &   ((φ z A) → (zRXzRY))       (φX = Y)
 
Theoremsymd 5924 Symmetric relationship in natural deduction form. (Contributed by SF, 20-Feb-2015.)
(φR Sym A)    &   (φX A)    &   (φY A)    &   (φXRY)       (φYRX)
 
Theoremtrrd 5925* Deduce transitivity from its properties. (Contributed by SF, 22-Feb-2015.)
(φR V)    &   (φA W)    &   ((φ (x A y A z A) (xRy yRz)) → xRz)       (φR Trans A)
 
Theoremrefrd 5926* Deduce reflexitiviy from its properties. (Contributed by SF, 12-Mar-2015.)
(φR V)    &   (φA W)    &   ((φ x A) → xRx)       (φR Ref A)
 
Theoremrefd 5927 Natural deduction form of reflexitivity. (Contributed by SF, 20-Mar-2015.)
(φR Ref A)    &   (φX A)       (φXRX)
 
Theoremantird 5928* Deduce antisymmetry from its properties. (Contributed by SF, 12-Mar-2015.)
(φR V)    &   (φA W)    &   ((φ (x A y A) (xRy yRx)) → x = y)       (φR Antisym A)
 
Theoremantid 5929 The antisymmetry property. (Contributed by SF, 18-Mar-2015.)
(φR Antisym A)    &   (φX A)    &   (φY A)    &   (φXRY)    &   (φYRX)       (φX = Y)
 
Theoremconnexrd 5930* Deduce connectivity from its properties. (Contributed by SF, 12-Mar-2015.)
(φR V)    &   (φA W)    &   ((φ x A y A) → (xRy yRx))       (φR Connex A)
 
Theoremconnexd 5931 The connectivity property. (Contributed by SF, 18-Mar-2015.)
(φR Connex A)    &   (φX A)    &   (φY A)       (φ → (XRY YRX))
 
Theoremersymtr 5932 Equivalence relationship as symmetric, transitive relationship. (Contributed by SF, 22-Feb-2015.)
(R Er A ↔ (R Sym A R Trans A))
 
Theoremporta 5933 Partial ordering as reflexive, transitive, antisymmetric relationship. (Contributed by SF, 12-Mar-2015.)
(R Po A ↔ (R Ref A R Trans A R Antisym A))
 
Theoremsopc 5934 Linear ordering as partial, connected relationship. (Contributed by SF, 12-Mar-2015.)
(R Or A ↔ (R Po A R Connex A))
 
Theoremfrds 5935* Substitution schema verson of frd 5922. (Contributed by SF, 19-Mar-2015.)
{x ψ} V    &   (x = y → (ψχ))    &   (x = z → (ψθ))    &   (φR Fr A)    &   (φx A ψ)       (φy A (χ z A ((θ zRy) → z = y)))
 
Theorempod 5936* A reflexive, transitive, and anti-symmetric ordering is a partial ordering. (Contributed by SF, 22-Feb-2015.)
(φR V)    &   (φA W)    &   ((φ x A) → xRx)    &   ((φ (x A y A z A) (xRy yRz)) → xRz)    &   ((φ (x A y A) (xRy yRx)) → x = y)       (φR Po A)
 
Theoremsod 5937* A reflexive, transitive, antisymmetric, and connected relationship is a strict ordering. (Contributed by SF, 12-Mar-2015.)
(φR V)    &   (φA W)    &   ((φ x A) → xRx)    &   ((φ (x A y A z A) (xRy yRz)) → xRz)    &   ((φ (x A y A) (xRy yRx)) → x = y)    &   ((φ x A y A) → (xRy yRx))       (φR Or A)
 
Theoremweds 5938* Any property that holds for some element of a well-ordered set A has an R minimal element satisfying that property. (Contributed by SF, 20-Mar-2015.)
{x ψ} V    &   (x = y → (ψχ))    &   (x = z → (ψθ))    &   (φR We A)    &   (φx A ψ)       (φy A (χ z A (θyRz)))
 
Theorempo0 5939 Anything partially orders the empty set. (Contributed by SF, 12-Mar-2015.)
(φR V)       (φR Po )
 
Theoremconnex0 5940 Anything is connected over the empty set. (Contributed by SF, 12-Mar-2015.)
(φR V)       (φR Connex )
 
Theoremso0 5941 Anything totally orders the empty set. (Contributed by SF, 12-Mar-2015.)
(φR V)       (φR Or )
 
Theoremiserd 5942* A symmetric, transitive relationship is an equivalence relationship. (Contributed by SF, 22-Feb-2015.)
(φR V)    &   (φA W)    &   ((φ (x A y A) xRy) → yRx)    &   ((φ (x A y A z A) (xRy yRz)) → xRz)       (φR Er A)
 
Theoremider 5943 The identity relationship is an equivalence relationship over the universe. (Contributed by SF, 22-Feb-2015.)
I Er V
 
Theoremssetpov 5944 The subset relationship partially orders the universe. (Contributed by SF, 12-Mar-2015.)
S Po V
 
2.4.2  Equivalence relations and classes
 
Syntaxcec 5945 Extend the definition of a class to include equivalence class.
class [A]R
 
Syntaxcqs 5946 Extend the definition of a class to include quotient set.
class (A / R)
 
Definitiondf-ec 5947 Define the R-coset of A. Exercise 35 of [Enderton] p. 61. This is called the equivalence class of A modulo R when R is an equivalence relation. In this case, A is a representative (member) of the equivalence class [A]R, which contains all sets that are equivalent to A. Definition of [Enderton] p. 57 uses the notation [A] (subscript) R, although we simply follow the brackets by R since we don't have subscripted expressions. For an alternate definition, see dfec2 5948. (Contributed by set.mm contributors, 22-Feb-2015.)
[A]R = (R “ {A})
 
Theoremdfec2 5948* Alternate definition of R-coset of A. Definition 34 of [Suppes] p. 81. (Contributed by set.mm contributors, 22-Feb-2015.)
[A]R = {y ARy}
 
Theoremecexg 5949 An equivalence class modulo a set is a set. (Contributed by set.mm contributors, 24-Jul-1995.)
(R B → [A]R V)
 
Theoremecexr 5950 A nonempty equivalence class implies the representative is a set. (Contributed by set.mm contributors, 9-Jul-2014.)
(A [B]RB V)
 
Definitiondf-qs 5951* Define quotient set. R is usually an equivalence relation. Definition of [Enderton] p. 58. (Contributed by set.mm contributors, 22-Feb-2015.)
(A / R) = {y x A y = [x]R}
 
Theoremersym 5952 An equivalence relation is symmetric. (Contributed by set.mm contributors, 22-Feb-2015.)
(φR Er A)    &   (φX A)    &   (φY A)    &   (φXRY)       (φYRX)
 
Theoremersymb 5953 An equivalence relation is symmetric. (Contributed by set.mm contributors, 30-Jul-1995.) (Revised by set.mm contributors, 9-Jul-2014.)
(φR Er A)    &   (φX A)    &   (φY A)       (φ → (XRYYRX))
 
Theoremertr 5954 An equivalence relation is transitive. (Contributed by set.mm contributors, 4-Jun-1995.) (Revised by set.mm contributors, 9-Jul-2014.)
(φR Er A)    &   (φX A)    &   (φY A)    &   (φZ A)       (φ → ((XRY YRZ) → XRZ))
 
Theoremertrd 5955 A transitivity relation for equivalences. (Contributed by set.mm contributors, 9-Jul-2014.)
(φR Er A)    &   (φX A)    &   (φY A)    &   (φZ A)    &   (φXRY)    &   (φYRZ)       (φXRZ)
 
Theoremertr2d 5956 A transitivity relation for equivalences. (Contributed by set.mm contributors, 9-Jul-2014.)
(φR Er A)    &   (φX A)    &   (φY A)    &   (φZ A)    &   (φXRY)    &   (φYRZ)       (φZRX)
 
Theoremertr3d 5957 A transitivity relation for equivalences. (Contributed by set.mm contributors, 9-Jul-2014.)
(φR Er A)    &   (φX A)    &   (φY A)    &   (φZ A)    &   (φYRX)    &   (φYRZ)       (φXRZ)
 
Theoremertr4d 5958 A transitivity relation for equivalences. (Contributed by set.mm contributors, 9-Jul-2014.)
(φR Er A)    &   (φX A)    &   (φY A)    &   (φZ A)    &   (φXRY)    &   (φZRY)       (φXRZ)
 
Theoremerref 5959 An equivalence relation is reflexive on its field. Compare Theorem 3M of [Enderton] p. 56. (Contributed by set.mm contributors, 6-May-2013.)
(φR Er V)    &   (φ → dom R = A)    &   (φX A)       (φXRX)
 
Theoremeqerlem 5960* Lemma for eqer 5961. (Contributed by set.mm contributors, 17-Mar-2008.)
(x = yA = B)    &   R = {x, y A = B}       (zRw[z / x]A = [w / x]A)
 
Theoremeqer 5961* Equivalence relation involving equality of dependent classes A(x) and B(y). (Contributed by set.mm contributors, 17-Mar-2008.)
(x = yA = B)    &   R = {x, y A = B}    &   R V       R Er V
 
Theoremeceq1 5962 Equality theorem for equivalence class. (Contributed by set.mm contributors, 23-Jul-1995.)
(A = B → [A]C = [B]C)
 
Theoremeceq2 5963 Equality theorem for equivalence class. (Contributed by set.mm contributors, 23-Jul-1995.)
(A = B → [C]A = [C]B)
 
Theoremelec 5964 Membership in an equivalence class. Theorem 72 of [Suppes] p. 82. (Contributed by set.mm contributors, 9-Jul-2014.)
(A [B]RBRA)
 
Theoremerdmrn 5965 The range and domain of an equivalence relation are equal. (Contributed by Rodolfo Medina, 11-Oct-2010.)
(R Er V → dom R = ran R)
 
Theoremecss 5966 An equivalence class is a subset of the domain. (Contributed by set.mm contributors, 6-Aug-1995.) (Revised by set.mm contributors, 9-Jul-2014.)
(φR Er V)    &   (φ → dom R = X)       (φ → [A]R X)
 
Theoremecdmn0 5967 A representative of a nonempty equivalence class belongs to the domain of the equivalence relation. (Contributed by set.mm contributors, 15-Feb-1996.) (Revised by set.mm contributors, 9-Jul-2014.)
(A dom R ↔ [A]R)
 
Theoremerth 5968 Basic property of equivalence relations. Theorem 73 of [Suppes] p. 82. (Contributed by set.mm contributors, 23-Jul-1995.) (Revised by Mario Carneiro, 9-Jul-2014.)
(φR Er V)    &   (φ → dom R = X)    &   (φA X)    &   (φB V)       (φ → (ARB ↔ [A]R = [B]R))
 
Theoremerth2 5969 Basic property of equivalence relations. Compare Theorem 73 of [Suppes] p. 82. Assumes membership of the second argument in the domain. (Contributed by set.mm contributors, 30-Jul-1995.) (Revised by set.mm contributors, 9-Jul-2014.)
(φR Er V)    &   (φ → dom R = X)    &   (φA V)    &   (φB X)       (φ → (ARB ↔ [A]R = [B]R))
 
Theoremerthi 5970 Basic property of equivalence relations. Part of Lemma 3N of [Enderton] p. 57. (Contributed by set.mm contributors, 30-Jul-1995.) (Revised by set.mm contributors, 9-Jul-2014.)
(φR Er V)    &   (φARB)       (φ → [A]R = [B]R)
 
Theoremereldm 5971 Equality of equivalence classes implies equivalence of domain membership. (Contributed by set.mm contributors, 28-Jan-1996.) (Revised by set.mm contributors, 9-Jul-2014.)
(φR Er V)    &   (φ → dom R = X)    &   (φ → [A]R = [B]R)    &   (φA V)    &   (φB W)       (φ → (A XB X))
 
Theoremerdisj 5972 Equivalence classes do not overlap. In other words, two equivalence classes are either equal or disjoint. Theorem 74 of [Suppes] p. 83. (Contributed by set.mm contributors, 15-Jun-2004.) (Revised by Mario Carneiro, 9-Jul-2014.)
(R Er V → ([A]R = [B]R ([A]R ∩ [B]R) = ))
 
Theoremecidsn 5973 An equivalence class modulo the identity relation is a singleton. (Contributed by set.mm contributors, 24-Oct-2004.)
[A] I = {A}
 
Theoremqseq1 5974 Equality theorem for quotient set. (Contributed by set.mm contributors, 23-Jul-1995.)
(A = B → (A / C) = (B / C))
 
Theoremqseq2 5975 Equality theorem for quotient set. (Contributed by set.mm contributors, 23-Jul-1995.)
(A = B → (C / A) = (C / B))
 
Theoremelqsg 5976* Closed form of elqs 5977. (Contributed by Rodolfo Medina, 12-Oct-2010.)
(B V → (B (A / R) ↔ x A B = [x]R))
 
Theoremelqs 5977* Membership in a quotient set. (Contributed by set.mm contributors, 23-Jul-1995.) (Revised by set.mm contributors, 12-Nov-2008.)
B V       (B (A / R) ↔ x A B = [x]R)
 
Theoremelqsi 5978* Membership in a quotient set. (Contributed by set.mm contributors, 23-Jul-1995.)
(B (A / R) → x A B = [x]R)
 
Theoremecelqsg 5979 Membership of an equivalence class in a quotient set. (Contributed by Jeff Madsen, 10-Jun-2010.)
((R V B A) → [B]R (A / R))
 
Theoremecelqsi 5980 Membership of an equivalence class in a quotient set. (Contributed by set.mm contributors, 25-Jul-1995.) (Revised by set.mm contributors, 9-Jul-2014.)
R V       (B A → [B]R (A / R))
 
Theoremecopqsi 5981 "Closure" law for equivalence class of ordered pairs. (Contributed by set.mm contributors, 25-Mar-1996.)
R V    &   S = ((A × A) / R)       ((B A C A) → [B, C]R S)
 
Theoremqsexg 5982 A quotient set exists. (Contributed by FL, 19-May-2007.)
((R V A W) → (A / R) V)
 
Theoremqsex 5983 A quotient set exists. (Contributed by set.mm contributors, 14-Aug-1995.)
R V    &   A V       (A / R) V
 
Theoremuniqs 5984 The union of a quotient set. (Contributed by set.mm contributors, 9-Dec-2008.)
(R V(A / R) = (RA))
 
Theoremuniqs2 5985 The union of a quotient set. (Contributed by set.mm contributors, 11-Jul-2014.)
(φR Er V)    &   (φ → dom R = A)    &   (φR V)       (φ(A / R) = A)
 
Theoremqsss 5986 A quotient set is a set of subsets of the base set. (Contributed by Mario Carneiro, 9-Jul-2014.)
(φR Er V)    &   (φ → dom R = A)    &   (φR V)       (φ → (A / R) A)
 
Theoremsnec 5987 The singleton of an equivalence class. (Contributed by set.mm contributors, 29-Jan-1999.) (Revised by set.mm contributors, 9-Jul-2014.)
A V       {[A]R} = ({A} / R)
 
Theoremecqs 5988 Equivalence class in terms of quotient set. (Contributed by set.mm contributors, 29-Jan-1999.) (Revised by set.mm contributors, 15-Jan-2009.)
R V       [A]R = ({A} / R)
 
Theoremecid 5989 A set is equal to its converse epsilon coset. (Note: converse epsilon is not an equivalence relation.) (Contributed by set.mm contributors, 13-Aug-1995.) (Revised by set.mm contributors, 9-Jul-2014.)
A V       [A] E = A
 
Theoremqsid 5990 A set is equal to its quotient set mod converse epsilon. (Note: converse epsilon is not an equivalence relation.) (Contributed by set.mm contributors, 13-Aug-1995.) (Revised by set.mm contributors, 9-Jul-2014.)
(A / E ) = A
 
Theoremectocld 5991* Implicit substitution of class for equivalence class. (Contributed by set.mm contributors, 9-Jul-2014.)
S = (B / R)    &   ([x]R = A → (φψ))    &   ((χ x B) → φ)       ((χ A S) → ψ)
 
Theoremectocl 5992* Implicit substitution of class for equivalence class. (Contributed by set.mm contributors, 23-Jul-1995.) (Revised by set.mm contributors, 9-Jul-2014.)
S = (B / R)    &   ([x]R = A → (φψ))    &   (x Bφ)       (A Sψ)
 
Theoremelqsn0 5993 A quotient set doesn't contain the empty set. (Contributed by set.mm contributors, 24-Aug-1995.) (Revised by set.mm contributors, 21-Mar-2007.)
((dom R = A B (A / R)) → B)
 
Theoremecelqsdm 5994 Membership of an equivalence class in a quotient set. (Contributed by set.mm contributors, 30-Jul-1995.) (Revised by set.mm contributors, 21-Mar-2007.)
((dom R = A [B]R (A / R)) → B A)
 
Theoremqsdisj 5995 Members of a quotient set do not overlap. (Contributed by Rodolfo Medina, 12-Oct-2010.) (Revised by Mario Carneiro, 11-Jul-2014.)
(φR Er V)    &   (φB (A / R))    &   (φC (A / R))       (φ → (B = C (BC) = ))
 
Theoremecoptocl 5996* Implicit substitution of class for equivalence class of ordered pair. (Contributed by set.mm contributors, 23-Jul-1995.)
S = ((B × C) / R)    &   ([x, y]R = A → (φψ))    &   ((x B y C) → φ)       (A Sψ)
 
Theorem2ecoptocl 5997* Implicit substitution of classes for equivalence classes of ordered pairs. (Contributed by set.mm contributors, 23-Jul-1995.)
S = ((C × D) / R)    &   ([x, y]R = A → (φψ))    &   ([z, w]R = B → (ψχ))    &   (((x C y D) (z C w D)) → φ)       ((A S B S) → χ)
 
Theorem3ecoptocl 5998* Implicit substitution of classes for equivalence classes of ordered pairs. (Contributed by set.mm contributors, 9-Aug-1995.)
S = ((D × D) / R)    &   ([x, y]R = A → (φψ))    &   ([z, w]R = B → (ψχ))    &   ([v, u]R = C → (χθ))    &   (((x D y D) (z D w D) (v D u D)) → φ)       ((A S B S C S) → θ)
 
2.4.3  The mapping operation
 
Syntaxcmap 5999 Extend the definition of a class to include the mapping operation. (Read for Am B, "the set of all functions that map from B to A.)
class m
 
Syntaxcpm 6000 Extend the definition of a class to include the partial mapping operation. (Read for Am B, "the set of all partial functions that map from B to A.)
class pm
    < 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 >