HomeHome Intuitionistic Logic Explorer
Theorem List (p. 46 of 108)
< Previous  Next >
Bad symbols? Try the
GIF version.

Mirrors  >  Metamath Home Page  >  ILE Home Page  >  Theorem List Contents  >  Recent Proofs       This page: Page List

Theorem List for Intuitionistic Logic Explorer - 4501-4600   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremrexxpf 4501* Version of rexxp 4498 with bound-variable hypotheses. (Contributed by NM, 19-Dec-2008.) (Revised by Mario Carneiro, 15-Oct-2016.)
𝑦𝜑    &   𝑧𝜑    &   𝑥𝜓    &   (𝑥 = ⟨𝑦, 𝑧⟩ → (𝜑𝜓))       (∃𝑥 ∈ (𝐴 × 𝐵)𝜑 ↔ ∃𝑦𝐴𝑧𝐵 𝜓)
 
Theoremiunxpf 4502* Indexed union on a cross product is equals a double indexed union. The hypothesis specifies an implicit substitution. (Contributed by NM, 19-Dec-2008.)
𝑦𝐶    &   𝑧𝐶    &   𝑥𝐷    &   (𝑥 = ⟨𝑦, 𝑧⟩ → 𝐶 = 𝐷)        𝑥 ∈ (𝐴 × 𝐵)𝐶 = 𝑦𝐴 𝑧𝐵 𝐷
 
Theoremopabbi2dv 4503* Deduce equality of a relation and an ordered-pair class builder. Compare abbi2dv 2197. (Contributed by NM, 24-Feb-2014.)
Rel 𝐴    &   (𝜑 → (⟨𝑥, 𝑦⟩ ∈ 𝐴𝜓))       (𝜑𝐴 = {⟨𝑥, 𝑦⟩ ∣ 𝜓})
 
Theoremrelop 4504* A necessary and sufficient condition for a Kuratowski ordered pair to be a relation. (Contributed by NM, 3-Jun-2008.) (Avoid depending on this detail.)
𝐴 ∈ V    &   𝐵 ∈ V       (Rel ⟨𝐴, 𝐵⟩ ↔ ∃𝑥𝑦(𝐴 = {𝑥} ∧ 𝐵 = {𝑥, 𝑦}))
 
