![]() |
Intuitionistic Logic Explorer Theorem List (p. 18 of 95) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > ILE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Type | Label | Description |
---|---|---|
Statement | ||
Axiom | ax-11o 1701 |
Axiom ax-11o 1701 ("o" for "old") was the
original version of ax-11 1394,
before it was discovered (in Jan. 2007) that the shorter ax-11 1394 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 "¬ ∀xx = y
→..." as informally
meaning "if x and y are distinct variables then..." The
antecedent becomes false if the same variable is substituted for x and
y, ensuring
the theorem is sound whenever this is the case. In some
later theorems, we call an antecedent of the form ¬
∀xx = y a
"distinctor."
This axiom is redundant, as shown by theorem ax11o 1700. This axiom is obsolete and should no longer be used. It is proved above as theorem ax11o 1700. (Contributed by NM, 5-Aug-1993.) (New usage is discouraged.) |
⊢ (¬ ∀x x = y → (x = y → (φ → ∀x(x = y → φ)))) | ||
Theorem | albidv 1702* | Formula-building rule for universal quantifier (deduction rule). (Contributed by NM, 5-Aug-1993.) |
⊢ (φ → (ψ ↔ χ)) ⇒ ⊢ (φ → (∀xψ ↔ ∀xχ)) | ||
Theorem | exbidv 1703* | Formula-building rule for existential quantifier (deduction rule). (Contributed by NM, 5-Aug-1993.) |
⊢ (φ → (ψ ↔ χ)) ⇒ ⊢ (φ → (∃xψ ↔ ∃xχ)) | ||
Theorem | ax11b 1704 | A bidirectional version of ax-11o 1701. (Contributed by NM, 30-Jun-2006.) |
⊢ ((¬ ∀x x = y ∧ x = y) → (φ ↔ ∀x(x = y → φ))) | ||
Theorem | ax11v 1705* | This is a version of ax-11o 1701 when the variables are distinct. Axiom (C8) of [Monk2] p. 105. (Contributed by NM, 5-Aug-1993.) (Revised by Jim Kingdon, 15-Dec-2017.) |
⊢ (x = y → (φ → ∀x(x = y → φ))) | ||
Theorem | ax11ev 1706* | Analogue to ax11v 1705 for existential quantification. (Contributed by Jim Kingdon, 9-Jan-2018.) |
⊢ (x = y → (∃x(x = y ∧ φ) → φ)) | ||
Theorem | equs5 1707 | Lemma used in proofs of substitution properties. (Contributed by NM, 5-Aug-1993.) |
⊢ (¬ ∀x x = y → (∃x(x = y ∧ φ) → ∀x(x = y → φ))) | ||
Theorem | equs5or 1708 | Lemma used in proofs of substitution properties. Like equs5 1707 but, in intuitionistic logic, replacing negation and implication with disjunction makes this a stronger result. (Contributed by Jim Kingdon, 2-Feb-2018.) |
⊢ (∀x x = y ∨ (∃x(x = y ∧ φ) → ∀x(x = y → φ))) | ||
Theorem | sb3 1709 | One direction of a simplified definition of substitution when variables are distinct. (Contributed by NM, 5-Aug-1993.) |
⊢ (¬ ∀x x = y → (∃x(x = y ∧ φ) → [y / x]φ)) | ||
Theorem | sb4 1710 | One direction of a simplified definition of substitution when variables are distinct. (Contributed by NM, 5-Aug-1993.) |
⊢ (¬ ∀x x = y → ([y / x]φ → ∀x(x = y → φ))) | ||
Theorem | sb4or 1711 | One direction of a simplified definition of substitution when variables are distinct. Similar to sb4 1710 but stronger in intuitionistic logic. (Contributed by Jim Kingdon, 2-Feb-2018.) |
⊢ (∀x x = y ∨ ∀x([y / x]φ → ∀x(x = y → φ))) | ||
Theorem | sb4b 1712 | Simplified definition of substitution when variables are distinct. (Contributed by NM, 27-May-1997.) |
⊢ (¬ ∀x x = y → ([y / x]φ ↔ ∀x(x = y → φ))) | ||
Theorem | sb4bor 1713 | Simplified definition of substitution when variables are distinct, expressed via disjunction. (Contributed by Jim Kingdon, 18-Mar-2018.) |
⊢ (∀x x = y ∨ ∀x([y / x]φ ↔ ∀x(x = y → φ))) | ||
Theorem | hbsb2 1714 | Bound-variable hypothesis builder for substitution. (Contributed by NM, 5-Aug-1993.) |
⊢ (¬ ∀x x = y → ([y / x]φ → ∀x[y / x]φ)) | ||
Theorem | nfsb2or 1715 | Bound-variable hypothesis builder for substitution. Similar to hbsb2 1714 but in intuitionistic logic a disjunction is stronger than an implication. (Contributed by Jim Kingdon, 2-Feb-2018.) |
⊢ (∀x x = y ∨ Ⅎx[y / x]φ) | ||
Theorem | sbequilem 1716 | Propositional logic lemma used in the sbequi 1717 proof. (Contributed by Jim Kingdon, 1-Feb-2018.) |
⊢ (φ ∨ (ψ → (χ → θ))) & ⊢ (τ ∨ (ψ → (θ → η))) ⇒ ⊢ (φ ∨ (τ ∨ (ψ → (χ → η)))) | ||
Theorem | sbequi 1717 | An equality theorem for substitution. (Contributed by NM, 5-Aug-1993.) (Proof modified by Jim Kingdon, 1-Feb-2018.) |
⊢ (x = y → ([x / z]φ → [y / z]φ)) | ||
Theorem | sbequ 1718 | 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.) |
⊢ (x = y → ([x / z]φ ↔ [y / z]φ)) | ||
Theorem | drsb2 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.) |
⊢ (∀x x = y → ([x / z]φ ↔ [y / z]φ)) | ||
Theorem | spsbe 1720 | 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.) |
⊢ ([y / x]φ → ∃xφ) | ||
Theorem | spsbim 1721 | Specialization of implication. (Contributed by NM, 5-Aug-1993.) (Proof rewritten by Jim Kingdon, 21-Jan-2018.) |
⊢ (∀x(φ → ψ) → ([y / x]φ → [y / x]ψ)) | ||
Theorem | spsbbi 1722 | Specialization of biconditional. (Contributed by NM, 5-Aug-1993.) (Proof rewritten by Jim Kingdon, 21-Jan-2018.) |
⊢ (∀x(φ ↔ ψ) → ([y / x]φ ↔ [y / x]ψ)) | ||
Theorem | sbbid 1723 | Deduction substituting both sides of a biconditional. (Contributed by NM, 5-Aug-1993.) |
⊢ (φ → ∀xφ) & ⊢ (φ → (ψ ↔ χ)) ⇒ ⊢ (φ → ([y / x]ψ ↔ [y / x]χ)) | ||
Theorem | sbequ8 1724 | Elimination of equality from antecedent after substitution. (Contributed by NM, 5-Aug-1993.) (Proof revised by Jim Kingdon, 20-Jan-2018.) |
⊢ ([y / x]φ ↔ [y / x](x = y → φ)) | ||
Theorem | sbft 1725 | 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.) |
⊢ (Ⅎxφ → ([y / x]φ ↔ φ)) | ||
Theorem | sbid2h 1726 | An identity law for substitution. (Contributed by NM, 5-Aug-1993.) |
⊢ (φ → ∀xφ) ⇒ ⊢ ([y / x][x / y]φ ↔ φ) | ||
Theorem | sbid2 1727 | An identity law for substitution. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 6-Oct-2016.) |
⊢ Ⅎxφ ⇒ ⊢ ([y / x][x / y]φ ↔ φ) | ||
Theorem | sbidm 1728 | An idempotent law for substitution. (Contributed by NM, 30-Jun-1994.) (Proof rewritten by Jim Kingdon, 21-Jan-2018.) |
⊢ ([y / x][y / x]φ ↔ [y / x]φ) | ||
Theorem | sb5rf 1729 | Reversed substitution. (Contributed by NM, 3-Feb-2005.) (Proof shortened by Andrew Salmon, 25-May-2011.) |
⊢ (φ → ∀yφ) ⇒ ⊢ (φ ↔ ∃y(y = x ∧ [y / x]φ)) | ||
Theorem | sb6rf 1730 | Reversed substitution. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.) |
⊢ (φ → ∀yφ) ⇒ ⊢ (φ ↔ ∀y(y = x → [y / x]φ)) | ||
Theorem | sb8h 1731 | 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.) |
⊢ (φ → ∀yφ) ⇒ ⊢ (∀xφ ↔ ∀y[y / x]φ) | ||
Theorem | sb8eh 1732 | Substitution of variable in existential quantifier. (Contributed by NM, 12-Aug-1993.) (Proof rewritten by Jim Kingdon, 15-Jan-2018.) |
⊢ (φ → ∀yφ) ⇒ ⊢ (∃xφ ↔ ∃y[y / x]φ) | ||
Theorem | sb8 1733 | 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.) |
⊢ Ⅎyφ ⇒ ⊢ (∀xφ ↔ ∀y[y / x]φ) | ||
Theorem | sb8e 1734 | 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.) |
⊢ Ⅎyφ ⇒ ⊢ (∃xφ ↔ ∃y[y / x]φ) | ||
Theorem | ax16i 1735* | Inference with ax-16 1692 as its conclusion, that doesn't require ax-10 1393, ax-11 1394, or ax-12 1399 for its proof. The hypotheses may be eliminable without one or more of these axioms in special cases. (Contributed by NM, 20-May-2008.) |
⊢ (x = z → (φ ↔ ψ)) & ⊢ (ψ → ∀xψ) ⇒ ⊢ (∀x x = y → (φ → ∀xφ)) | ||
Theorem | ax16ALT 1736* | Version of ax16 1691 that doesn't require ax-10 1393 or ax-12 1399 for its proof. (Contributed by NM, 17-May-2008.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (∀x x = y → (φ → ∀xφ)) | ||
Theorem | spv 1737* | Specialization, using implicit substitition. (Contributed by NM, 30-Aug-1993.) |
⊢ (x = y → (φ ↔ ψ)) ⇒ ⊢ (∀xφ → ψ) | ||
Theorem | spimev 1738* | Distinct-variable version of spime 1626. (Contributed by NM, 5-Aug-1993.) |
⊢ (x = y → (φ → ψ)) ⇒ ⊢ (φ → ∃xψ) | ||
Theorem | speiv 1739* | Inference from existential specialization, using implicit substitition. (Contributed by NM, 19-Aug-1993.) |
⊢ (x = y → (φ ↔ ψ)) & ⊢ ψ ⇒ ⊢ ∃xφ | ||
Theorem | equvin 1740* | A variable introduction law for equality. Lemma 15 of [Monk2] p. 109. (Contributed by NM, 5-Aug-1993.) |
⊢ (x = y ↔ ∃z(x = z ∧ z = y)) | ||
Theorem | a16g 1741* | A generalization of axiom ax-16 1692. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.) |
⊢ (∀x x = y → (φ → ∀zφ)) | ||
Theorem | a16gb 1742* | A generalization of axiom ax-16 1692. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀x x = y → (φ ↔ ∀zφ)) | ||
Theorem | a16nf 1743* | If there is only one element in the universe, then everything satisfies Ⅎ. (Contributed by Mario Carneiro, 7-Oct-2016.) |
⊢ (∀x x = y → Ⅎzφ) | ||
Theorem | 2albidv 1744* | Formula-building rule for 2 existential quantifiers (deduction rule). (Contributed by NM, 4-Mar-1997.) |
⊢ (φ → (ψ ↔ χ)) ⇒ ⊢ (φ → (∀x∀yψ ↔ ∀x∀yχ)) | ||
Theorem | 2exbidv 1745* | Formula-building rule for 2 existential quantifiers (deduction rule). (Contributed by NM, 1-May-1995.) |
⊢ (φ → (ψ ↔ χ)) ⇒ ⊢ (φ → (∃x∃yψ ↔ ∃x∃yχ)) | ||
Theorem | 3exbidv 1746* | Formula-building rule for 3 existential quantifiers (deduction rule). (Contributed by NM, 1-May-1995.) |
⊢ (φ → (ψ ↔ χ)) ⇒ ⊢ (φ → (∃x∃y∃zψ ↔ ∃x∃y∃zχ)) | ||
Theorem | 4exbidv 1747* | Formula-building rule for 4 existential quantifiers (deduction rule). (Contributed by NM, 3-Aug-1995.) |
⊢ (φ → (ψ ↔ χ)) ⇒ ⊢ (φ → (∃x∃y∃z∃wψ ↔ ∃x∃y∃z∃wχ)) | ||
Theorem | 19.9v 1748* | Special case of Theorem 19.9 of [Margaris] p. 89. (Contributed by NM, 28-May-1995.) (Revised by NM, 21-May-2007.) |
⊢ (∃xφ ↔ φ) | ||
Theorem | exlimdd 1749 | Existential elimination rule of natural deduction. (Contributed by Mario Carneiro, 9-Feb-2017.) |
⊢ Ⅎxφ & ⊢ Ⅎxχ & ⊢ (φ → ∃xψ) & ⊢ ((φ ∧ ψ) → χ) ⇒ ⊢ (φ → χ) | ||
Theorem | 19.21v 1750* | 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 (φ → ∀xφ) in 19.21 1472 via the use of distinct variable conditions combined with ax-17 1416. 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 1902 derived from df-eu 1900. The "f" stands for "not free in" which is less restrictive than "does not occur in." (Contributed by NM, 5-Aug-1993.) |
⊢ (∀x(φ → ψ) ↔ (φ → ∀xψ)) | ||
Theorem | alrimiv 1751* | Inference from Theorem 19.21 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) |
⊢ (φ → ψ) ⇒ ⊢ (φ → ∀xψ) | ||
Theorem | alrimivv 1752* | Inference from Theorem 19.21 of [Margaris] p. 90. (Contributed by NM, 31-Jul-1995.) |
⊢ (φ → ψ) ⇒ ⊢ (φ → ∀x∀yψ) | ||
Theorem | alrimdv 1753* | Deduction from Theorem 19.21 of [Margaris] p. 90. (Contributed by NM, 10-Feb-1997.) |
⊢ (φ → (ψ → χ)) ⇒ ⊢ (φ → (ψ → ∀xχ)) | ||
Theorem | nfdv 1754* | Apply the definition of not-free in a context. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (φ → (ψ → ∀xψ)) ⇒ ⊢ (φ → Ⅎxψ) | ||
Theorem | 2ax17 1755* | Quantification of two variables over a formula in which they do not occur. (Contributed by Alan Sare, 12-Apr-2011.) |
⊢ (φ → ∀x∀yφ) | ||
Theorem | alimdv 1756* | Deduction from Theorem 19.20 of [Margaris] p. 90. (Contributed by NM, 3-Apr-1994.) |
⊢ (φ → (ψ → χ)) ⇒ ⊢ (φ → (∀xψ → ∀xχ)) | ||
Theorem | eximdv 1757* | Deduction from Theorem 19.22 of [Margaris] p. 90. (Contributed by NM, 27-Apr-1994.) |
⊢ (φ → (ψ → χ)) ⇒ ⊢ (φ → (∃xψ → ∃xχ)) | ||
Theorem | 2alimdv 1758* | Deduction from Theorem 19.22 of [Margaris] p. 90. (Contributed by NM, 27-Apr-2004.) |
⊢ (φ → (ψ → χ)) ⇒ ⊢ (φ → (∀x∀yψ → ∀x∀yχ)) | ||
Theorem | 2eximdv 1759* | Deduction from Theorem 19.22 of [Margaris] p. 90. (Contributed by NM, 3-Aug-1995.) |
⊢ (φ → (ψ → χ)) ⇒ ⊢ (φ → (∃x∃yψ → ∃x∃yχ)) | ||
Theorem | 19.23v 1760* | Special case of Theorem 19.23 of [Margaris] p. 90. (Contributed by NM, 28-Jun-1998.) |
⊢ (∀x(φ → ψ) ↔ (∃xφ → ψ)) | ||
Theorem | 19.23vv 1761* | Theorem 19.23 of [Margaris] p. 90 extended to two variables. (Contributed by NM, 10-Aug-2004.) |
⊢ (∀x∀y(φ → ψ) ↔ (∃x∃yφ → ψ)) | ||
Theorem | sb56 1762* | Two equivalent ways of expressing the proper substitution of y for x in φ, when x and y are distinct. Theorem 6.2 of [Quine] p. 40. The proof does not involve df-sb 1643. (Contributed by NM, 14-Apr-2008.) |
⊢ (∃x(x = y ∧ φ) ↔ ∀x(x = y → φ)) | ||
Theorem | sb6 1763* | Equivalence for substitution. Compare Theorem 6.2 of [Quine] p. 40. Also proved as Lemmas 16 and 17 of [Tarski] p. 70. (Contributed by NM, 18-Aug-1993.) (Revised by NM, 14-Apr-2008.) |
⊢ ([y / x]φ ↔ ∀x(x = y → φ)) | ||
Theorem | sb5 1764* | Equivalence for substitution. Similar to Theorem 6.1 of [Quine] p. 40. (Contributed by NM, 18-Aug-1993.) (Revised by NM, 14-Apr-2008.) |
⊢ ([y / x]φ ↔ ∃x(x = y ∧ φ)) | ||
Theorem | sbnv 1765* | Version of sbn 1823 where x and y are distinct. (Contributed by Jim Kingdon, 18-Dec-2017.) |
⊢ ([y / x] ¬ φ ↔ ¬ [y / x]φ) | ||
Theorem | sbanv 1766* | Version of sban 1826 where x and y are distinct. (Contributed by Jim Kingdon, 24-Dec-2017.) |
⊢ ([y / x](φ ∧ ψ) ↔ ([y / x]φ ∧ [y / x]ψ)) | ||
Theorem | sborv 1767* | Version of sbor 1825 where x and y are distinct. (Contributed by Jim Kingdon, 3-Feb-2018.) |
⊢ ([y / x](φ ∨ ψ) ↔ ([y / x]φ ∨ [y / x]ψ)) | ||
Theorem | sbi1v 1768* | Forward direction of sbimv 1770. (Contributed by Jim Kingdon, 25-Dec-2017.) |
⊢ ([y / x](φ → ψ) → ([y / x]φ → [y / x]ψ)) | ||
Theorem | sbi2v 1769* | Reverse direction of sbimv 1770. (Contributed by Jim Kingdon, 18-Jan-2018.) |
⊢ (([y / x]φ → [y / x]ψ) → [y / x](φ → ψ)) | ||
Theorem | sbimv 1770* | Intuitionistic proof of sbim 1824 where x and y are distinct. (Contributed by Jim Kingdon, 18-Jan-2018.) |
⊢ ([y / x](φ → ψ) ↔ ([y / x]φ → [y / x]ψ)) | ||
Theorem | sblimv 1771* | Version of sblim 1828 where x and y are distinct. (Contributed by Jim Kingdon, 19-Jan-2018.) |
⊢ (ψ → ∀xψ) ⇒ ⊢ ([y / x](φ → ψ) ↔ ([y / x]φ → ψ)) | ||
Theorem | pm11.53 1772* | Theorem *11.53 in [WhiteheadRussell] p. 164. (Contributed by Andrew Salmon, 24-May-2011.) |
⊢ (∀x∀y(φ → ψ) ↔ (∃xφ → ∀yψ)) | ||
Theorem | exlimivv 1773* | Inference from Theorem 19.23 of [Margaris] p. 90. (Contributed by NM, 1-Aug-1995.) |
⊢ (φ → ψ) ⇒ ⊢ (∃x∃yφ → ψ) | ||
Theorem | exlimdvv 1774* | Deduction from Theorem 19.23 of [Margaris] p. 90. (Contributed by NM, 31-Jul-1995.) |
⊢ (φ → (ψ → χ)) ⇒ ⊢ (φ → (∃x∃yψ → χ)) | ||
Theorem | exlimddv 1775* | Existential elimination rule of natural deduction. (Contributed by Mario Carneiro, 15-Jun-2016.) |
⊢ (φ → ∃xψ) & ⊢ ((φ ∧ ψ) → χ) ⇒ ⊢ (φ → χ) | ||
Theorem | 19.27v 1776* | Theorem 19.27 of [Margaris] p. 90. (Contributed by NM, 3-Jun-2004.) |
⊢ (∀x(φ ∧ ψ) ↔ (∀xφ ∧ ψ)) | ||
Theorem | 19.28v 1777* | Theorem 19.28 of [Margaris] p. 90. (Contributed by NM, 25-Mar-2004.) |
⊢ (∀x(φ ∧ ψ) ↔ (φ ∧ ∀xψ)) | ||
Theorem | 19.36aiv 1778* | Inference from Theorem 19.36 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) |
⊢ ∃x(φ → ψ) ⇒ ⊢ (∀xφ → ψ) | ||
Theorem | 19.41v 1779* | Special case of Theorem 19.41 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) |
⊢ (∃x(φ ∧ ψ) ↔ (∃xφ ∧ ψ)) | ||
Theorem | 19.41vv 1780* | Theorem 19.41 of [Margaris] p. 90 with 2 quantifiers. (Contributed by NM, 30-Apr-1995.) |
⊢ (∃x∃y(φ ∧ ψ) ↔ (∃x∃yφ ∧ ψ)) | ||
Theorem | 19.41vvv 1781* | Theorem 19.41 of [Margaris] p. 90 with 3 quantifiers. (Contributed by NM, 30-Apr-1995.) |
⊢ (∃x∃y∃z(φ ∧ ψ) ↔ (∃x∃y∃zφ ∧ ψ)) | ||
Theorem | 19.41vvvv 1782* | Theorem 19.41 of [Margaris] p. 90 with 4 quantifiers. (Contributed by FL, 14-Jul-2007.) |
⊢ (∃w∃x∃y∃z(φ ∧ ψ) ↔ (∃w∃x∃y∃zφ ∧ ψ)) | ||
Theorem | 19.42v 1783* | Special case of Theorem 19.42 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) |
⊢ (∃x(φ ∧ ψ) ↔ (φ ∧ ∃xψ)) | ||
Theorem | exdistr 1784* | Distribution of existential quantifiers. (Contributed by NM, 9-Mar-1995.) |
⊢ (∃x∃y(φ ∧ ψ) ↔ ∃x(φ ∧ ∃yψ)) | ||
Theorem | 19.42vv 1785* | Theorem 19.42 of [Margaris] p. 90 with 2 quantifiers. (Contributed by NM, 16-Mar-1995.) |
⊢ (∃x∃y(φ ∧ ψ) ↔ (φ ∧ ∃x∃yψ)) | ||
Theorem | 19.42vvv 1786* | Theorem 19.42 of [Margaris] p. 90 with 3 quantifiers. (Contributed by NM, 21-Sep-2011.) |
⊢ (∃x∃y∃z(φ ∧ ψ) ↔ (φ ∧ ∃x∃y∃zψ)) | ||
Theorem | 19.42vvvv 1787* | Theorem 19.42 of [Margaris] p. 90 with 4 quantifiers. (Contributed by Jim Kingdon, 23-Nov-2019.) |
⊢ (∃w∃x∃y∃z(φ ∧ ψ) ↔ (φ ∧ ∃w∃x∃y∃zψ)) | ||
Theorem | exdistr2 1788* | Distribution of existential quantifiers. (Contributed by NM, 17-Mar-1995.) |
⊢ (∃x∃y∃z(φ ∧ ψ) ↔ ∃x(φ ∧ ∃y∃zψ)) | ||
Theorem | 3exdistr 1789* | Distribution of existential quantifiers. (Contributed by NM, 9-Mar-1995.) (Proof shortened by Andrew Salmon, 25-May-2011.) |
⊢ (∃x∃y∃z(φ ∧ ψ ∧ χ) ↔ ∃x(φ ∧ ∃y(ψ ∧ ∃zχ))) | ||
Theorem | 4exdistr 1790* | Distribution of existential quantifiers. (Contributed by NM, 9-Mar-1995.) |
⊢ (∃x∃y∃z∃w((φ ∧ ψ) ∧ (χ ∧ θ)) ↔ ∃x(φ ∧ ∃y(ψ ∧ ∃z(χ ∧ ∃wθ)))) | ||
Theorem | cbvalv 1791* | Rule used to change bound variables, using implicit substitition. (Contributed by NM, 5-Aug-1993.) |
⊢ (x = y → (φ ↔ ψ)) ⇒ ⊢ (∀xφ ↔ ∀yψ) | ||
Theorem | cbvexv 1792* | Rule used to change bound variables, using implicit substitition. (Contributed by NM, 5-Aug-1993.) |
⊢ (x = y → (φ ↔ ψ)) ⇒ ⊢ (∃xφ ↔ ∃yψ) | ||
Theorem | cbval2 1793* | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 22-Dec-2003.) (Revised by Mario Carneiro, 6-Oct-2016.) (Proof shortened by Wolf Lammen, 22-Apr-2018.) |
⊢ Ⅎzφ & ⊢ Ⅎwφ & ⊢ Ⅎxψ & ⊢ Ⅎyψ & ⊢ ((x = z ∧ y = w) → (φ ↔ ψ)) ⇒ ⊢ (∀x∀yφ ↔ ∀z∀wψ) | ||
Theorem | cbvex2 1794* | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 14-Sep-2003.) (Revised by Mario Carneiro, 6-Oct-2016.) |
⊢ Ⅎzφ & ⊢ Ⅎwφ & ⊢ Ⅎxψ & ⊢ Ⅎyψ & ⊢ ((x = z ∧ y = w) → (φ ↔ ψ)) ⇒ ⊢ (∃x∃yφ ↔ ∃z∃wψ) | ||
Theorem | cbval2v 1795* | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 4-Feb-2005.) |
⊢ ((x = z ∧ y = w) → (φ ↔ ψ)) ⇒ ⊢ (∀x∀yφ ↔ ∀z∀wψ) | ||
Theorem | cbvex2v 1796* | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 26-Jul-1995.) |
⊢ ((x = z ∧ y = w) → (φ ↔ ψ)) ⇒ ⊢ (∃x∃yφ ↔ ∃z∃wψ) | ||
Theorem | cbvald 1797* | Deduction used to change bound variables, using implicit substitution, particularly useful in conjunction with dvelim 1890. (Contributed by NM, 2-Jan-2002.) (Revised by Mario Carneiro, 6-Oct-2016.) (Revised by Wolf Lammen, 13-May-2018.) |
⊢ Ⅎyφ & ⊢ (φ → Ⅎyψ) & ⊢ (φ → (x = y → (ψ ↔ χ))) ⇒ ⊢ (φ → (∀xψ ↔ ∀yχ)) | ||
Theorem | cbvexdh 1798* | Deduction used to change bound variables, using implicit substitition, particularly useful in conjunction with dvelim 1890. (Contributed by NM, 2-Jan-2002.) (Proof rewritten by Jim Kingdon, 30-Dec-2017.) |
⊢ (φ → ∀yφ) & ⊢ (φ → (ψ → ∀yψ)) & ⊢ (φ → (x = y → (ψ ↔ χ))) ⇒ ⊢ (φ → (∃xψ ↔ ∃yχ)) | ||
Theorem | cbvexd 1799* | Deduction used to change bound variables, using implicit substitution, particularly useful in conjunction with dvelim 1890. (Contributed by NM, 2-Jan-2002.) (Revised by Mario Carneiro, 6-Oct-2016.) (Proof rewritten by Jim Kingdon, 10-Jun-2018.) |
⊢ Ⅎyφ & ⊢ (φ → Ⅎyψ) & ⊢ (φ → (x = y → (ψ ↔ χ))) ⇒ ⊢ (φ → (∃xψ ↔ ∃yχ)) | ||
Theorem | cbvaldva 1800* | Rule used to change the bound variable in a universal quantifier with implicit substitution. Deduction form. (Contributed by David Moews, 1-May-2017.) |
⊢ ((φ ∧ x = y) → (ψ ↔ χ)) ⇒ ⊢ (φ → (∀xψ ↔ ∀yχ)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |