![]() |
Intuitionistic Logic Explorer Theorem List (p. 18 of 147) | < 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 | ||
Theorem | equid 1701 |
Identity law for equality (reflexivity). Lemma 6 of [Tarski] p. 68.
This is often an axiom of equality in textbook systems, but we don't
need it as an axiom since it can be proved from our other axioms.
This proof is similar to Tarski's and makes use of a dummy variable 𝑦. It also works in intuitionistic logic, unlike some other possible ways of proving this theorem. (Contributed by NM, 1-Apr-2005.) |
⊢ 𝑥 = 𝑥 | ||
Theorem | nfequid 1702 | Bound-variable hypothesis builder for 𝑥 = 𝑥. This theorem tells us that any variable, including 𝑥, is effectively not free in 𝑥 = 𝑥, even though 𝑥 is technically free according to the traditional definition of free variable. (Contributed by NM, 13-Jan-2011.) (Revised by NM, 21-Aug-2017.) |
⊢ Ⅎ𝑦 𝑥 = 𝑥 | ||
Theorem | stdpc6 1703 | One of the two equality axioms of standard predicate calculus, called reflexivity of equality. (The other one is stdpc7 1770.) Axiom 6 of [Mendelson] p. 95. Mendelson doesn't say why he prepended the redundant quantifier, but it was probably to be compatible with free logic (which is valid in the empty domain). (Contributed by NM, 16-Feb-2005.) |
⊢ ∀𝑥 𝑥 = 𝑥 | ||
Theorem | equcomi 1704 | Commutative law for equality. Lemma 7 of [Tarski] p. 69. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → 𝑦 = 𝑥) | ||
Theorem | ax6evr 1705* | A commuted form of a9ev 1697. The naming reflects how axioms were numbered in the Metamath Proof Explorer as of 2020 (a numbering which we eventually plan to adopt here too, but until this happens everywhere only some theorems will have it). (Contributed by BJ, 7-Dec-2020.) |
⊢ ∃𝑥 𝑦 = 𝑥 | ||
Theorem | equcom 1706 | Commutative law for equality. (Contributed by NM, 20-Aug-1993.) |
⊢ (𝑥 = 𝑦 ↔ 𝑦 = 𝑥) | ||
Theorem | equcomd 1707 | Deduction form of equcom 1706, symmetry of equality. For the versions for classes, see eqcom 2179 and eqcomd 2183. (Contributed by BJ, 6-Oct-2019.) |
⊢ (𝜑 → 𝑥 = 𝑦) ⇒ ⊢ (𝜑 → 𝑦 = 𝑥) | ||
Theorem | equcoms 1708 | An inference commuting equality in antecedent. Used to eliminate the need for a syllogism. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → 𝜑) ⇒ ⊢ (𝑦 = 𝑥 → 𝜑) | ||
Theorem | equtr 1709 | A transitive law for equality. (Contributed by NM, 23-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝑦 = 𝑧 → 𝑥 = 𝑧)) | ||
Theorem | equtrr 1710 | A transitive law for equality. Lemma L17 in [Megill] p. 446 (p. 14 of the preprint). (Contributed by NM, 23-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝑧 = 𝑥 → 𝑧 = 𝑦)) | ||
Theorem | equtr2 1711 | A transitive law for equality. (Contributed by NM, 12-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.) |
⊢ ((𝑥 = 𝑧 ∧ 𝑦 = 𝑧) → 𝑥 = 𝑦) | ||
Theorem | equequ1 1712 | An equivalence law for equality. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝑥 = 𝑧 ↔ 𝑦 = 𝑧)) | ||
Theorem | equequ2 1713 | An equivalence law for equality. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝑧 = 𝑥 ↔ 𝑧 = 𝑦)) | ||
Theorem | ax11i 1714 | Inference that has ax-11 1506 (without ∀𝑦) as its conclusion and does not require ax-10 1505, ax-11 1506, or ax12 1512 for its proof. The hypotheses may be eliminable without one or more of these axioms in special cases. Proof similar to Lemma 16 of [Tarski] p. 70. (Contributed by NM, 20-May-2008.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝜓 → ∀𝑥𝜓) ⇒ ⊢ (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | ax10o 1715 |
Show that ax-10o 1716 can be derived from ax-10 1505. An open problem is
whether this theorem can be derived from ax-10 1505 and the others when
ax-11 1506 is replaced with ax-11o 1823. See Theorem ax10 1717
for the
rederivation of ax-10 1505 from ax10o 1715.
Normally, ax10o 1715 should be used rather than ax-10o 1716, except by theorems specifically studying the latter's properties. (Contributed by NM, 16-May-2008.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (∀𝑥𝜑 → ∀𝑦𝜑)) | ||
Axiom | ax-10o 1716 |
Axiom ax-10o 1716 ("o" for "old") was the
original version of ax-10 1505,
before it was discovered (in May 2008) that the shorter ax-10 1505 could
replace it. It appears as Axiom scheme C11' in [Megill] p. 448 (p. 16 of
the preprint).
This axiom is redundant, as shown by Theorem ax10o 1715. Normally, ax10o 1715 should be used rather than ax-10o 1716, except by theorems specifically studying the latter's properties. (Contributed by NM, 5-Aug-1993.) (New usage is discouraged.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (∀𝑥𝜑 → ∀𝑦𝜑)) | ||
Theorem | ax10 1717 |
Rederivation of ax-10 1505 from original version ax-10o 1716. See Theorem
ax10o 1715 for the derivation of ax-10o 1716 from ax-10 1505.
This theorem should not be referenced in any proof. Instead, use ax-10 1505 above so that uses of ax-10 1505 can be more easily identified. (Contributed by NM, 16-May-2008.) (New usage is discouraged.) |
⊢ (∀𝑥 𝑥 = 𝑦 → ∀𝑦 𝑦 = 𝑥) | ||
Theorem | hbae 1718 | All variables are effectively bound in an identical variable specifier. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 3-Feb-2015.) |
⊢ (∀𝑥 𝑥 = 𝑦 → ∀𝑧∀𝑥 𝑥 = 𝑦) | ||
Theorem | nfae 1719 | All variables are effectively bound in an identical variable specifier. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ Ⅎ𝑧∀𝑥 𝑥 = 𝑦 | ||
Theorem | hbaes 1720 | Rule that applies hbae 1718 to antecedent. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑧∀𝑥 𝑥 = 𝑦 → 𝜑) ⇒ ⊢ (∀𝑥 𝑥 = 𝑦 → 𝜑) | ||
Theorem | hbnae 1721 | All variables are effectively bound in a distinct variable specifier. Lemma L19 in [Megill] p. 446 (p. 14 of the preprint). (Contributed by NM, 5-Aug-1993.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → ∀𝑧 ¬ ∀𝑥 𝑥 = 𝑦) | ||
Theorem | nfnae 1722 | All variables are effectively bound in a distinct variable specifier. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ Ⅎ𝑧 ¬ ∀𝑥 𝑥 = 𝑦 | ||
Theorem | hbnaes 1723 | Rule that applies hbnae 1721 to antecedent. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑧 ¬ ∀𝑥 𝑥 = 𝑦 → 𝜑) ⇒ ⊢ (¬ ∀𝑥 𝑥 = 𝑦 → 𝜑) | ||
Theorem | naecoms 1724 | A commutation rule for distinct variable specifiers. (Contributed by NM, 2-Jan-2002.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → 𝜑) ⇒ ⊢ (¬ ∀𝑦 𝑦 = 𝑥 → 𝜑) | ||
Theorem | equs4 1725 | Lemma used in proofs of substitution properties. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Mario Carneiro, 20-May-2014.) |
⊢ (∀𝑥(𝑥 = 𝑦 → 𝜑) → ∃𝑥(𝑥 = 𝑦 ∧ 𝜑)) | ||
Theorem | equsalh 1726 | A useful equivalence related to substitution. New proofs should use equsal 1727 instead. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 12-Aug-2011.) (New usage is discouraged.) |
⊢ (𝜓 → ∀𝑥𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥(𝑥 = 𝑦 → 𝜑) ↔ 𝜓) | ||
Theorem | equsal 1727 | A useful equivalence related to substitution. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 12-Aug-2011.) (Revised by Mario Carneiro, 3-Oct-2016.) (Proof shortened by Wolf Lammen, 5-Feb-2018.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥(𝑥 = 𝑦 → 𝜑) ↔ 𝜓) | ||
Theorem | equsex 1728 | A useful equivalence related to substitution. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 3-Feb-2015.) |
⊢ (𝜓 → ∀𝑥𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑥(𝑥 = 𝑦 ∧ 𝜑) ↔ 𝜓) | ||
Theorem | equsexd 1729 | Deduction form of equsex 1728. (Contributed by Jim Kingdon, 29-Dec-2017.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → (𝜒 → ∀𝑥𝜒)) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → (∃𝑥(𝑥 = 𝑦 ∧ 𝜓) ↔ 𝜒)) | ||
Theorem | dral1 1730 | 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, 24-Nov-1994.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 𝑥 = 𝑦 → (∀𝑥𝜑 ↔ ∀𝑦𝜓)) | ||
Theorem | dral2 1731 | 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.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 𝑥 = 𝑦 → (∀𝑧𝜑 ↔ ∀𝑧𝜓)) | ||
Theorem | drex2 1732 | 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.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 𝑥 = 𝑦 → (∃𝑧𝜑 ↔ ∃𝑧𝜓)) | ||
Theorem | drnf1 1733 | Formula-building lemma for use with the Distinctor Reduction Theorem. (Contributed by Mario Carneiro, 4-Oct-2016.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 𝑥 = 𝑦 → (Ⅎ𝑥𝜑 ↔ Ⅎ𝑦𝜓)) | ||
Theorem | drnf2 1734 | Formula-building lemma for use with the Distinctor Reduction Theorem. (Contributed by Mario Carneiro, 4-Oct-2016.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 𝑥 = 𝑦 → (Ⅎ𝑧𝜑 ↔ Ⅎ𝑧𝜓)) | ||
Theorem | spimth 1735 | Closed theorem form of spim 1738. (Contributed by NM, 15-Jan-2008.) (New usage is discouraged.) |
⊢ (∀𝑥((𝜓 → ∀𝑥𝜓) ∧ (𝑥 = 𝑦 → (𝜑 → 𝜓))) → (∀𝑥𝜑 → 𝜓)) | ||
Theorem | spimt 1736 | Closed theorem form of spim 1738. (Contributed by NM, 15-Jan-2008.) (Revised by Mario Carneiro, 17-Oct-2016.) (Proof shortened by Wolf Lammen, 24-Feb-2018.) |
⊢ ((Ⅎ𝑥𝜓 ∧ ∀𝑥(𝑥 = 𝑦 → (𝜑 → 𝜓))) → (∀𝑥𝜑 → 𝜓)) | ||
Theorem | spimh 1737 | Specialization, using implicit substitition. Compare Lemma 14 of [Tarski] p. 70. The spim 1738 series of theorems requires that only one direction of the substitution hypothesis hold. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 8-May-2008.) (New usage is discouraged.) |
⊢ (𝜓 → ∀𝑥𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (∀𝑥𝜑 → 𝜓) | ||
Theorem | spim 1738 | Specialization, using implicit substitution. Compare Lemma 14 of [Tarski] p. 70. The spim 1738 series of theorems requires that only one direction of the substitution hypothesis hold. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 3-Oct-2016.) (Proof rewritten by Jim Kingdon, 10-Jun-2018.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (∀𝑥𝜑 → 𝜓) | ||
Theorem | spimeh 1739 | Existential introduction, using implicit substitition. Compare Lemma 14 of [Tarski] p. 70. (Contributed by NM, 7-Aug-1994.) (Revised by NM, 3-Feb-2015.) (New usage is discouraged.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (𝜑 → ∃𝑥𝜓) | ||
Theorem | spimed 1740 | Deduction version of spime 1741. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 3-Oct-2016.) (Proof shortened by Wolf Lammen, 19-Feb-2018.) |
⊢ (𝜒 → Ⅎ𝑥𝜑) & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (𝜒 → (𝜑 → ∃𝑥𝜓)) | ||
Theorem | spime 1741 | Existential introduction, using implicit substitution. Compare Lemma 14 of [Tarski] p. 70. (Contributed by NM, 7-Aug-1994.) (Revised by Mario Carneiro, 3-Oct-2016.) (Proof shortened by Wolf Lammen, 6-Mar-2018.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (𝜑 → ∃𝑥𝜓) | ||
Theorem | cbv3 1742 | Rule used to change bound variables, using implicit substitution. Usage of this theorem is discouraged because proofs are encouraged to use the weaker cbv3v 1744 if possible. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 12-May-2018.) (New usage is discouraged.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (∀𝑥𝜑 → ∀𝑦𝜓) | ||
Theorem | cbv3h 1743 | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.) (Proof shortened by Wolf Lammen, 12-May-2018.) |
⊢ (𝜑 → ∀𝑦𝜑) & ⊢ (𝜓 → ∀𝑥𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (∀𝑥𝜑 → ∀𝑦𝜓) | ||
Theorem | cbv3v 1744* | Rule used to change bound variables, using implicit substitution. Version of cbv3 1742 with a disjoint variable condition. (Contributed by NM, 5-Aug-1993.) (Revised by BJ, 31-May-2019.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (∀𝑥𝜑 → ∀𝑦𝜓) | ||
Theorem | cbv1 1745 | Rule used to change bound variables, using implicit substitution. Revised to format hypotheses to common style. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 3-Oct-2016.) (Revised by Wolf Lammen, 13-May-2018.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑦𝜓) & ⊢ (𝜑 → Ⅎ𝑥𝜒) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 → 𝜒))) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 → ∀𝑦𝜒)) | ||
Theorem | cbv1h 1746 | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 13-May-2018.) |
⊢ (𝜑 → (𝜓 → ∀𝑦𝜓)) & ⊢ (𝜑 → (𝜒 → ∀𝑥𝜒)) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 → 𝜒))) ⇒ ⊢ (∀𝑥∀𝑦𝜑 → (∀𝑥𝜓 → ∀𝑦𝜒)) | ||
Theorem | cbv1v 1747* | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 5-Aug-1993.) (Revised by BJ, 16-Jun-2019.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑦𝜓) & ⊢ (𝜑 → Ⅎ𝑥𝜒) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 → 𝜒))) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 → ∀𝑦𝜒)) | ||
Theorem | cbv2h 1748 | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 → ∀𝑦𝜓)) & ⊢ (𝜑 → (𝜒 → ∀𝑥𝜒)) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (∀𝑥∀𝑦𝜑 → (∀𝑥𝜓 ↔ ∀𝑦𝜒)) | ||
Theorem | cbv2 1749 | Rule used to change bound variables, using implicit substitution. Revised to align format of hypotheses to common style. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 3-Oct-2016.) (Revised by Wolf Lammen, 13-May-2018.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑦𝜓) & ⊢ (𝜑 → Ⅎ𝑥𝜒) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 ↔ ∀𝑦𝜒)) | ||
Theorem | cbv2w 1750* | Rule used to change bound variables, using implicit substitution. Version of cbv2 1749 with a disjoint variable condition. (Contributed by NM, 5-Aug-1993.) (Revised by Gino Giotto, 10-Jan-2024.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑦𝜓) & ⊢ (𝜑 → Ⅎ𝑥𝜒) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 ↔ ∀𝑦𝜒)) | ||
Theorem | cbvalv1 1751* | Rule used to change bound variables, using implicit substitution. Version of cbval 1754 with a disjoint variable condition. See cbvalvw 1919 for a version with two disjoint variable conditions, and cbvalv 1917 for another variant. (Contributed by NM, 13-May-1993.) (Revised by BJ, 31-May-2019.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥𝜑 ↔ ∀𝑦𝜓) | ||
Theorem | cbvexv1 1752* | Rule used to change bound variables, using implicit substitution. Version of cbvex 1756 with a disjoint variable condition. See cbvexvw 1920 for a version with two disjoint variable conditions, and cbvexv 1918 for another variant. (Contributed by NM, 21-Jun-1993.) (Revised by BJ, 31-May-2019.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑥𝜑 ↔ ∃𝑦𝜓) | ||
Theorem | cbvalh 1753 | Rule used to change bound variables, using implicit substitition. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.) |
⊢ (𝜑 → ∀𝑦𝜑) & ⊢ (𝜓 → ∀𝑥𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥𝜑 ↔ ∀𝑦𝜓) | ||
Theorem | cbval 1754 | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 3-Oct-2016.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥𝜑 ↔ ∀𝑦𝜓) | ||
Theorem | cbvexh 1755 | Rule used to change bound variables, using implicit substitition. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 3-Feb-2015.) |
⊢ (𝜑 → ∀𝑦𝜑) & ⊢ (𝜓 → ∀𝑥𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑥𝜑 ↔ ∃𝑦𝜓) | ||
Theorem | cbvex 1756 | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 5-Aug-1993.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑥𝜑 ↔ ∃𝑦𝜓) | ||
Theorem | chvar 1757 | Implicit substitution of 𝑦 for 𝑥 into a theorem. (Contributed by Raph Levien, 9-Jul-2003.) (Revised by Mario Carneiro, 3-Oct-2016.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ 𝜑 ⇒ ⊢ 𝜓 | ||
Theorem | equvini 1758 | A variable introduction law for equality. Lemma 15 of [Monk2] p. 109, however we do not require 𝑧 to be distinct from 𝑥 and 𝑦 (making the proof longer). (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.) |
⊢ (𝑥 = 𝑦 → ∃𝑧(𝑥 = 𝑧 ∧ 𝑧 = 𝑦)) | ||
Theorem | equveli 1759 | A variable elimination law for equality with no distinct variable requirements. (Compare equvini 1758.) (Contributed by NM, 1-Mar-2013.) (Revised by NM, 3-Feb-2015.) |
⊢ (∀𝑧(𝑧 = 𝑥 ↔ 𝑧 = 𝑦) → 𝑥 = 𝑦) | ||
Theorem | nfald 1760 | If 𝑥 is not free in 𝜑, it is not free in ∀𝑦𝜑. (Contributed by Mario Carneiro, 24-Sep-2016.) (Proof shortened by Wolf Lammen, 6-Jan-2018.) |
⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥∀𝑦𝜓) | ||
Theorem | nfexd 1761 | If 𝑥 is not free in 𝜑, it is not free in ∃𝑦𝜑. (Contributed by Mario Carneiro, 24-Sep-2016.) (Proof rewritten by Jim Kingdon, 7-Feb-2018.) |
⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥∃𝑦𝜓) | ||
Syntax | wsb 1762 | Extend wff definition to include proper substitution (read "the wff that results when 𝑦 is properly substituted for 𝑥 in wff 𝜑"). (Contributed by NM, 24-Jan-2006.) |
wff [𝑦 / 𝑥]𝜑 | ||
Definition | df-sb 1763 |
Define proper substitution. Remark 9.1 in [Megill] p. 447 (p. 15 of the
preprint). For our notation, we use [𝑦 / 𝑥]𝜑 to mean "the wff
that results when 𝑦 is properly substituted for 𝑥 in the
wff
𝜑". We can also use [𝑦 / 𝑥]𝜑 in place of the "free for"
side condition used in traditional predicate calculus; see, for example,
stdpc4 1775.
Our notation was introduced in Haskell B. Curry's Foundations of Mathematical Logic (1977), p. 316 and is frequently used in textbooks of lambda calculus and combinatory logic. This notation improves the common but ambiguous notation, "𝜑(𝑦) is the wff that results when 𝑦 is properly substituted for 𝑥 in 𝜑(𝑥)". For example, if the original 𝜑(𝑥) is 𝑥 = 𝑦, then 𝜑(𝑦) is 𝑦 = 𝑦, from which we obtain that 𝜑(𝑥) is 𝑥 = 𝑥. So what exactly does 𝜑(𝑥) mean? Curry's notation solves this problem. In most books, proper substitution has a somewhat complicated recursive definition with multiple cases based on the occurrences of free and bound variables in the wff. Instead, we use a single formula that is exactly equivalent and gives us a direct definition. We later prove that our definition has the properties we expect of proper substitution (see Theorems sbequ 1840, sbcom2 1987 and sbid2v 1996). Note that our definition is valid even when 𝑥 and 𝑦 are replaced with the same variable, as sbid 1774 shows. We achieve this by having 𝑥 free in the first conjunct and bound in the second. We can also achieve this by using a dummy variable, as the alternate definition dfsb7 1991 shows (which some logicians may prefer because it doesn't mix free and bound variables). Another alternate definition which uses a dummy variable is dfsb7a 1994. When 𝑥 and 𝑦 are distinct, we can express proper substitution with the simpler expressions of sb5 1887 and sb6 1886. In classical logic, another possible definition is (𝑥 = 𝑦 ∧ 𝜑) ∨ ∀𝑥(𝑥 = 𝑦 → 𝜑) but we do not have an intuitionistic proof that this is equivalent. There are no restrictions on any of the variables, including what variables may occur in wff 𝜑. (Contributed by NM, 5-Aug-1993.) |
⊢ ([𝑦 / 𝑥]𝜑 ↔ ((𝑥 = 𝑦 → 𝜑) ∧ ∃𝑥(𝑥 = 𝑦 ∧ 𝜑))) | ||
Theorem | sbimi 1764 | Infer substitution into antecedent and consequent of an implication. (Contributed by NM, 25-Jun-1998.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 → [𝑦 / 𝑥]𝜓) | ||
Theorem | sbbii 1765 | Infer substitution into both sides of a logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 ↔ [𝑦 / 𝑥]𝜓) | ||
Theorem | sb1 1766 | One direction of a simplified definition of substitution. (Contributed by NM, 5-Aug-1993.) |
⊢ ([𝑦 / 𝑥]𝜑 → ∃𝑥(𝑥 = 𝑦 ∧ 𝜑)) | ||
Theorem | sb2 1767 | One direction of a simplified definition of substitution. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥(𝑥 = 𝑦 → 𝜑) → [𝑦 / 𝑥]𝜑) | ||
Theorem | sbequ1 1768 | An equality theorem for substitution. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝜑 → [𝑦 / 𝑥]𝜑)) | ||
Theorem | sbequ2 1769 | An equality theorem for substitution. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → ([𝑦 / 𝑥]𝜑 → 𝜑)) | ||
Theorem | stdpc7 1770 | One of the two equality axioms of standard predicate calculus, called substitutivity of equality. (The other one is stdpc6 1703.) Translated to traditional notation, it can be read: "𝑥 = 𝑦 → (𝜑(𝑥, 𝑥) → 𝜑(𝑥, 𝑦)), provided that 𝑦 is free for 𝑥 in 𝜑(𝑥, 𝑦)". Axiom 7 of [Mendelson] p. 95. (Contributed by NM, 15-Feb-2005.) |
⊢ (𝑥 = 𝑦 → ([𝑥 / 𝑦]𝜑 → 𝜑)) | ||
Theorem | sbequ12 1771 | An equality theorem for substitution. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ [𝑦 / 𝑥]𝜑)) | ||
Theorem | sbequ12r 1772 | An equality theorem for substitution. (Contributed by NM, 6-Oct-2004.) (Proof shortened by Andrew Salmon, 21-Jun-2011.) |
⊢ (𝑥 = 𝑦 → ([𝑥 / 𝑦]𝜑 ↔ 𝜑)) | ||
Theorem | sbequ12a 1773 | An equality theorem for substitution. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → ([𝑦 / 𝑥]𝜑 ↔ [𝑥 / 𝑦]𝜑)) | ||
Theorem | sbid 1774 | An identity theorem for substitution. Remark 9.1 in [Megill] p. 447 (p. 15 of the preprint). (Contributed by NM, 5-Aug-1993.) |
⊢ ([𝑥 / 𝑥]𝜑 ↔ 𝜑) | ||
Theorem | stdpc4 1775 | The specialization axiom of standard predicate calculus. It states that if a statement 𝜑 holds for all 𝑥, then it also holds for the specific case of 𝑦 (properly) substituted for 𝑥. Translated to traditional notation, it can be read: "∀𝑥𝜑(𝑥) → 𝜑(𝑦), provided that 𝑦 is free for 𝑥 in 𝜑(𝑥)". Axiom 4 of [Mendelson] p. 69. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥𝜑 → [𝑦 / 𝑥]𝜑) | ||
Theorem | sbh 1776 | Substitution for a variable not free in a wff does not affect it. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 17-Oct-2004.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 ↔ 𝜑) | ||
Theorem | sbf 1777 | Substitution for a variable not free in a wff does not affect it. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 4-Oct-2016.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ ([𝑦 / 𝑥]𝜑 ↔ 𝜑) | ||
Theorem | sbf2 1778 | Substitution has no effect on a bound variable. (Contributed by NM, 1-Jul-2005.) |
⊢ ([𝑦 / 𝑥]∀𝑥𝜑 ↔ ∀𝑥𝜑) | ||
Theorem | sb6x 1779 | Equivalence involving substitution for a variable not free. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 12-Aug-2011.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝑦 → 𝜑)) | ||
Theorem | nfs1f 1780 | If 𝑥 is not free in 𝜑, it is not free in [𝑦 / 𝑥]𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ Ⅎ𝑥[𝑦 / 𝑥]𝜑 | ||
Theorem | hbs1f 1781 | If 𝑥 is not free in 𝜑, it is not free in [𝑦 / 𝑥]𝜑. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon, 25-May-2011.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 → ∀𝑥[𝑦 / 𝑥]𝜑) | ||
Theorem | sbequ5 1782 | Substitution does not change an identical variable specifier. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 21-Dec-2004.) |
⊢ ([𝑤 / 𝑧]∀𝑥 𝑥 = 𝑦 ↔ ∀𝑥 𝑥 = 𝑦) | ||
Theorem | sbequ6 1783 | Substitution does not change a distinctor. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 14-May-2005.) |
⊢ ([𝑤 / 𝑧] ¬ ∀𝑥 𝑥 = 𝑦 ↔ ¬ ∀𝑥 𝑥 = 𝑦) | ||
Theorem | sbt 1784 | A substitution into a theorem remains true. (See chvar 1757 and chvarv 1937 for versions using implicit substitition.) (Contributed by NM, 21-Jan-2004.) (Proof shortened by Andrew Salmon, 25-May-2011.) |
⊢ 𝜑 ⇒ ⊢ [𝑦 / 𝑥]𝜑 | ||
Theorem | equsb1 1785 | Substitution applied to an atomic wff. (Contributed by NM, 5-Aug-1993.) |
⊢ [𝑦 / 𝑥]𝑥 = 𝑦 | ||
Theorem | equsb2 1786 | Substitution applied to an atomic wff. (Contributed by NM, 5-Aug-1993.) |
⊢ [𝑦 / 𝑥]𝑦 = 𝑥 | ||
Theorem | sbiedh 1787 | Conversion of implicit substitution to explicit substitution (deduction version of sbieh 1790). New proofs should use sbied 1788 instead. (Contributed by NM, 30-Jun-1994.) (Proof shortened by Andrew Salmon, 25-May-2011.) (New usage is discouraged.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → (𝜒 → ∀𝑥𝜒)) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → ([𝑦 / 𝑥]𝜓 ↔ 𝜒)) | ||
Theorem | sbied 1788 | Conversion of implicit substitution to explicit substitution (deduction version of sbie 1791). (Contributed by NM, 30-Jun-1994.) (Revised by Mario Carneiro, 4-Oct-2016.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜒) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → ([𝑦 / 𝑥]𝜓 ↔ 𝜒)) | ||
Theorem | sbiedv 1789* | Conversion of implicit substitution to explicit substitution (deduction version of sbie 1791). (Contributed by NM, 7-Jan-2017.) |
⊢ ((𝜑 ∧ 𝑥 = 𝑦) → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → ([𝑦 / 𝑥]𝜓 ↔ 𝜒)) | ||
Theorem | sbieh 1790 | Conversion of implicit substitution to explicit substitution. New proofs should use sbie 1791 instead. (Contributed by NM, 30-Jun-1994.) (New usage is discouraged.) |
⊢ (𝜓 → ∀𝑥𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 ↔ 𝜓) | ||
Theorem | sbie 1791 | 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.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 ↔ 𝜓) | ||
Theorem | sbiev 1792* | Conversion of implicit substitution to explicit substitution. Version of sbie 1791 with a disjoint variable condition. (Contributed by Wolf Lammen, 18-Jan-2023.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 ↔ 𝜓) | ||
Theorem | equsalv 1793* | An equivalence related to implicit substitution. Version of equsal 1727 with a disjoint variable condition. (Contributed by NM, 2-Jun-1993.) (Revised by BJ, 31-May-2019.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥(𝑥 = 𝑦 → 𝜑) ↔ 𝜓) | ||
Theorem | equs5a 1794 | A property related to substitution that unlike equs5 1829 doesn't require a distinctor antecedent. (Contributed by NM, 2-Feb-2007.) |
⊢ (∃𝑥(𝑥 = 𝑦 ∧ ∀𝑦𝜑) → ∀𝑥(𝑥 = 𝑦 → 𝜑)) | ||
Theorem | equs5e 1795 | A property related to substitution that unlike equs5 1829 doesn't require a distinctor antecedent. (Contributed by NM, 2-Feb-2007.) (Revised by NM, 3-Feb-2015.) |
⊢ (∃𝑥(𝑥 = 𝑦 ∧ 𝜑) → ∀𝑥(𝑥 = 𝑦 → ∃𝑦𝜑)) | ||
Theorem | ax11e 1796 | Analogue to ax-11 1506 but for existential quantification. (Contributed by Mario Carneiro and Jim Kingdon, 31-Dec-2017.) (Proved by Mario Carneiro, 9-Feb-2018.) |
⊢ (𝑥 = 𝑦 → (∃𝑥(𝑥 = 𝑦 ∧ 𝜑) → ∃𝑦𝜑)) | ||
Theorem | ax10oe 1797 | Quantifier Substitution for existential quantifiers. Analogue to ax10o 1715 but for ∃ rather than ∀. (Contributed by Jim Kingdon, 21-Dec-2017.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (∃𝑥𝜓 → ∃𝑦𝜓)) | ||
Theorem | drex1 1798 | 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.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 𝑥 = 𝑦 → (∃𝑥𝜑 ↔ ∃𝑦𝜓)) | ||
Theorem | drsb1 1799 | 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.) |
⊢ (∀𝑥 𝑥 = 𝑦 → ([𝑧 / 𝑥]𝜑 ↔ [𝑧 / 𝑦]𝜑)) | ||
Theorem | exdistrfor 1800 | 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.) |
⊢ (∀𝑥 𝑥 = 𝑦 ∨ ∀𝑥Ⅎ𝑦𝜑) ⇒ ⊢ (∃𝑥∃𝑦(𝜑 ∧ 𝜓) → ∃𝑥(𝜑 ∧ ∃𝑦𝜓)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |