HomeHome Intuitionistic Logic Explorer
Theorem List (p. 18 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 - 1701-1800   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremsbf2 1701 Substitution has no effect on a bound variable. (Contributed by NM, 1-Jul-2005.)
([𝑦 / 𝑥]∀𝑥𝜑 ↔ ∀𝑥𝜑)
 
Theoremsb6x 1702 Equivalence involving substitution for a variable not free. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 12-Aug-2011.)
(𝜑 → ∀𝑥𝜑)       ([𝑦 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝑦𝜑))
 
Theoremnfs1f 1703 If 𝑥 is not free in 𝜑, it is not free in [𝑦 / 𝑥]𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.)
𝑥𝜑       𝑥[𝑦 / 𝑥]𝜑
 
Theoremhbs1f 1704 If 𝑥 is not free in 𝜑, it is not free in [𝑦 / 𝑥]𝜑. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.)
(𝜑 → ∀𝑥𝜑)       ([𝑦 / 𝑥]𝜑 → ∀𝑥[𝑦 / 𝑥]𝜑)
 
Theoremsbequ5 1705 Substitution does not change an identical variable specifier. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 21-Dec-2004.)
([𝑤 / 𝑧]∀𝑥 𝑥 = 𝑦 ↔ ∀𝑥 𝑥 = 𝑦)
 
Theoremsbequ6 1706 Substitution does not change a distinctor. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 14-May-2005.)
([𝑤 / 𝑧] ¬ ∀𝑥 𝑥 = 𝑦 ↔ ¬ ∀𝑥 𝑥 = 𝑦)
 
Theoremsbt 1707 A substitution into a theorem remains true. (See chvar 1680 and chvarv 1853 for versions using implicit substitition.) (Contributed by NM, 21-Jan-2004.) (Proof shortened by Andrew Salmon, 25-May-2011.)
𝜑       [𝑦 / 𝑥]𝜑
 
Theoremequsb1 1708 Substitution applied to an atomic wff. (Contributed by NM, 5-Aug-1993.)
[𝑦 / 𝑥]𝑥 = 𝑦
 
Theoremequsb2 1709 Substitution applied to an atomic wff. (Contributed by NM, 5-Aug-1993.)
[𝑦 / 𝑥]𝑦 = 𝑥
 
Theoremsbiedh 1710 Conversion of implicit substitution to explicit substitution (deduction version of sbieh 1713). New proofs should use sbied 1711 instead. (Contributed by NM, 30-Jun-1994.) (Proof shortened by Andrew Salmon, 25-May-2011.) (New usage is discouraged.)
(𝜑 → ∀𝑥𝜑)    &   (𝜑 → (𝜒 → ∀𝑥𝜒))    &   (𝜑 → (𝑥 = 𝑦 → (𝜓𝜒)))       (𝜑 → ([𝑦 / 𝑥]𝜓𝜒))
 
Theoremsbied 1711 Conversion of implicit substitution to explicit substitution (deduction version of sbie 1714). (Contributed by NM, 30-Jun-1994.) (Revised by Mario Carneiro, 4-Oct-2016.)
𝑥𝜑    &   (𝜑 → Ⅎ𝑥𝜒)    &   (𝜑 → (𝑥 = 𝑦 → (𝜓𝜒)))       (𝜑 → ([𝑦 / 𝑥]𝜓𝜒))
 
Theoremsbiedv 1712* Conversion of implicit substitution to explicit substitution (deduction version of sbie 1714). (Contributed by NM, 7-Jan-2017.)
((𝜑𝑥 = 𝑦) → (𝜓𝜒))       (𝜑 → ([𝑦 / 𝑥]𝜓𝜒))
 
Theoremsbieh 1713 Conversion of implicit substitution to explicit substitution. New proofs should use sbie 1714 instead. (Contributed by NM, 30-Jun-1994.) (New usage is discouraged.)
(𝜓 → ∀𝑥𝜓)    &   (𝑥 = 𝑦 → (𝜑𝜓))       ([𝑦 / 𝑥]𝜑𝜓)
 
Theoremsbie 1714 Conversion of implicit substitution to explicit substitution. (Contributed by NM, 30-Jun-1994.) (Revised by Mario Carneiro, 4-Oct-2016.) (Revised by Wolf Lammen, 30-Apr-2018.)
𝑥𝜓    &   (𝑥 = 𝑦 → (𝜑𝜓))       ([𝑦 / 𝑥]𝜑𝜓)
 
1.3.11  Theorems using axiom ax-11
 
Theoremequs5a 1715 A property related to substitution that unlike equs5 1750 doesn't require a distinctor antecedent. (Contributed by NM, 2-Feb-2007.)
(∃𝑥(𝑥 = 𝑦 ∧ ∀𝑦𝜑) → ∀𝑥(𝑥 = 𝑦𝜑))
 
Theoremequs5e 1716 A property related to substitution that unlike equs5 1750 doesn't require a distinctor antecedent. (Contributed by NM, 2-Feb-2007.) (Revised by NM, 3-Feb-2015.)
(∃𝑥(𝑥 = 𝑦𝜑) → ∀𝑥(𝑥 = 𝑦 → ∃𝑦𝜑))
 
Theoremax11e 1717 Analogue to ax-11 1437 but for existential quantification. (Contributed by Mario Carneiro and Jim Kingdon, 31-Dec-2017.) (Proved by Mario Carneiro, 9-Feb-2018.)
(𝑥 = 𝑦 → (∃𝑥(𝑥 = 𝑦𝜑) → ∃𝑦𝜑))
 
Theoremax10oe 1718 Quantifier Substitution for existential quantifiers. Analogue to ax10o 1643 but for rather than . (Contributed by Jim Kingdon, 21-Dec-2017.)
(∀𝑥 𝑥 = 𝑦 → (∃𝑥𝜓 → ∃𝑦𝜓))
 
Theoremdrex1 1719 Formula-building lemma for use with the Distinctor Reduction Theorem. Part of Theorem 9.4 of [Megill] p. 448 (p. 16 of preprint). (Contributed by NM, 27-Feb-2005.) (Revised by NM, 3-Feb-2015.)
(∀𝑥 𝑥 = 𝑦 → (𝜑𝜓))       (∀𝑥 𝑥 = 𝑦 → (∃𝑥𝜑 ↔ ∃𝑦𝜓))
 
Theoremdrsb1 1720 Formula-building lemma for use with the Distinctor Reduction Theorem. Part of Theorem 9.4 of [Megill] p. 448 (p. 16 of preprint). (Contributed by NM, 5-Aug-1993.)
(∀𝑥 𝑥 = 𝑦 → ([𝑧 / 𝑥]𝜑 ↔ [𝑧 / 𝑦]𝜑))
 
Theoremexdistrfor 1721 Distribution of existential quantifiers, with a bound-variable hypothesis saying that 𝑦 is not free in 𝜑, but 𝑥 can be free in 𝜑 (and there is no distinct variable condition on 𝑥 and 𝑦). (Contributed by Jim Kingdon, 25-Feb-2018.)
(∀𝑥 𝑥 = 𝑦 ∨ ∀𝑥𝑦𝜑)       (∃𝑥𝑦(𝜑𝜓) → ∃𝑥(𝜑 ∧ ∃𝑦𝜓))
 
Theoremsb4a 1722 A version of sb4 1753 that doesn't require a distinctor antecedent. (Contributed by NM, 2-Feb-2007.)
([𝑦 / 𝑥]∀𝑦𝜑 → ∀𝑥(𝑥 = 𝑦𝜑))
 
Theoremequs45f 1723 Two ways of expressing substitution when 𝑦 is not free in 𝜑. (Contributed by NM, 25-Apr-2008.)
(𝜑 → ∀𝑦𝜑)       (∃𝑥(𝑥 = 𝑦𝜑) ↔ ∀𝑥(𝑥 = 𝑦𝜑))
 
Theoremsb6f 1724 Equivalence for substitution when 𝑦 is not free in 𝜑. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 30-Apr-2008.)
(𝜑 → ∀𝑦𝜑)       ([𝑦 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝑦𝜑))
 
Theoremsb5f 1725 Equivalence for substitution when 𝑦 is not free in 𝜑. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 18-May-2008.)
(𝜑 → ∀𝑦𝜑)       ([𝑦 / 𝑥]𝜑 ↔ ∃𝑥(𝑥 = 𝑦𝜑))
 
Theoremsb4e 1726 One direction of a simplified definition of substitution that unlike sb4 1753 doesn't require a distinctor antecedent. (Contributed by NM, 2-Feb-2007.)
([𝑦 / 𝑥]𝜑 → ∀𝑥(𝑥 = 𝑦 → ∃𝑦𝜑))
 
Theoremhbsb2a 1727 Special case of a bound-variable hypothesis builder for substitution. (Contributed by NM, 2-Feb-2007.)
([𝑦 / 𝑥]∀𝑦𝜑 → ∀𝑥[𝑦 / 𝑥]𝜑)
 
Theoremhbsb2e 1728 Special case of a bound-variable hypothesis builder for substitution. (Contributed by NM, 2-Feb-2007.)
([𝑦 / 𝑥]𝜑 → ∀𝑥[𝑦 / 𝑥]∃𝑦𝜑)
 
Theoremhbsb3 1729 If 𝑦 is not free in 𝜑, 𝑥 is not free in [𝑦 / 𝑥]𝜑. (Contributed by NM, 5-Aug-1993.)
(𝜑 → ∀𝑦𝜑)       ([𝑦 / 𝑥]𝜑 → ∀𝑥[𝑦 / 𝑥]𝜑)
 
Theoremnfs1 1730 If 𝑦 is not free in 𝜑, 𝑥 is not free in [𝑦 / 𝑥]𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.)
𝑦𝜑       𝑥[𝑦 / 𝑥]𝜑
 
Theoremsbcof2 1731 Version of sbco 1883 where 𝑥 is not free in 𝜑. (Contributed by Jim Kingdon, 28-Dec-2017.)
(𝜑 → ∀𝑥𝜑)       ([𝑦 / 𝑥][𝑥 / 𝑦]𝜑 ↔ [𝑦 / 𝑥]𝜑)
 
1.4  Predicate calculus with distinct variables
 
1.4.1  Derive the axiom of distinct variables ax-16
 
Theoremspimv 1732* A version of spim 1666 with a distinct variable requirement instead of a bound variable hypothesis. (Contributed by NM, 5-Aug-1993.)
(𝑥 = 𝑦 → (𝜑𝜓))       (∀𝑥𝜑𝜓)
 
Theoremaev 1733* A "distinctor elimination" lemma with no restrictions on variables in the consequent, proved without using ax-16 1735. (Contributed by NM, 8-Nov-2006.) (Proof shortened by Andrew Salmon, 21-Jun-2011.)
(∀𝑥 𝑥 = 𝑦 → ∀𝑧 𝑤 = 𝑣)
 
Theoremax16 1734* Theorem showing that ax-16 1735 is redundant if ax-17 1459 is included in the axiom system. The important part of the proof is provided by aev 1733.

See ax16ALT 1780 for an alternate proof that does not require ax-10 1436 or ax-12 1442.

This theorem should not be referenced in any proof. Instead, use ax-16 1735 below so that theorems needing ax-16 1735 can be more easily identified. (Contributed by NM, 8-Nov-2006.)

(∀𝑥 𝑥 = 𝑦 → (𝜑 → ∀𝑥𝜑))
 
Axiomax-16 1735* Axiom of Distinct Variables. The only axiom of predicate calculus requiring that variables be distinct (if we consider ax-17 1459 to be a metatheorem and not an axiom). Axiom scheme C16' in [Megill] p. 448 (p. 16 of the preprint). It apparently does not otherwise appear in the literature but is easily proved from textbook predicate calculus by cases. It is a somewhat bizarre axiom since the antecedent is always false in set theory, but nonetheless it is technically necessary as you can see from its uses.

This axiom is redundant if we include ax-17 1459; see theorem ax16 1734.

This axiom is obsolete and should no longer be used. It is proved above as theorem ax16 1734. (Contributed by NM, 5-Aug-1993.) (New usage is discouraged.)

(∀𝑥 𝑥 = 𝑦 → (𝜑 → ∀𝑥𝜑))
 
Theoremdveeq2 1736* Quantifier introduction when one pair of variables is distinct. (Contributed by NM, 2-Jan-2002.)
(¬ ∀𝑥 𝑥 = 𝑦 → (𝑧 = 𝑦 → ∀𝑥 𝑧 = 𝑦))
 
Theoremdveeq2or 1737* Quantifier introduction when one pair of variables is distinct. Like dveeq2 1736 but connecting 𝑥𝑥 = 𝑦 by a disjunction rather than negation and implication makes the theorem stronger in intuitionistic logic. (Contributed by Jim Kingdon, 1-Feb-2018.)
(∀𝑥 𝑥 = 𝑦 ∨ Ⅎ𝑥 𝑧 = 𝑦)
 