Theoremideqg 4505 For sets, the identity relation is the same as equality. (Contributed by NM, 30-Apr-2004.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
(𝐵𝑉 → (𝐴 I 𝐵𝐴 = 𝐵))
 
Theoremideq 4506 For sets, the identity relation is the same as equality. (Contributed by NM, 13-Aug-1995.)
𝐵 ∈ V       (𝐴 I 𝐵𝐴 = 𝐵)
 
Theoremididg 4507 A set is identical to itself. (Contributed by NM, 28-May-2008.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
(𝐴𝑉𝐴 I 𝐴)
 
Theoremissetid 4508 Two ways of expressing set existence. (Contributed by NM, 16-Feb-2008.) (Proof shortened by Andrew Salmon, 27-Aug-2011.) (Revised by Mario Carneiro, 26-Apr-2015.)
(𝐴 ∈ V ↔ 𝐴 I 𝐴)
 
Theoremcoss1 4509 Subclass theorem for composition. (Contributed by FL, 30-Dec-2010.)
(𝐴𝐵 → (𝐴𝐶) ⊆ (𝐵𝐶))
 
Theoremcoss2 4510 Subclass theorem for composition. (Contributed by NM, 5-Apr-2013.)
(𝐴𝐵 → (𝐶𝐴) ⊆ (𝐶𝐵))
 
Theoremcoeq1 4511 Equality theorem for composition of two classes. (Contributed by NM, 3-Jan-1997.)
(𝐴 = 𝐵 → (𝐴𝐶) = (𝐵𝐶))
 
Theoremcoeq2 4512 Equality theorem for composition of two classes. (Contributed by NM, 3-Jan-1997.)
(𝐴 = 𝐵 → (𝐶𝐴) = (𝐶𝐵))
 
Theoremcoeq1i 4513 Equality inference for composition of two classes. (Contributed by NM, 16-Nov-2000.)
𝐴 = 𝐵       (𝐴𝐶) = (𝐵𝐶)
 
Theoremcoeq2i 4514 Equality inference for composition of two classes. (Contributed by NM, 16-Nov-2000.)
𝐴 = 𝐵       (𝐶𝐴) = (𝐶𝐵)
 
Theoremcoeq1d 4515 Equality deduction for composition of two classes. (Contributed by NM, 16-Nov-2000.)
(𝜑𝐴 = 𝐵)       (𝜑 → (𝐴𝐶) = (𝐵𝐶))
 
Theoremcoeq2d 4516 Equality deduction for composition of two classes. (Contributed by NM, 16-Nov-2000.)
(𝜑𝐴 = 𝐵)       (𝜑 → (𝐶𝐴) = (𝐶𝐵))
 
Theoremcoeq12i 4517 Equality inference for composition of two classes. (Contributed by FL, 7-Jun-2012.)
𝐴 = 𝐵    &   𝐶 = 𝐷       (𝐴𝐶) = (𝐵𝐷)
 
Theoremcoeq12d 4518 Equality deduction for composition of two classes. (Contributed by FL, 7-Jun-2012.)
(𝜑𝐴 = 𝐵)    &   (𝜑𝐶 = 𝐷)       (𝜑 → (𝐴𝐶) = (𝐵𝐷))
 
Theoremnfco 4519 Bound-variable hypothesis builder for function value. (Contributed by NM, 1-Sep-1999.)
𝑥𝐴    &   𝑥𝐵       𝑥(𝐴𝐵)
 
Theorembrcog 4520* Ordered pair membership in a composition. (Contributed by NM, 24-Feb-2015.)
((𝐴𝑉𝐵𝑊) → (𝐴(𝐶𝐷)𝐵 ↔ ∃𝑥(𝐴𝐷𝑥𝑥𝐶𝐵)))
 
Theoremopelco2g 4521* Ordered pair membership in a composition. (Contributed by NM, 27-Jan-1997.) (Revised by Mario Carneiro, 24-Feb-2015.)
((𝐴𝑉𝐵𝑊) → (⟨𝐴, 𝐵⟩ ∈ (𝐶𝐷) ↔ ∃𝑥(⟨𝐴, 𝑥⟩ ∈ 𝐷 ∧ ⟨𝑥, 𝐵⟩ ∈ 𝐶)))
 
Theorembrcogw 4522 Ordered pair membership in a composition. (Contributed by Thierry Arnoux, 14-Jan-2018.)
(((𝐴𝑉𝐵𝑊𝑋𝑍) ∧ (𝐴𝐷𝑋𝑋𝐶𝐵)) → 𝐴(𝐶𝐷)𝐵)
 
Theoremeqbrrdva 4523* Deduction from extensionality principle for relations, given an equivalence only on the relation's domain and range. (Contributed by Thierry Arnoux, 28-Nov-2017.)
(𝜑𝐴 ⊆ (𝐶 × 𝐷))    &   (𝜑𝐵 ⊆ (𝐶 × 𝐷))    &   ((𝜑𝑥𝐶𝑦𝐷) → (𝑥𝐴𝑦𝑥𝐵𝑦))       (𝜑𝐴 = 𝐵)
 
Theorembrco 4524* Binary relation on a composition. (Contributed by NM, 21-Sep-2004.) (Revised by Mario Carneiro, 24-Feb-2015.)
𝐴 ∈ V    &   𝐵 ∈ V       (𝐴(𝐶𝐷)𝐵 ↔ ∃𝑥(𝐴𝐷𝑥𝑥𝐶𝐵))
 
Theoremopelco 4525* Ordered pair membership in a composition. (Contributed by NM, 27-Dec-1996.) (Revised by Mario Carneiro, 24-Feb-2015.)
𝐴 ∈ V    &   𝐵 ∈ V       (⟨𝐴, 𝐵⟩ ∈ (𝐶𝐷) ↔ ∃𝑥(𝐴𝐷𝑥𝑥𝐶𝐵))
 
Theoremcnvss 4526 Subset theorem for converse. (Contributed by NM, 22-Mar-1998.)
(𝐴𝐵𝐴𝐵)
 
Theoremcnveq 4527 Equality theorem for converse. (Contributed by NM, 13-Aug-1995.)
(𝐴 = 𝐵𝐴 = 𝐵)
 
Theoremcnveqi 4528 Equality inference for converse. (Contributed by NM, 23-Dec-2008.)
𝐴 = 𝐵       𝐴 = 𝐵
 
Theoremcnveqd 4529 Equality deduction for converse. (Contributed by NM, 6-Dec-2013.)
(𝜑𝐴 = 𝐵)       (𝜑𝐴 = 𝐵)
 
Theoremelcnv 4530* Membership in a converse. Equation 5 of [Suppes] p. 62. (Contributed by NM, 24-Mar-1998.)
(𝐴𝑅 ↔ ∃𝑥𝑦(𝐴 = ⟨𝑥, 𝑦⟩ ∧ 𝑦𝑅𝑥))
 
Theoremelcnv2 4531* Membership in a converse. Equation 5 of [Suppes] p. 62. (Contributed by NM, 11-Aug-2004.)
(𝐴𝑅 ↔ ∃𝑥𝑦(𝐴 = ⟨𝑥, 𝑦⟩ ∧ ⟨𝑦, 𝑥⟩ ∈ 𝑅))
 
Theoremnfcnv 4532 Bound-variable hypothesis builder for converse. (Contributed by NM, 31-Jan-2004.) (Revised by Mario Carneiro, 15-Oct-2016.)
𝑥𝐴       𝑥𝐴
 
Theoremopelcnvg 4533 Ordered-pair membership in converse. (Contributed by NM, 13-May-1999.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
((𝐴𝐶𝐵𝐷) → (⟨𝐴, 𝐵⟩ ∈ 𝑅 ↔ ⟨𝐵, 𝐴⟩ ∈ 𝑅))
 
Theorembrcnvg 4534 The converse of a binary relation swaps arguments. Theorem 11 of [Suppes] p. 61. (Contributed by NM, 10-Oct-2005.)
((𝐴𝐶𝐵𝐷) → (𝐴𝑅𝐵𝐵𝑅𝐴))
 
Theoremopelcnv 4535 Ordered-pair membership in converse. (Contributed by NM, 13-Aug-1995.)
𝐴 ∈ V    &   𝐵 ∈ V       (⟨𝐴, 𝐵⟩ ∈ 𝑅 ↔ ⟨𝐵, 𝐴⟩ ∈ 𝑅)
 
Theorembrcnv 4536 The converse of a binary relation swaps arguments. Theorem 11 of [Suppes] p. 61. (Contributed by NM, 13-Aug-1995.)
𝐴 ∈ V    &   𝐵 ∈ V       (𝐴𝑅𝐵𝐵𝑅𝐴)
 
Theoremcsbcnvg 4537 Move class substitution in and out of the converse of a function. (Contributed by Thierry Arnoux, 8-Feb-2017.)
(𝐴𝑉𝐴 / 𝑥𝐹 = 𝐴 / 𝑥𝐹)
 
Theoremcnvco 4538 Distributive law of converse over class composition. Theorem 26 of [Suppes] p. 64. (Contributed by NM, 19-Mar-1998.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
(𝐴𝐵) = (𝐵𝐴)
 
Theoremcnvuni 4539* The converse of a class union is the (indexed) union of the converses of its members. (Contributed by NM, 11-Aug-2004.)
𝐴 = 𝑥𝐴 𝑥
 
Theoremdfdm3 4540* Alternate definition of domain. Definition 6.5(1) of [TakeutiZaring] p. 24. (Contributed by NM, 28-Dec-1996.)
dom 𝐴 = {𝑥 ∣ ∃𝑦𝑥, 𝑦⟩ ∈ 𝐴}
 
Theoremdfrn2 4541* Alternate definition of range. Definition 4 of [Suppes] p. 60. (Contributed by NM, 27-Dec-1996.)
ran 𝐴 = {𝑦 ∣ ∃𝑥 𝑥𝐴𝑦}
 
Theoremdfrn3 4542* Alternate definition of range. Definition 6.5(2) of [TakeutiZaring] p. 24. (Contributed by NM, 28-Dec-1996.)
ran 𝐴 = {𝑦 ∣ ∃𝑥𝑥, 𝑦⟩ ∈ 𝐴}
 
Theoremelrn2g 4543* Membership in a range. (Contributed by Scott Fenton, 2-Feb-2011.)
(𝐴𝑉 → (𝐴 ∈ ran 𝐵 ↔ ∃𝑥𝑥, 𝐴⟩ ∈ 𝐵))
 
Theoremelrng 4544* Membership in a range. (Contributed by Scott Fenton, 2-Feb-2011.)
(𝐴𝑉 → (𝐴 ∈ ran 𝐵 ↔ ∃𝑥 𝑥𝐵𝐴))
 
Theoremdfdm4 4545 Alternate definition of domain. (Contributed by NM, 28-Dec-1996.)
dom 𝐴 = ran 𝐴
 
Theoremdfdmf 4546* Definition of domain, using bound-variable hypotheses instead of distinct variable conditions. (Contributed by NM, 8-Mar-1995.) (Revised by Mario Carneiro, 15-Oct-2016.)
𝑥𝐴    &   𝑦𝐴       dom 𝐴 = {𝑥 ∣ ∃𝑦 𝑥𝐴𝑦}
 
Theoremcsbdmg 4547 Distribute proper substitution through the domain of a class. (Contributed by Jim Kingdon, 8-Dec-2018.)
(𝐴𝑉𝐴 / 𝑥dom 𝐵 = dom 𝐴 / 𝑥𝐵)
 
Theoremeldmg 4548* Domain membership. Theorem 4 of [Suppes] p. 59. (Contributed by Mario Carneiro, 9-Jul-2014.)
(𝐴𝑉 → (𝐴 ∈ dom 𝐵 ↔ ∃𝑦 𝐴𝐵𝑦))
 
Theoremeldm2g 4549* Domain membership. Theorem 4 of [Suppes] p. 59. (Contributed by NM, 27-Jan-1997.) (Revised by Mario Carneiro, 9-Jul-2014.)
(𝐴𝑉 → (𝐴 ∈ dom 𝐵 ↔ ∃𝑦𝐴, 𝑦⟩ ∈ 𝐵))
 
Theoremeldm 4550* Membership in a domain. Theorem 4 of [Suppes] p. 59. (Contributed by NM, 2-Apr-2004.)
𝐴 ∈ V       (𝐴 ∈ dom 𝐵 ↔ ∃𝑦 𝐴𝐵𝑦)
 
Theoremeldm2 4551* Membership in a domain. Theorem 4 of [Suppes] p. 59. (Contributed by NM, 1-Aug-1994.)
𝐴 ∈ V       (𝐴 ∈ dom 𝐵 ↔ ∃𝑦𝐴, 𝑦⟩ ∈ 𝐵)
 
Theoremdmss 4552 Subset theorem for domain. (Contributed by NM, 11-Aug-1994.)
(𝐴𝐵 → dom 𝐴 ⊆ dom 𝐵)
 
Theoremdmeq 4553 Equality theorem for domain. (Contributed by NM, 11-Aug-1994.)
(𝐴 = 𝐵 → dom 𝐴 = dom 𝐵)
 
Theoremdmeqi 4554 Equality inference for domain. (Contributed by NM, 4-Mar-2004.)
𝐴 = 𝐵       dom 𝐴 = dom 𝐵
 
Theoremdmeqd 4555 Equality deduction for domain. (Contributed by NM, 4-Mar-2004.)
(𝜑𝐴 = 𝐵)       (𝜑 → dom 𝐴 = dom 𝐵)
 
Theoremopeldm 4556 Membership of first of an ordered pair in a domain. (Contributed by NM, 30-Jul-1995.)
𝐴 ∈ V    &   𝐵 ∈ V       (⟨𝐴, 𝐵⟩ ∈ 𝐶𝐴 ∈ dom 𝐶)
 
Theorembreldm 4557 Membership of first of a binary relation in a domain. (Contributed by NM, 30-Jul-1995.)
𝐴 ∈ V    &   𝐵 ∈ V       (𝐴𝑅𝐵𝐴 ∈ dom 𝑅)
 
Theoremopeldmg 4558 Membership of first of an ordered pair in a domain. (Contributed by Jim Kingdon, 9-Jul-2019.)
((𝐴𝑉𝐵𝑊) → (⟨𝐴, 𝐵⟩ ∈ 𝐶𝐴 ∈ dom 𝐶))
 
Theorembreldmg 4559 Membership of first of a binary relation in a domain. (Contributed by NM, 21-Mar-2007.)
((𝐴𝐶𝐵𝐷𝐴𝑅𝐵) → 𝐴 ∈ dom 𝑅)
 
Theoremdmun 4560 The domain of a union is the union of domains. Exercise 56(a) of [Enderton] p. 65. (Contributed by NM, 12-Aug-1994.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
dom (𝐴𝐵) = (dom 𝐴 ∪ dom 𝐵)
 
Theoremdmin 4561 The domain of an intersection belong to the intersection of domains. Theorem 6 of [Suppes] p. 60. (Contributed by NM, 15-Sep-2004.)
dom (𝐴𝐵) ⊆ (dom 𝐴 ∩ dom 𝐵)
 
Theoremdmiun 4562 The domain of an indexed union. (Contributed by Mario Carneiro, 26-Apr-2016.)
dom 𝑥𝐴 𝐵 = 𝑥𝐴 dom 𝐵
 
Theoremdmuni 4563* The domain of a union. Part of Exercise 8 of [Enderton] p. 41. (Contributed by NM, 3-Feb-2004.)
dom 𝐴 = 𝑥𝐴 dom 𝑥
 
Theoremdmopab 4564* The domain of a class of ordered pairs. (Contributed by NM, 16-May-1995.) (Revised by Mario Carneiro, 4-Dec-2016.)
dom {⟨𝑥, 𝑦⟩ ∣ 𝜑} = {𝑥 ∣ ∃𝑦𝜑}
 
Theoremdmopabss 4565* Upper bound for the domain of a restricted class of ordered pairs. (Contributed by NM, 31-Jan-2004.)
dom {⟨𝑥, 𝑦⟩ ∣ (𝑥𝐴𝜑)} ⊆ 𝐴
 
Theoremdmopab3 4566* The domain of a restricted class of ordered pairs. (Contributed by NM, 31-Jan-2004.)
(∀𝑥𝐴𝑦𝜑 ↔ dom {⟨𝑥, 𝑦⟩ ∣ (𝑥𝐴𝜑)} = 𝐴)
 
Theoremdm0 4567 The domain of the empty set is empty. Part of Theorem 3.8(v) of [Monk1] p. 36. (Contributed by NM, 4-Jul-1994.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
dom ∅ = ∅
 
Theoremdmi 4568 The domain of the identity relation is the universe. (Contributed by NM, 30-Apr-1998.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
dom I = V
 
Theoremdmv 4569 The domain of the universe is the universe. (Contributed by NM, 8-Aug-2003.)
dom V = V
 
Theoremdm0rn0 4570 An empty domain implies an empty range. (Contributed by NM, 21-May-1998.)
(dom 𝐴 = ∅ ↔ ran 𝐴 = ∅)
 
Theoremreldm0 4571 A relation is empty iff its domain is empty. (Contributed by NM, 15-Sep-2004.)
(Rel 𝐴 → (𝐴 = ∅ ↔ dom 𝐴 = ∅))
 
Theoremdmmrnm 4572* A domain is inhabited if and only if the range is inhabited. (Contributed by Jim Kingdon, 15-Dec-2018.)
(∃𝑥 𝑥 ∈ dom 𝐴 ↔ ∃𝑦 𝑦 ∈ ran 𝐴)
 
Theoremdmxpm 4573* The domain of a cross product. Part of Theorem 3.13(x) of [Monk1] p. 37. (Contributed by NM, 28-Jul-1995.) (Proof shortened by Andrew Salmon, 27-Aug-2011.)
(∃𝑥 𝑥𝐵 → dom (𝐴 × 𝐵) = 𝐴)
 
Theoremdmxpinm 4574* The domain of the intersection of two square cross products. Unlike dmin 4561, equality holds. (Contributed by NM, 29-Jan-2008.)
(∃𝑥 𝑥 ∈ (𝐴𝐵) → dom ((𝐴 × 𝐴) ∩ (𝐵 × 𝐵)) = (𝐴𝐵))
 
Theoremxpid11m 4575* The cross product of a class with itself is one-to-one. (Contributed by Jim Kingdon, 8-Dec-2018.)
((∃𝑥 𝑥𝐴 ∧ ∃𝑥 𝑥𝐵) → ((𝐴 × 𝐴) = (𝐵 × 𝐵) ↔ 𝐴 = 𝐵))
 
Theoremdmcnvcnv 4576 The domain of the double converse of a class (which doesn't have to be a relation as in dfrel2 4791). (Contributed by NM, 8-Apr-2007.)
dom 𝐴 = dom 𝐴
 
Theoremrncnvcnv 4577 The range of the double converse of a class. (Contributed by NM, 8-Apr-2007.)
ran 𝐴 = ran 𝐴
 
Theoremelreldm 4578 The first member of an ordered pair in a relation belongs to the domain of the relation. (Contributed by NM, 28-Jul-2004.)
((Rel 𝐴𝐵𝐴) → 𝐵 ∈ dom 𝐴)
 
Theoremrneq 4579 Equality theorem for range. (Contributed by NM, 29-Dec-1996.)
(𝐴 = 𝐵 → ran 𝐴 = ran 𝐵)
 
Theoremrneqi 4580 Equality inference for range. (Contributed by NM, 4-Mar-2004.)
𝐴 = 𝐵       ran 𝐴 = ran 𝐵
 
Theoremrneqd 4581 Equality deduction for range. (Contributed by NM, 4-Mar-2004.)
(𝜑𝐴 = 𝐵)       (𝜑 → ran 𝐴 = ran 𝐵)
 
Theoremrnss 4582 Subset theorem for range. (Contributed by NM, 22-Mar-1998.)
(𝐴𝐵 → ran 𝐴 ⊆ ran 𝐵)
 
Theorembrelrng 4583 The second argument of a binary relation belongs to its range. (Contributed by NM, 29-Jun-2008.)
((𝐴𝐹𝐵𝐺𝐴𝐶𝐵) → 𝐵 ∈ ran 𝐶)
 
Theoremopelrng 4584 Membership of second member of an ordered pair in a range. (Contributed by Jim Kingdon, 26-Jan-2019.)
((𝐴𝐹𝐵𝐺 ∧ ⟨𝐴, 𝐵⟩ ∈ 𝐶) → 𝐵 ∈ ran 𝐶)
 
Theorembrelrn 4585 The second argument of a binary relation belongs to its range. (Contributed by NM, 13-Aug-2004.)
𝐴 ∈ V    &   𝐵 ∈ V       (𝐴𝐶𝐵𝐵 ∈ ran 𝐶)
 
Theoremopelrn 4586 Membership of second member of an ordered pair in a range. (Contributed by NM, 23-Feb-1997.)
𝐴 ∈ V    &   𝐵 ∈ V       (⟨𝐴, 𝐵⟩ ∈ 𝐶𝐵 ∈ ran 𝐶)
 
Theoremreleldm 4587 The first argument of a binary relation belongs to its domain. (Contributed by NM, 2-Jul-2008.)
((Rel 𝑅𝐴𝑅𝐵) → 𝐴 ∈ dom 𝑅)
 
Theoremrelelrn 4588 The second argument of a binary relation belongs to its range. (Contributed by NM, 2-Jul-2008.)
((Rel 𝑅𝐴𝑅𝐵) → 𝐵 ∈ ran 𝑅)
 
Theoremreleldmb 4589* Membership in a domain. (Contributed by Mario Carneiro, 5-Nov-2015.)
(Rel 𝑅 → (𝐴 ∈ dom 𝑅 ↔ ∃𝑥 𝐴𝑅𝑥))
 
Theoremrelelrnb 4590* Membership in a range. (Contributed by Mario Carneiro, 5-Nov-2015.)
(Rel 𝑅 → (𝐴 ∈ ran 𝑅 ↔ ∃𝑥 𝑥𝑅𝐴))
 
Theoremreleldmi 4591 The first argument of a binary relation belongs to its domain. (Contributed by NM, 28-Apr-2015.)
Rel 𝑅       (𝐴𝑅𝐵𝐴 ∈ dom 𝑅)
 
Theoremrelelrni 4592 The second argument of a binary relation belongs to its range. (Contributed by NM, 28-Apr-2015.)
Rel 𝑅       (𝐴𝑅𝐵𝐵 ∈ ran 𝑅)
 
Theoremdfrnf 4593* Definition of range, using bound-variable hypotheses instead of distinct variable conditions. (Contributed by NM, 14-Aug-1995.) (Revised by Mario Carneiro, 15-Oct-2016.)
𝑥𝐴    &   𝑦𝐴       ran 𝐴 = {𝑦 ∣ ∃𝑥 𝑥𝐴𝑦}
 
Theoremelrn2 4594* Membership in a range. (Contributed by NM, 10-Jul-1994.)
𝐴 ∈ V       (𝐴 ∈ ran 𝐵 ↔ ∃𝑥𝑥, 𝐴⟩ ∈ 𝐵)
 
Theoremelrn 4595* Membership in a range. (Contributed by NM, 2-Apr-2004.)
𝐴 ∈ V       (𝐴 ∈ ran 𝐵 ↔ ∃𝑥 𝑥𝐵𝐴)
 
Theoremnfdm 4596 Bound-variable hypothesis builder for domain. (Contributed by NM, 30-Jan-2004.) (Revised by Mario Carneiro, 15-Oct-2016.)
𝑥𝐴       𝑥dom 𝐴
 
Theoremnfrn 4597 Bound-variable hypothesis builder for range. (Contributed by NM, 1-Sep-1999.) (Revised by Mario Carneiro, 15-Oct-2016.)
𝑥𝐴       𝑥ran 𝐴
 
Theoremdmiin 4598 Domain of an intersection. (Contributed by FL, 15-Oct-2012.)
dom 𝑥𝐴 𝐵 𝑥𝐴 dom 𝐵
 
Theoremrnopab 4599* The range of a class of ordered pairs. (Contributed by NM, 14-Aug-1995.) (Revised by Mario Carneiro, 4-Dec-2016.)
ran {⟨𝑥, 𝑦⟩ ∣ 𝜑} = {𝑦 ∣ ∃𝑥𝜑}
 
Theoremrnmpt 4600* The range of a function in maps-to notation. (Contributed by Scott Fenton, 21-Mar-2011.) (Revised by Mario Carneiro, 31-Aug-2015.)
𝐹 = (𝑥𝐴𝐵)       ran 𝐹 = {𝑦 ∣ ∃𝑥𝐴 𝑦 = 𝐵}
    < 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-6400 65 6401-6500 66 6501-6600 67 6601-6700 68 6701-6800 69 6801-6900 70 6901-7000 71 7001-7100 72 7101-7200 73 7201-7300 74 7301-7400 75 7401-7500 76 7501-7600 77 7601-7700 78 7701-7800 79 7801-7900 80 7901-8000 81 8001-8100 82 8101-8200 83 8201-8300 84 8301-8400 85 8401-8500 86 8501-8600 87 8601-8700 88 8701-8800 89 8801-8900 90 8901-9000 91 9001-9100 92 9101-9200 93 9201-9300 94 9301-9400 95 9401-9500 96 9501-9600 97 9601-9700 98 9701-9800 99 9801-9900 100 9901-10000 101 10001-10100 102 10101-10200 103 10201-10300 104 10301-10400 105 10401-10500 106 10501-10600 107 10601-10700 108 10701-10795
  Copyright terms: Public domain < Previous  Next >