TheoremdvelimfALT2 1738* Proof of dvelimf 1932 using dveeq2 1736 (shown as the last hypothesis) instead of ax-12 1442. This shows that ax-12 1442 could be replaced by dveeq2 1736 (the last hypothesis). (Contributed by Andrew Salmon, 21-Jul-2011.)
(𝜑 → ∀𝑥𝜑)    &   (𝜓 → ∀𝑧𝜓)    &   (𝑧 = 𝑦 → (𝜑𝜓))    &   (¬ ∀𝑥 𝑥 = 𝑦 → (𝑧 = 𝑦 → ∀𝑥 𝑧 = 𝑦))       (¬ ∀𝑥 𝑥 = 𝑦 → (𝜓 → ∀𝑥𝜓))
 
Theoremnd5 1739* A lemma for proving conditionless ZFC axioms. (Contributed by NM, 8-Jan-2002.)
(¬ ∀𝑦 𝑦 = 𝑥 → (𝑧 = 𝑦 → ∀𝑥 𝑧 = 𝑦))
 
Theoremexlimdv 1740* Deduction from Theorem 19.23 of [Margaris] p. 90. (Contributed by NM, 27-Apr-1994.)
(𝜑 → (𝜓𝜒))       (𝜑 → (∃𝑥𝜓𝜒))
 
Theoremax11v2 1741* Recovery of ax11o 1743 from ax11v 1748 without using ax-11 1437. The hypothesis is even weaker than ax11v 1748, with 𝑧 both distinct from 𝑥 and not occurring in 𝜑. Thus the hypothesis provides an alternate axiom that can be used in place of ax11o 1743. (Contributed by NM, 2-Feb-2007.)
(𝑥 = 𝑧 → (𝜑 → ∀𝑥(𝑥 = 𝑧𝜑)))       (¬ ∀𝑥 𝑥 = 𝑦 → (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦𝜑))))
 
Theoremax11a2 1742* Derive ax-11o 1744 from a hypothesis in the form of ax-11 1437. The hypothesis is even weaker than ax-11 1437, with 𝑧 both distinct from 𝑥 and not occurring in 𝜑. Thus the hypothesis provides an alternate axiom that can be used in place of ax11o 1743. (Contributed by NM, 2-Feb-2007.)
(𝑥 = 𝑧 → (∀𝑧𝜑 → ∀𝑥(𝑥 = 𝑧𝜑)))       (¬ ∀𝑥 𝑥 = 𝑦 → (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦𝜑))))
 
1.4.2  Derive the obsolete axiom of variable substitution ax-11o
 
Theoremax11o 1743 Derivation of set.mm's original ax-11o 1744 from the shorter ax-11 1437 that has replaced it.

An open problem is whether this theorem can be proved without relying on ax-16 1735 or ax-17 1459.

Normally, ax11o 1743 should be used rather than ax-11o 1744, except by theorems specifically studying the latter's properties. (Contributed by NM, 3-Feb-2007.)

(¬ ∀𝑥 𝑥 = 𝑦 → (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦𝜑))))
 
Axiomax-11o 1744 Axiom ax-11o 1744 ("o" for "old") was the original version of ax-11 1437, before it was discovered (in Jan. 2007) that the shorter ax-11 1437 could replace it. It appears as Axiom scheme C15' in [Megill] p. 448 (p. 16 of the preprint). It is based on Lemma 16 of [Tarski] p. 70 and Axiom C8 of [Monk2] p. 105, from which it can be proved by cases. To understand this theorem more easily, think of "¬ ∀𝑥𝑥 = 𝑦..." as informally meaning "if 𝑥 and 𝑦 are distinct variables then..." The antecedent becomes false if the same variable is substituted for 𝑥 and 𝑦, ensuring the theorem is sound whenever this is the case. In some later theorems, we call an antecedent of the form ¬ ∀𝑥𝑥 = 𝑦 a "distinctor."

This axiom is redundant, as shown by theorem ax11o 1743.

This axiom is obsolete and should no longer be used. It is proved above as theorem ax11o 1743. (Contributed by NM, 5-Aug-1993.) (New usage is discouraged.)

(¬ ∀𝑥 𝑥 = 𝑦 → (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦𝜑))))
 
1.4.3  More theorems related to ax-11 and substitution
 
Theoremalbidv 1745* Formula-building rule for universal quantifier (deduction rule). (Contributed by NM, 5-Aug-1993.)
(𝜑 → (𝜓𝜒))       (𝜑 → (∀𝑥𝜓 ↔ ∀𝑥𝜒))
 
Theoremexbidv 1746* Formula-building rule for existential quantifier (deduction rule). (Contributed by NM, 5-Aug-1993.)
(𝜑 → (𝜓𝜒))       (𝜑 → (∃𝑥𝜓 ↔ ∃𝑥𝜒))
 
Theoremax11b 1747 A bidirectional version of ax-11o 1744. (Contributed by NM, 30-Jun-2006.)
((¬ ∀𝑥 𝑥 = 𝑦𝑥 = 𝑦) → (𝜑 ↔ ∀𝑥(𝑥 = 𝑦𝜑)))
 
Theoremax11v 1748* This is a version of ax-11o 1744 when the variables are distinct. Axiom (C8) of [Monk2] p. 105. (Contributed by NM, 5-Aug-1993.) (Revised by Jim Kingdon, 15-Dec-2017.)
(𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦𝜑)))
 
Theoremax11ev 1749* Analogue to ax11v 1748 for existential quantification. (Contributed by Jim Kingdon, 9-Jan-2018.)
(𝑥 = 𝑦 → (∃𝑥(𝑥 = 𝑦𝜑) → 𝜑))
 
Theoremequs5 1750 Lemma used in proofs of substitution properties. (Contributed by NM, 5-Aug-1993.)
(¬ ∀𝑥 𝑥 = 𝑦 → (∃𝑥(𝑥 = 𝑦𝜑) → ∀𝑥(𝑥 = 𝑦𝜑)))
 
Theoremequs5or 1751 Lemma used in proofs of substitution properties. Like equs5 1750 but, in intuitionistic logic, replacing negation and implication with disjunction makes this a stronger result. (Contributed by Jim Kingdon, 2-Feb-2018.)
(∀𝑥 𝑥 = 𝑦 ∨ (∃𝑥(𝑥 = 𝑦𝜑) → ∀𝑥(𝑥 = 𝑦𝜑)))
 
Theoremsb3 1752 One direction of a simplified definition of substitution when variables are distinct. (Contributed by NM, 5-Aug-1993.)
(¬ ∀𝑥 𝑥 = 𝑦 → (∃𝑥(𝑥 = 𝑦𝜑) → [𝑦 / 𝑥]𝜑))
 
Theoremsb4 1753 One direction of a simplified definition of substitution when variables are distinct. (Contributed by NM, 5-Aug-1993.)
(¬ ∀𝑥 𝑥 = 𝑦 → ([𝑦 / 𝑥]𝜑 → ∀𝑥(𝑥 = 𝑦𝜑)))
 
Theoremsb4or 1754 One direction of a simplified definition of substitution when variables are distinct. Similar to sb4 1753 but stronger in intuitionistic logic. (Contributed by Jim Kingdon, 2-Feb-2018.)
(∀𝑥 𝑥 = 𝑦 ∨ ∀𝑥([𝑦 / 𝑥]𝜑 → ∀𝑥(𝑥 = 𝑦𝜑)))
 
Theoremsb4b 1755 Simplified definition of substitution when variables are distinct. (Contributed by NM, 27-May-1997.)
(¬ ∀𝑥 𝑥 = 𝑦 → ([𝑦 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝑦𝜑)))
 
Theoremsb4bor 1756 Simplified definition of substitution when variables are distinct, expressed via disjunction. (Contributed by Jim Kingdon, 18-Mar-2018.)
(∀𝑥 𝑥 = 𝑦 ∨ ∀𝑥([𝑦 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝑦𝜑)))
 
Theoremhbsb2 1757 Bound-variable hypothesis builder for substitution. (Contributed by NM, 5-Aug-1993.)
(¬ ∀𝑥 𝑥 = 𝑦 → ([𝑦 / 𝑥]𝜑 → ∀𝑥[𝑦 / 𝑥]𝜑))
 
Theoremnfsb2or 1758 Bound-variable hypothesis builder for substitution. Similar to hbsb2 1757 but in intuitionistic logic a disjunction is stronger than an implication. (Contributed by Jim Kingdon, 2-Feb-2018.)
(∀𝑥 𝑥 = 𝑦 ∨ Ⅎ𝑥[𝑦 / 𝑥]𝜑)
 
Theoremsbequilem 1759 Propositional logic lemma used in the sbequi 1760 proof. (Contributed by Jim Kingdon, 1-Feb-2018.)
(𝜑 ∨ (𝜓 → (𝜒𝜃)))    &   (𝜏 ∨ (𝜓 → (𝜃𝜂)))       (𝜑 ∨ (𝜏 ∨ (𝜓 → (𝜒𝜂))))
 
Theoremsbequi 1760 An equality theorem for substitution. (Contributed by NM, 5-Aug-1993.) (Proof modified by Jim Kingdon, 1-Feb-2018.)
(𝑥 = 𝑦 → ([𝑥 / 𝑧]𝜑 → [𝑦 / 𝑧]𝜑))
 
Theoremsbequ 1761 An equality theorem for substitution. Used in proof of Theorem 9.7 in [Megill] p. 449 (p. 16 of the preprint). (Contributed by NM, 5-Aug-1993.)
(𝑥 = 𝑦 → ([𝑥 / 𝑧]𝜑 ↔ [𝑦 / 𝑧]𝜑))
 
Theoremdrsb2 1762 Formula-building lemma for use with the Distinctor Reduction Theorem. Part of Theorem 9.4 of [Megill] p. 448 (p. 16 of preprint). (Contributed by NM, 27-Feb-2005.)
(∀𝑥 𝑥 = 𝑦 → ([𝑥 / 𝑧]𝜑 ↔ [𝑦 / 𝑧]𝜑))
 
Theoremspsbe 1763 A specialization theorem, mostly the same as Theorem 19.8 of [Margaris] p. 89. (Contributed by NM, 5-Aug-1993.) (Proof rewritten by Jim Kingdon, 29-Dec-2017.)
([𝑦 / 𝑥]𝜑 → ∃𝑥𝜑)
 
Theoremspsbim 1764 Specialization of implication. (Contributed by NM, 5-Aug-1993.) (Proof rewritten by Jim Kingdon, 21-Jan-2018.)
(∀𝑥(𝜑𝜓) → ([𝑦 / 𝑥]𝜑 → [𝑦 / 𝑥]𝜓))
 
Theoremspsbbi 1765 Specialization of biconditional. (Contributed by NM, 5-Aug-1993.) (Proof rewritten by Jim Kingdon, 21-Jan-2018.)
(∀𝑥(𝜑𝜓) → ([𝑦 / 𝑥]𝜑 ↔ [𝑦 / 𝑥]𝜓))
 
Theoremsbbidh 1766 Deduction substituting both sides of a biconditional. New proofs should use sbbid 1767 instead. (Contributed by NM, 5-Aug-1993.) (New usage is discouraged.)
(𝜑 → ∀𝑥𝜑)    &   (𝜑 → (𝜓𝜒))       (𝜑 → ([𝑦 / 𝑥]𝜓 ↔ [𝑦 / 𝑥]𝜒))
 
Theoremsbbid 1767 Deduction substituting both sides of a biconditional. (Contributed by NM, 30-Jun-1993.)
𝑥𝜑    &   (𝜑 → (𝜓𝜒))       (𝜑 → ([𝑦 / 𝑥]𝜓 ↔ [𝑦 / 𝑥]𝜒))
 
Theoremsbequ8 1768 Elimination of equality from antecedent after substitution. (Contributed by NM, 5-Aug-1993.) (Proof revised by Jim Kingdon, 20-Jan-2018.)
([𝑦 / 𝑥]𝜑 ↔ [𝑦 / 𝑥](𝑥 = 𝑦𝜑))
 
Theoremsbft 1769 Substitution has no effect on a non-free variable. (Contributed by NM, 30-May-2009.) (Revised by Mario Carneiro, 12-Oct-2016.) (Proof shortened by Wolf Lammen, 3-May-2018.)
(Ⅎ𝑥𝜑 → ([𝑦 / 𝑥]𝜑𝜑))
 
Theoremsbid2h 1770 An identity law for substitution. (Contributed by NM, 5-Aug-1993.)
(𝜑 → ∀𝑥𝜑)       ([𝑦 / 𝑥][𝑥 / 𝑦]𝜑𝜑)
 
Theoremsbid2 1771 An identity law for substitution. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 6-Oct-2016.)
𝑥𝜑       ([𝑦 / 𝑥][𝑥 / 𝑦]𝜑𝜑)
 
Theoremsbidm 1772 An idempotent law for substitution. (Contributed by NM, 30-Jun-1994.) (Proof rewritten by Jim Kingdon, 21-Jan-2018.)
([𝑦 / 𝑥][𝑦 / 𝑥]𝜑 ↔ [𝑦 / 𝑥]𝜑)
 
Theoremsb5rf 1773 Reversed substitution. (Contributed by NM, 3-Feb-2005.) (Proof shortened by Andrew Salmon, 25-May-2011.)
(𝜑 → ∀𝑦𝜑)       (𝜑 ↔ ∃𝑦(𝑦 = 𝑥 ∧ [𝑦 / 𝑥]𝜑))
 
Theoremsb6rf 1774 Reversed substitution. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.)
(𝜑 → ∀𝑦𝜑)       (𝜑 ↔ ∀𝑦(𝑦 = 𝑥 → [𝑦 / 𝑥]𝜑))
 
Theoremsb8h 1775 Substitution of variable in universal quantifier. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.) (Proof shortened by Jim Kingdon, 15-Jan-2018.)
(𝜑 → ∀𝑦𝜑)       (∀𝑥𝜑 ↔ ∀𝑦[𝑦 / 𝑥]𝜑)
 
Theoremsb8eh 1776 Substitution of variable in existential quantifier. (Contributed by NM, 12-Aug-1993.) (Proof rewritten by Jim Kingdon, 15-Jan-2018.)
(𝜑 → ∀𝑦𝜑)       (∃𝑥𝜑 ↔ ∃𝑦[𝑦 / 𝑥]𝜑)
 
Theoremsb8 1777 Substitution of variable in universal quantifier. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 6-Oct-2016.) (Proof shortened by Jim Kingdon, 15-Jan-2018.)
𝑦𝜑       (∀𝑥𝜑 ↔ ∀𝑦[𝑦 / 𝑥]𝜑)
 
Theoremsb8e 1778 Substitution of variable in existential quantifier. (Contributed by NM, 12-Aug-1993.) (Revised by Mario Carneiro, 6-Oct-2016.) (Proof shortened by Jim Kingdon, 15-Jan-2018.)
𝑦𝜑       (∃𝑥𝜑 ↔ ∃𝑦[𝑦 / 𝑥]𝜑)
 
1.4.4  Predicate calculus with distinct variables (cont.)
 
Theoremax16i 1779* Inference with ax-16 1735 as its conclusion, that doesn't require ax-10 1436, ax-11 1437, or ax-12 1442 for its proof. The hypotheses may be eliminable without one or more of these axioms in special cases. (Contributed by NM, 20-May-2008.)
(𝑥 = 𝑧 → (𝜑𝜓))    &   (𝜓 → ∀𝑥𝜓)       (∀𝑥 𝑥 = 𝑦 → (𝜑 → ∀𝑥𝜑))
 
Theoremax16ALT 1780* Version of ax16 1734 that doesn't require ax-10 1436 or ax-12 1442 for its proof. (Contributed by NM, 17-May-2008.) (Proof modification is discouraged.) (New usage is discouraged.)
(∀𝑥 𝑥 = 𝑦 → (𝜑 → ∀𝑥𝜑))
 
Theoremspv 1781* Specialization, using implicit substitition. (Contributed by NM, 30-Aug-1993.)
(𝑥 = 𝑦 → (𝜑𝜓))       (∀𝑥𝜑𝜓)
 
Theoremspimev 1782* Distinct-variable version of spime 1669. (Contributed by NM, 5-Aug-1993.)
(𝑥 = 𝑦 → (𝜑𝜓))       (𝜑 → ∃𝑥𝜓)
 
Theoremspeiv 1783* Inference from existential specialization, using implicit substitition. (Contributed by NM, 19-Aug-1993.)
(𝑥 = 𝑦 → (𝜑𝜓))    &   𝜓       𝑥𝜑
 
Theoremequvin 1784* A variable introduction law for equality. Lemma 15 of [Monk2] p. 109. (Contributed by NM, 5-Aug-1993.)
(𝑥 = 𝑦 ↔ ∃𝑧(𝑥 = 𝑧𝑧 = 𝑦))
 
Theorema16g 1785* A generalization of axiom ax-16 1735. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.)
(∀𝑥 𝑥 = 𝑦 → (𝜑 → ∀𝑧𝜑))
 
Theorema16gb 1786* A generalization of axiom ax-16 1735. (Contributed by NM, 5-Aug-1993.)
(∀𝑥 𝑥 = 𝑦 → (𝜑 ↔ ∀𝑧𝜑))
 
Theorema16nf 1787* If there is only one element in the universe, then everything satisfies . (Contributed by Mario Carneiro, 7-Oct-2016.)
(∀𝑥 𝑥 = 𝑦 → Ⅎ𝑧𝜑)
 
Theorem2albidv 1788* Formula-building rule for 2 existential quantifiers (deduction rule). (Contributed by NM, 4-Mar-1997.)
(𝜑 → (𝜓𝜒))       (𝜑 → (∀𝑥𝑦𝜓 ↔ ∀𝑥𝑦𝜒))
 
Theorem2exbidv 1789* Formula-building rule for 2 existential quantifiers (deduction rule). (Contributed by NM, 1-May-1995.)
(𝜑 → (𝜓𝜒))       (𝜑 → (∃𝑥𝑦𝜓 ↔ ∃𝑥𝑦𝜒))
 
Theorem3exbidv 1790* Formula-building rule for 3 existential quantifiers (deduction rule). (Contributed by NM, 1-May-1995.)
(𝜑 → (𝜓𝜒))       (𝜑 → (∃𝑥𝑦𝑧𝜓 ↔ ∃𝑥𝑦𝑧𝜒))
 
Theorem4exbidv 1791* Formula-building rule for 4 existential quantifiers (deduction rule). (Contributed by NM, 3-Aug-1995.)
(𝜑 → (𝜓𝜒))       (𝜑 → (∃𝑥𝑦𝑧𝑤𝜓 ↔ ∃𝑥𝑦𝑧𝑤𝜒))
 
Theorem19.9v 1792* Special case of Theorem 19.9 of [Margaris] p. 89. (Contributed by NM, 28-May-1995.) (Revised by NM, 21-May-2007.)
(∃𝑥𝜑𝜑)
 
Theoremexlimdd 1793 Existential elimination rule of natural deduction. (Contributed by Mario Carneiro, 9-Feb-2017.)
𝑥𝜑    &   𝑥𝜒    &   (𝜑 → ∃𝑥𝜓)    &   ((𝜑𝜓) → 𝜒)       (𝜑𝜒)
 
Theorem19.21v 1794* Special case of Theorem 19.21 of [Margaris] p. 90. Notational convention: We sometimes suffix with "v" the label of a theorem eliminating a hypothesis such as (𝜑 → ∀𝑥𝜑) in 19.21 1515 via the use of distinct variable conditions combined with ax-17 1459. Conversely, we sometimes suffix with "f" the label of a theorem introducing such a hypothesis to eliminate the need for the distinct variable condition; e.g. euf 1946 derived from df-eu 1944. The "f" stands for "not free in" which is less restrictive than "does not occur in." (Contributed by NM, 5-Aug-1993.)
(∀𝑥(𝜑𝜓) ↔ (𝜑 → ∀𝑥𝜓))
 
Theoremalrimiv 1795* Inference from Theorem 19.21 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.)
(𝜑𝜓)       (𝜑 → ∀𝑥𝜓)
 
Theoremalrimivv 1796* Inference from Theorem 19.21 of [Margaris] p. 90. (Contributed by NM, 31-Jul-1995.)
(𝜑𝜓)       (𝜑 → ∀𝑥𝑦𝜓)
 
Theoremalrimdv 1797* Deduction from Theorem 19.21 of [Margaris] p. 90. (Contributed by NM, 10-Feb-1997.)
(𝜑 → (𝜓𝜒))       (𝜑 → (𝜓 → ∀𝑥𝜒))
 
Theoremnfdv 1798* Apply the definition of not-free in a context. (Contributed by Mario Carneiro, 11-Aug-2016.)
(𝜑 → (𝜓 → ∀𝑥𝜓))       (𝜑 → Ⅎ𝑥𝜓)
 
Theorem2ax17 1799* Quantification of two variables over a formula in which they do not occur. (Contributed by Alan Sare, 12-Apr-2011.)
(𝜑 → ∀𝑥𝑦𝜑)
 
Theoremalimdv 1800* Deduction from Theorem 19.20 of [Margaris] p. 90. (Contributed by NM, 3-Apr-1994.)
(𝜑 → (𝜓𝜒))       (𝜑 → (∀𝑥𝜓 → ∀𝑥𝜒))
    < 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 >