![]() |
Metamath
Proof Explorer Theorem List (p. 21 of 431) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > MPE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Color key: | ![]() (1-28080) |
![]() (28081-29605) |
![]() (29606-43060) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | alrimivv 2001* | Inference form of Theorem 19.21 of [Margaris] p. 90. See 19.21 2218 and 19.21v 2013. (Contributed by NM, 31-Jul-1995.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (𝜑 → ∀𝑥∀𝑦𝜓) | ||
Theorem | alrimdv 2002* | Deduction form of Theorem 19.21 of [Margaris] p. 90. See 19.21 2218 and 19.21v 2013. (Contributed by NM, 10-Feb-1997.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → ∀𝑥𝜒)) | ||
Theorem | exlimiv 2003* |
Inference form of Theorem 19.23 of [Margaris]
p. 90, see 19.23 2223.
See exlimi 2229 for a more general version requiring more axioms. This inference, along with its many variants such as rexlimdv 3164, is used to implement a metatheorem called "Rule C" that is given in many logic textbooks. See, for example, Rule C in [Mendelson] p. 81, Rule C in [Margaris] p. 40, or Rule C in Hirst and Hirst's A Primer for Logic and Proof p. 59 (PDF p. 65) at http://www.appstate.edu/~hirstjl/primer/hirst.pdf. In informal proofs, the statement "Let 𝐶 be an element such that..." almost always means an implicit application of Rule C. In essence, Rule C states that if we can prove that some element 𝑥 exists satisfying a wff, i.e. ∃𝑥𝜑(𝑥) where 𝜑(𝑥) has 𝑥 free, then we can use 𝜑(𝐶) as a hypothesis for the proof where 𝐶 is a new (fictitious) constant not appearing previously in the proof, nor in any axioms used, nor in the theorem to be proved. The purpose of Rule C is to get rid of the existential quantifier. We cannot do this in Metamath directly. Instead, we use the original 𝜑 (containing 𝑥) as an antecedent for the main part of the proof. We eventually arrive at (𝜑 → 𝜓) where 𝜓 is the theorem to be proved and does not contain 𝑥. Then we apply exlimiv 2003 to arrive at (∃𝑥𝜑 → 𝜓). Finally, we separately prove ∃𝑥𝜑 and detach it with modus ponens ax-mp 5 to arrive at the final theorem 𝜓. (Contributed by NM, 21-Jun-1993.) Remove dependencies on ax-6 2050 and ax-8 2137. (Revised by Wolf Lammen, 4-Dec-2017.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (∃𝑥𝜑 → 𝜓) | ||
Theorem | exlimiiv 2004* | Inference associated with exlimiv 2003. (Contributed by BJ, 19-Dec-2020.) |
⊢ (𝜑 → 𝜓) & ⊢ ∃𝑥𝜑 ⇒ ⊢ 𝜓 | ||
Theorem | exlimivv 2005* | Inference form of Theorem 19.23 of [Margaris] p. 90, see 19.23 2223. (Contributed by NM, 1-Aug-1995.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (∃𝑥∃𝑦𝜑 → 𝜓) | ||
Theorem | exlimdv 2006* | Deduction form of Theorem 19.23 of [Margaris] p. 90, see 19.23 2223. (Contributed by NM, 27-Apr-1994.) Remove dependencies on ax-6 2050, ax-7 2086. (Revised by Wolf Lammen, 4-Dec-2017.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (∃𝑥𝜓 → 𝜒)) | ||
Theorem | exlimdvv 2007* | Deduction form of Theorem 19.23 of [Margaris] p. 90, see 19.23 2223. (Contributed by NM, 31-Jul-1995.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (∃𝑥∃𝑦𝜓 → 𝜒)) | ||
Theorem | exlimddv 2008* | Existential elimination rule of natural deduction. (Contributed by Mario Carneiro, 15-Jun-2016.) |
⊢ (𝜑 → ∃𝑥𝜓) & ⊢ ((𝜑 ∧ 𝜓) → 𝜒) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | nexdv 2009* | Deduction for generalization rule for negated wff. (Contributed by NM, 5-Aug-1993.) Reduce dependencies on axioms. (Revised by Wolf Lammen, 13-Jul-2020.) (Proof shortened by Wolf Lammen, 10-Oct-2021.) |
⊢ (𝜑 → ¬ 𝜓) ⇒ ⊢ (𝜑 → ¬ ∃𝑥𝜓) | ||
Theorem | nexdvOLD 2010* | Obsolete proof of nexdv 2009 as of 10-Oct-2021. (Contributed by NM, 5-Aug-1993.) Reduce dependencies on axioms. (Revised by Wolf Lammen, 13-Jul-2020.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝜑 → ¬ 𝜓) ⇒ ⊢ (𝜑 → ¬ ∃𝑥𝜓) | ||
Theorem | 2ax5 2011* | Quantification of two variables over a formula in which they do not occur. (Contributed by Alan Sare, 12-Apr-2011.) |
⊢ (𝜑 → ∀𝑥∀𝑦𝜑) | ||
Theorem | stdpc5v 2012* | Version of stdpc5 2219 with a dv condition, requiring fewer axioms. (Contributed by BJ, 7-Mar-2020.) Revised to shorten 19.21v 2013. (Revised by Wolf Lammen, 12-Jul-2020.) |
⊢ (∀𝑥(𝜑 → 𝜓) → (𝜑 → ∀𝑥𝜓)) | ||
Theorem | 19.21v 2013* |
Version of 19.21 2218 with a dv condition, requiring fewer axioms.
Notational convention: We sometimes suffix with "v" the label of a theorem using a distinct variable ("dv") condition instead of a non-freeness hypothesis such as Ⅎ𝑥𝜑. Conversely, we sometimes suffix with "f" the label of a theorem introducing such a non-freeness hypothesis ("f" stands for "not free in", see df-nf 1855) instead of a dv condition. For instance, 19.21v 2013 versus 19.21 2218 and vtoclf 3394 versus vtocl 3395. Note that "not free in" is less restrictive than "does not occur in." Note that the version with a dv condition is easily proved from the version with the corresponding non-freeness hypothesis, by using nfv 1988. However, the dv version can often be proved from fewer axioms. (Contributed by NM, 21-Jun-1993.) Reduce dependencies on axioms. (Revised by Wolf Lammen, 2-Jan-2020.) (Proof shortened by Wolf Lammen, 12-Jul-2020.) |
⊢ (∀𝑥(𝜑 → 𝜓) ↔ (𝜑 → ∀𝑥𝜓)) | ||
Theorem | 19.32v 2014* | Version of 19.32 2244 with a dv condition, requiring fewer axioms. (Contributed by BJ, 7-Mar-2020.) |
⊢ (∀𝑥(𝜑 ∨ 𝜓) ↔ (𝜑 ∨ ∀𝑥𝜓)) | ||
Theorem | 19.31v 2015* | Version of 19.31 2245 with a dv condition, requiring fewer axioms. (Contributed by BJ, 7-Mar-2020.) |
⊢ (∀𝑥(𝜑 ∨ 𝜓) ↔ (∀𝑥𝜑 ∨ 𝜓)) | ||
Theorem | 19.23v 2016* | Version of 19.23 2223 with a dv condition instead of a non-freeness hypothesis. (Contributed by NM, 28-Jun-1998.) Reduce dependencies on axioms. (Revised by Wolf Lammen, 11-Jan-2020.) Remove dependency on ax-6 2050. (Revised by Rohan Ridenour, 15-Apr-2022.) |
⊢ (∀𝑥(𝜑 → 𝜓) ↔ (∃𝑥𝜑 → 𝜓)) | ||
Theorem | 19.23vv 2017* | Theorem 19.23v 2016 extended to two variables. (Contributed by NM, 10-Aug-2004.) |
⊢ (∀𝑥∀𝑦(𝜑 → 𝜓) ↔ (∃𝑥∃𝑦𝜑 → 𝜓)) | ||
Theorem | 19.36imv 2018* | One direction of 19.36v 2065 that can be proven without ax-6 2050. (Contributed by Rohan Ridenour, 16-Apr-2022.) |
⊢ (∃𝑥(𝜑 → 𝜓) → (∀𝑥𝜑 → 𝜓)) | ||
Theorem | 19.36iv 2019* | Inference associated with 19.36v 2065. Version of 19.36i 2242 with a dv condition. (Contributed by NM, 5-Aug-1993.) Reduce dependencies on axioms. (Revised by Wolf Lammen, 17-Jan-2020.) Remove dependency on ax-6 2050. (Revised by Rohan Ridenour, 15-Apr-2022.) |
⊢ ∃𝑥(𝜑 → 𝜓) ⇒ ⊢ (∀𝑥𝜑 → 𝜓) | ||
Theorem | 19.37imv 2020* | One direction of 19.37v 2071 that can be proven without ax-6 2050. (Contributed by Rohan Ridenour, 16-Apr-2022.) |
⊢ (∃𝑥(𝜑 → 𝜓) → (𝜑 → ∃𝑥𝜓)) | ||
Theorem | 19.37iv 2021* | Inference associated with 19.37v 2071. (Contributed by NM, 5-Aug-1993.) Remove dependency on ax-6 2050. (Revised by Rohan Ridenour, 15-Apr-2022.) |
⊢ ∃𝑥(𝜑 → 𝜓) ⇒ ⊢ (𝜑 → ∃𝑥𝜓) | ||
Theorem | 19.41v 2022* | Version of 19.41 2246 with a dv condition, requiring fewer axioms. (Contributed by NM, 21-Jun-1993.) Remove dependency on ax-6 2050. (Revised by Rohan Ridenour, 15-Apr-2022.) |
⊢ (∃𝑥(𝜑 ∧ 𝜓) ↔ (∃𝑥𝜑 ∧ 𝜓)) | ||
Theorem | 19.41vv 2023* | Version of 19.41 2246 with two quantifiers and a dv condition requiring fewer axioms. (Contributed by NM, 30-Apr-1995.) |
⊢ (∃𝑥∃𝑦(𝜑 ∧ 𝜓) ↔ (∃𝑥∃𝑦𝜑 ∧ 𝜓)) | ||
Theorem | 19.41vvv 2024* | Version of 19.41 2246 with three quantifiers and a dv condition requiring fewer axioms. (Contributed by NM, 30-Apr-1995.) |
⊢ (∃𝑥∃𝑦∃𝑧(𝜑 ∧ 𝜓) ↔ (∃𝑥∃𝑦∃𝑧𝜑 ∧ 𝜓)) | ||
Theorem | 19.41vvvv 2025* | Version of 19.41 2246 with four quantifiers and a dv condition requiring fewer axioms. (Contributed by FL, 14-Jul-2007.) |
⊢ (∃𝑤∃𝑥∃𝑦∃𝑧(𝜑 ∧ 𝜓) ↔ (∃𝑤∃𝑥∃𝑦∃𝑧𝜑 ∧ 𝜓)) | ||
Theorem | 19.42v 2026* | Version of 19.42 2248 with a dv condition requiring fewer axioms. (Contributed by NM, 21-Jun-1993.) |
⊢ (∃𝑥(𝜑 ∧ 𝜓) ↔ (𝜑 ∧ ∃𝑥𝜓)) | ||
Theorem | exdistr 2027* | Distribution of existential quantifiers. (Contributed by NM, 9-Mar-1995.) |
⊢ (∃𝑥∃𝑦(𝜑 ∧ 𝜓) ↔ ∃𝑥(𝜑 ∧ ∃𝑦𝜓)) | ||
Theorem | 19.42vv 2028* | Version of 19.42 2248 with two quantifiers and a dv condition requiring fewer axioms. (Contributed by NM, 16-Mar-1995.) |
⊢ (∃𝑥∃𝑦(𝜑 ∧ 𝜓) ↔ (𝜑 ∧ ∃𝑥∃𝑦𝜓)) | ||
Theorem | 19.42vvv 2029* | Version of 19.42 2248 with three quantifiers and a dv condition requiring fewer axioms. (Contributed by NM, 21-Sep-2011.) |
⊢ (∃𝑥∃𝑦∃𝑧(𝜑 ∧ 𝜓) ↔ (𝜑 ∧ ∃𝑥∃𝑦∃𝑧𝜓)) | ||
Theorem | exdistr2 2030* | Distribution of existential quantifiers. (Contributed by NM, 17-Mar-1995.) |
⊢ (∃𝑥∃𝑦∃𝑧(𝜑 ∧ 𝜓) ↔ ∃𝑥(𝜑 ∧ ∃𝑦∃𝑧𝜓)) | ||
Theorem | 3exdistr 2031* | Distribution of existential quantifiers in a triple conjunction. (Contributed by NM, 9-Mar-1995.) (Proof shortened by Andrew Salmon, 25-May-2011.) |
⊢ (∃𝑥∃𝑦∃𝑧(𝜑 ∧ 𝜓 ∧ 𝜒) ↔ ∃𝑥(𝜑 ∧ ∃𝑦(𝜓 ∧ ∃𝑧𝜒))) | ||
Theorem | 4exdistr 2032* | Distribution of existential quantifiers in a quadruple conjunction. (Contributed by NM, 9-Mar-1995.) (Proof shortened by Wolf Lammen, 20-Jan-2018.) |
⊢ (∃𝑥∃𝑦∃𝑧∃𝑤((𝜑 ∧ 𝜓) ∧ (𝜒 ∧ 𝜃)) ↔ ∃𝑥(𝜑 ∧ ∃𝑦(𝜓 ∧ ∃𝑧(𝜒 ∧ ∃𝑤𝜃)))) | ||
Theorem | nfvOLD 2033* | Obsolete proof of nfv 1988 as of 5-Oct-2021. (Contributed by Mario Carneiro, 11-Aug-2016.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ Ⅎ𝑥𝜑 | ||
Theorem | nfvdOLD 2034* | Obsolete proof of nfvd 1989 as of 6-Oct-2021. (Contributed by Mario Carneiro, 6-Oct-2016.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
Theorem | nfdvOLD 2035* | Obsolete proof of nf5dv 2170 as of 6-Oct-2021. (Contributed by Mario Carneiro, 11-Aug-2016.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
The equality predicate was introduced above in wceq 1628 for use by df-tru 1631. See the comments in that section. In this section, we continue with the first "real" use of it. | ||
Theorem | weq 2036 |
Extend wff definition to include atomic formulas using the equality
predicate.
(Instead of introducing weq 2036 as an axiomatic statement, as was done in an older version of this database, we introduce it by "proving" a special case of set theory's more general wceq 1628. This lets us avoid overloading the = connective, thus preventing ambiguity that would complicate certain Metamath parsers. However, logically weq 2036 is considered to be a primitive syntax, even though here it is artificially "derived" from wceq 1628. Note: To see the proof steps of this syntax proof, type "show proof weq /all" in the Metamath program.) (Contributed by NM, 24-Jan-2006.) |
wff 𝑥 = 𝑦 | ||
Theorem | equs3 2037 | Lemma used in proofs of substitution properties. (Contributed by NM, 10-May-1993.) |
⊢ (∃𝑥(𝑥 = 𝑦 ∧ 𝜑) ↔ ¬ ∀𝑥(𝑥 = 𝑦 → ¬ 𝜑)) | ||
Theorem | speimfw 2038 | Specialization, with additional weakening (compared to 19.2 2054) to allow bundling of 𝑥 and 𝑦. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 23-Apr-2017.) (Proof shortened by Wolf Lammen, 5-Dec-2017.) |
⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (¬ ∀𝑥 ¬ 𝑥 = 𝑦 → (∀𝑥𝜑 → ∃𝑥𝜓)) | ||
Theorem | speimfwALT 2039 | Alternate proof of speimfw 2038 (longer compressed proof, but fewer essential steps). (Contributed by NM, 23-Apr-2017.) (Proof shortened by Wolf Lammen, 5-Aug-2017.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (¬ ∀𝑥 ¬ 𝑥 = 𝑦 → (∀𝑥𝜑 → ∃𝑥𝜓)) | ||
Theorem | spimfw 2040 | Specialization, with additional weakening (compared to sp 2196) to allow bundling of 𝑥 and 𝑦. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 23-Apr-2017.) (Proof shortened by Wolf Lammen, 7-Aug-2017.) |
⊢ (¬ 𝜓 → ∀𝑥 ¬ 𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (¬ ∀𝑥 ¬ 𝑥 = 𝑦 → (∀𝑥𝜑 → 𝜓)) | ||
Theorem | ax12i 2041 | Inference that has ax-12 2192 (without ∀𝑦) as its conclusion. Uses only Tarski's FOL axiom schemes. The hypotheses may be eliminable without using ax-12 2192 in special cases. Proof similar to Lemma 16 of [Tarski] p. 70. (Contributed by NM, 20-May-2008.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝜓 → ∀𝑥𝜓) ⇒ ⊢ (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Syntax | wsb 2042 | 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 2043 |
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 from the proper substitution of 𝑦 for 𝑥 in the
wff
𝜑." That is, 𝑦
properly replaces 𝑥. For example,
[𝑥 /
𝑦]𝑧 ∈ 𝑦 is the same as 𝑧 ∈ 𝑥, as shown in elsb4 2568. We
can also use [𝑦 / 𝑥]𝜑 in place of the "free for"
side condition
used in traditional predicate calculus; see, for example, stdpc4 2486.
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 2509, sbcom2 2578 and sbid2v 2590). Note that our definition is valid even when 𝑥 and 𝑦 are replaced with the same variable, as sbid 2257 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 2588 shows (which some logicians may prefer because it doesn't mix free and bound variables). Another version that mixes free and bound variables is dfsb3 2507. When 𝑥 and 𝑦 are distinct, we can express proper substitution with the simpler expressions of sb5 2563 and sb6 2562. There are no restrictions on any of the variables, including what variables may occur in wff 𝜑. (Contributed by NM, 10-May-1993.) |
⊢ ([𝑦 / 𝑥]𝜑 ↔ ((𝑥 = 𝑦 → 𝜑) ∧ ∃𝑥(𝑥 = 𝑦 ∧ 𝜑))) | ||
Theorem | sbequ2 2044 | An equality theorem for substitution. (Contributed by NM, 16-May-1993.) (Proof shortened by Wolf Lammen, 25-Feb-2018.) |
⊢ (𝑥 = 𝑦 → ([𝑦 / 𝑥]𝜑 → 𝜑)) | ||
Theorem | sb1 2045 | One direction of a simplified definition of substitution. The converse requires either a dv condition (sb5 2563) or a non-freeness hypothesis (sb5f 2519). (Contributed by NM, 13-May-1993.) |
⊢ ([𝑦 / 𝑥]𝜑 → ∃𝑥(𝑥 = 𝑦 ∧ 𝜑)) | ||
Theorem | spsbe 2046 | A specialization theorem. (Contributed by NM, 29-Jun-1993.) (Proof shortened by Wolf Lammen, 3-May-2018.) |
⊢ ([𝑦 / 𝑥]𝜑 → ∃𝑥𝜑) | ||
Theorem | sbequ8 2047 | Elimination of equality from antecedent after substitution. (Contributed by NM, 5-Aug-1993.) Reduce dependencies on axioms. (Revised by Wolf Lammen, 28-Jul-2018.) |
⊢ ([𝑦 / 𝑥]𝜑 ↔ [𝑦 / 𝑥](𝑥 = 𝑦 → 𝜑)) | ||
Theorem | sbimi 2048 | Infer substitution into antecedent and consequent of an implication. (Contributed by NM, 25-Jun-1998.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 → [𝑦 / 𝑥]𝜓) | ||
Theorem | sbbii 2049 | Infer substitution into both sides of a logical equivalence. (Contributed by NM, 14-May-1993.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 ↔ [𝑦 / 𝑥]𝜓) | ||
Axiom | ax-6 2050 |
Axiom of Existence. One of the equality and substitution axioms of
predicate calculus with equality. This axiom tells us is that at least
one thing exists. In this form (not requiring that 𝑥 and 𝑦 be
distinct) it was used in an axiom system of Tarski (see Axiom B7' in
footnote 1 of [KalishMontague] p.
81.) It is equivalent to axiom scheme
C10' in [Megill] p. 448 (p. 16 of the
preprint); the equivalence is
established by axc10 2393 and ax6fromc10 34681. A more convenient form of this
axiom is ax6e 2391, which has additional remarks.
Raph Levien proved the independence of this axiom from the other logical axioms on 12-Apr-2005. See item 16 at http://us.metamath.org/award2003.html. ax-6 2050 can be proved from the weaker version ax6v 2051 requiring that the variables be distinct; see theorem ax6 2392. ax-6 2050 can also be proved from the Axiom of Separation (in the form that we use that axiom, where free variables are not universally quantified). See theorem ax6vsep 4933. Except by ax6v 2051, this axiom should not be referenced directly. Instead, use theorem ax6 2392. (Contributed by NM, 10-Jan-1993.) (New usage is discouraged.) |
⊢ ¬ ∀𝑥 ¬ 𝑥 = 𝑦 | ||
Theorem | ax6v 2051* |
Axiom B7 of [Tarski] p. 75, which requires that
𝑥
and 𝑦 be
distinct. This trivial proof is intended merely to weaken axiom ax-6 2050
by adding a distinct variable restriction ($d). From here on, ax-6 2050
should not be referenced directly by any other proof, so that theorem
ax6 2392 will show that we can recover ax-6 2050
from this weaker version if
it were an axiom (as it is in the case of Tarski).
Note: Introducing 𝑥, 𝑦 as a distinct variable group "out of the blue" with no apparent justification has puzzled some people, but it is perfectly sound. All we are doing is adding an additional prerequisite, similar to adding an unnecessary logical hypothesis, that results in a weakening of the theorem. This means that any future theorem that references ax6v 2051 must have a $d specified for the two variables that get substituted for 𝑥 and 𝑦. The $d does not propagate "backwards" i.e. it does not impose a requirement on ax-6 2050. When possible, use of this theorem rather than ax6 2392 is preferred since its derivation is much shorter and requires fewer axioms. (Contributed by NM, 7-Aug-2015.) |
⊢ ¬ ∀𝑥 ¬ 𝑥 = 𝑦 | ||
Theorem | ax6ev 2052* | At least one individual exists. Weaker version of ax6e 2391. When possible, use of this theorem rather than ax6e 2391 is preferred since its derivation is much shorter and requires fewer axioms. (Contributed by NM, 3-Aug-2017.) |
⊢ ∃𝑥 𝑥 = 𝑦 | ||
Theorem | exiftru 2053 | Rule of existential generalization, similar to universal generalization ax-gen 1867, but valid only if an individual exists. Its proof requires ax-6 2050 but the equality predicate does not occur in its statement. Some fundamental theorems of predicate logic can be proven from ax-gen 1867, ax-4 1882 and this theorem alone, not requiring ax-7 2086 or excessive distinct variable conditions. (Contributed by Wolf Lammen, 12-Nov-2017.) (Proof shortened by Wolf Lammen, 9-Dec-2017.) |
⊢ 𝜑 ⇒ ⊢ ∃𝑥𝜑 | ||
Theorem | 19.2 2054 | Theorem 19.2 of [Margaris] p. 89. This corresponds to the axiom (D) of modal logic. Note: This proof is very different from Margaris' because we only have Tarski's FOL axiom schemes available at this point. See the later 19.2g 2201 for a more conventional proof of a more general result, which uses additional axioms. The reverse implication is the defining property of effective non-freeness (see df-nf 1855). (Contributed by NM, 2-Aug-2017.) Remove dependency on ax-7 2086. (Revised by Wolf Lammen, 4-Dec-2017.) |
⊢ (∀𝑥𝜑 → ∃𝑥𝜑) | ||
Theorem | 19.2d 2055 | Deduction associated with 19.2 2054. (Contributed by BJ, 12-May-2019.) |
⊢ (𝜑 → ∀𝑥𝜓) ⇒ ⊢ (𝜑 → ∃𝑥𝜓) | ||
Theorem | 19.8w 2056 | Weak version of 19.8a 2195 and instance of 19.2d 2055. (Contributed by NM, 1-Aug-2017.) (Proof shortened by Wolf Lammen, 4-Dec-2017.) (Revised by BJ, 31-Mar-2021.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ (𝜑 → ∃𝑥𝜑) | ||
Theorem | 19.8v 2057* | Version of 19.8a 2195 with a dv condition, requiring fewer axioms. Converse of ax5e 1986. (Contributed by BJ, 12-Mar-2020.) |
⊢ (𝜑 → ∃𝑥𝜑) | ||
Theorem | 19.9v 2058* | Version of 19.9 2215 with a dv condition, requiring fewer axioms. Any formula can be existentially quantified using a variable which it does not contain. See also 19.3v 2059. (Contributed by NM, 28-May-1995.) Remove dependency on ax-7 2086. (Revised by Wolf Lammen, 4-Dec-2017.) |
⊢ (∃𝑥𝜑 ↔ 𝜑) | ||
Theorem | 19.3v 2059* | Version of 19.3 2212 with a dv condition, requiring fewer axioms. Any formula can be universally quantified using a variable which it does not contain. See also 19.9v 2058. (Contributed by Anthony Hart, 13-Sep-2011.) Remove dependency on ax-7 2086. (Revised by Wolf Lammen, 4-Dec-2017.) |
⊢ (∀𝑥𝜑 ↔ 𝜑) | ||
Theorem | spvw 2060* | Version of sp 2196 when 𝑥 does not occur in 𝜑. Converse of ax-5 1984. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 10-Apr-2017.) (Proof shortened by Wolf Lammen, 4-Dec-2017.) |
⊢ (∀𝑥𝜑 → 𝜑) | ||
Theorem | 19.39 2061 | Theorem 19.39 of [Margaris] p. 90. (Contributed by NM, 12-Mar-1993.) |
⊢ ((∃𝑥𝜑 → ∃𝑥𝜓) → ∃𝑥(𝜑 → 𝜓)) | ||
Theorem | 19.24 2062 | Theorem 19.24 of [Margaris] p. 90. (Contributed by NM, 12-Mar-1993.) |
⊢ ((∀𝑥𝜑 → ∀𝑥𝜓) → ∃𝑥(𝜑 → 𝜓)) | ||
Theorem | 19.34 2063 | Theorem 19.34 of [Margaris] p. 90. (Contributed by NM, 12-Mar-1993.) |
⊢ ((∀𝑥𝜑 ∨ ∃𝑥𝜓) → ∃𝑥(𝜑 ∨ 𝜓)) | ||
Theorem | 19.23vOLD 2064* | Obsolete version of 19.23v 2016 as of 15-Apr-2022. (Contributed by NM, 28-Jun-1998.) Reduce dependencies on axioms. (Revised by Wolf Lammen, 11-Jan-2020.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (∀𝑥(𝜑 → 𝜓) ↔ (∃𝑥𝜑 → 𝜓)) | ||
Theorem | 19.36v 2065* | Version of 19.36 2241 with a dv condition instead of a non-freeness hypothesis. (Contributed by NM, 18-Aug-1993.) Reduce dependencies on axioms. (Revised by Wolf Lammen, 17-Jan-2020.) |
⊢ (∃𝑥(𝜑 → 𝜓) ↔ (∀𝑥𝜑 → 𝜓)) | ||
Theorem | 19.36ivOLD 2066* | Obsolete version of 19.36iv 2019 as of 15-Apr-2022. (Contributed by NM, 5-Aug-1993.) Reduce dependencies on axioms. (Revised by Wolf Lammen, 17-Jan-2020.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ ∃𝑥(𝜑 → 𝜓) ⇒ ⊢ (∀𝑥𝜑 → 𝜓) | ||
Theorem | pm11.53v 2067* | Version of pm11.53 2320 with a dv condition, requiring fewer axioms. (Contributed by BJ, 7-Mar-2020.) |
⊢ (∀𝑥∀𝑦(𝜑 → 𝜓) ↔ (∃𝑥𝜑 → ∀𝑦𝜓)) | ||
Theorem | 19.12vvv 2068* | Version of 19.12vv 2321 with a dv condition, requiring fewer axioms. See also 19.12 2305. (Contributed by BJ, 18-Mar-2020.) |
⊢ (∃𝑥∀𝑦(𝜑 → 𝜓) ↔ ∀𝑦∃𝑥(𝜑 → 𝜓)) | ||
Theorem | 19.27v 2069* | Version of 19.27 2238 with a dv condition, requiring fewer axioms. (Contributed by NM, 3-Jun-2004.) |
⊢ (∀𝑥(𝜑 ∧ 𝜓) ↔ (∀𝑥𝜑 ∧ 𝜓)) | ||
Theorem | 19.28v 2070* | Version of 19.28 2239 with a dv condition, requiring fewer axioms. (Contributed by NM, 25-Mar-2004.) |
⊢ (∀𝑥(𝜑 ∧ 𝜓) ↔ (𝜑 ∧ ∀𝑥𝜓)) | ||
Theorem | 19.37v 2071* | Version of 19.37 2243 with a dv condition, requiring fewer axioms. (Contributed by NM, 21-Jun-1993.) |
⊢ (∃𝑥(𝜑 → 𝜓) ↔ (𝜑 → ∃𝑥𝜓)) | ||
Theorem | 19.37ivOLD 2072* | Obsolete version of 19.37iv 2021 as of 15-Apr-2022. (Contributed by NM, 5-Aug-1993.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ ∃𝑥(𝜑 → 𝜓) ⇒ ⊢ (𝜑 → ∃𝑥𝜓) | ||
Theorem | 19.44v 2073* | Version of 19.44 2249 with a dv condition, requiring fewer axioms. (Contributed by NM, 12-Mar-1993.) |
⊢ (∃𝑥(𝜑 ∨ 𝜓) ↔ (∃𝑥𝜑 ∨ 𝜓)) | ||
Theorem | 19.45v 2074* | Version of 19.45 2250 with a dv condition, requiring fewer axioms. (Contributed by NM, 12-Mar-1993.) |
⊢ (∃𝑥(𝜑 ∨ 𝜓) ↔ (𝜑 ∨ ∃𝑥𝜓)) | ||
Theorem | 19.41vOLD 2075* | Obsolete version of 19.41v 2022 as of 15-Apr-2022. (Contributed by NM, 21-Jun-1993.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (∃𝑥(𝜑 ∧ 𝜓) ↔ (∃𝑥𝜑 ∧ 𝜓)) | ||
Theorem | spimeh 2076* | Existential introduction, using implicit substitution. Compare Lemma 14 of [Tarski] p. 70. (Contributed by NM, 7-Aug-1994.) (Proof shortened by Wolf Lammen, 10-Dec-2017.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (𝜑 → ∃𝑥𝜓) | ||
Theorem | spimw 2077* | Specialization. Lemma 8 of [KalishMontague] p. 87. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 19-Apr-2017.) (Proof shortened by Wolf Lammen, 7-Aug-2017.) |
⊢ (¬ 𝜓 → ∀𝑥 ¬ 𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (∀𝑥𝜑 → 𝜓) | ||
Theorem | spimvw 2078* | Specialization. Lemma 8 of [KalishMontague] p. 87. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 9-Apr-2017.) |
⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (∀𝑥𝜑 → 𝜓) | ||
Theorem | spnfw 2079 | Weak version of sp 2196. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 1-Aug-2017.) (Proof shortened by Wolf Lammen, 13-Aug-2017.) |
⊢ (¬ 𝜑 → ∀𝑥 ¬ 𝜑) ⇒ ⊢ (∀𝑥𝜑 → 𝜑) | ||
Theorem | spfalw 2080 | Version of sp 2196 when 𝜑 is false. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 23-Apr-2017.) (Proof shortened by Wolf Lammen, 25-Dec-2017.) |
⊢ ¬ 𝜑 ⇒ ⊢ (∀𝑥𝜑 → 𝜑) | ||
Theorem | equs4v 2081* | Version of equs4 2431 with a dv condition, which requires fewer axioms. (Contributed by BJ, 31-May-2019.) |
⊢ (∀𝑥(𝑥 = 𝑦 → 𝜑) → ∃𝑥(𝑥 = 𝑦 ∧ 𝜑)) | ||
Theorem | equsalvw 2082* | Version of equsalv 2251 with a dv condition, and of equsal 2432 with two dv conditions, which requires fewer axioms. See also the dual form equsexvw 2083. (Contributed by BJ, 31-May-2019.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥(𝑥 = 𝑦 → 𝜑) ↔ 𝜓) | ||
Theorem | equsexvw 2083* | Version of equsexv 2252 with a dv condition, and of equsex 2433 with two dv conditions, which requires fewer axioms. See also the dual form equsalvw 2082. (Contributed by BJ, 31-May-2019.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑥(𝑥 = 𝑦 ∧ 𝜑) ↔ 𝜓) | ||
Theorem | cbvaliw 2084* | Change bound variable. Uses only Tarski's FOL axiom schemes. Part of Lemma 7 of [KalishMontague] p. 86. (Contributed by NM, 19-Apr-2017.) |
⊢ (∀𝑥𝜑 → ∀𝑦∀𝑥𝜑) & ⊢ (¬ 𝜓 → ∀𝑥 ¬ 𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (∀𝑥𝜑 → ∀𝑦𝜓) | ||
Theorem | cbvalivw 2085* | Change bound variable. Uses only Tarski's FOL axiom schemes. Part of Lemma 7 of [KalishMontague] p. 86. (Contributed by NM, 9-Apr-2017.) |
⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (∀𝑥𝜑 → ∀𝑦𝜓) | ||
Axiom | ax-7 2086 |
Axiom of Equality. One of the equality and substitution axioms of
predicate calculus with equality. It states that equality is a
right-Euclidean binary relation (this is similar, but not identical, to
being transitive, which is proved as equtr 2099). This axiom scheme is a
sub-scheme of Axiom Scheme B8 of system S2 of [Tarski], p. 75, whose
general form cannot be represented with our notation. Also appears as
Axiom C7 of [Monk2] p. 105 and Axiom Scheme
C8' in [Megill] p. 448 (p. 16
of the preprint).
The equality symbol was invented in 1557 by Robert Recorde. He chose a pair of parallel lines of the same length because "noe .2. thynges, can be moare equalle." We prove in ax7 2094 that this axiom can be recovered from its weakened version ax7v 2087 where 𝑥 and 𝑦 are assumed to be disjoint variables. In particular, the only theorem referencing ax-7 2086 should be ax7v 2087. See the comment of ax7v 2087 for more details on these matters. (Contributed by NM, 10-Jan-1993.) (Revised by BJ, 7-Dec-2020.) Use ax7 2094 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝑥 = 𝑧 → 𝑦 = 𝑧)) | ||
Theorem | ax7v 2087* |
Weakened version of ax-7 2086, with a dv condition on 𝑥, 𝑦. This
should be the only proof referencing ax-7 2086,
and it should be
referenced only by its two weakened versions ax7v1 2088 and ax7v2 2089, from
which ax-7 2086 is then rederived as ax7 2094,
which shows that either ax7v 2087
or the conjunction of ax7v1 2088 and ax7v2 2089 is sufficient.
In ax7v 2087, it is still allowed to substitute the same variable for 𝑥 and 𝑧, or the same variable for 𝑦 and 𝑧. Therefore, ax7v 2087 "bundles" (a term coined by Raph Levien) its "principal instance" (𝑥 = 𝑦 → (𝑥 = 𝑧 → 𝑦 = 𝑧)) with 𝑥, 𝑦, 𝑧 distinct, and its "degenerate instances" (𝑥 = 𝑦 → (𝑥 = 𝑥 → 𝑦 = 𝑥)) and (𝑥 = 𝑦 → (𝑥 = 𝑦 → 𝑦 = 𝑦)) with 𝑥, 𝑦 distinct. These degenerate instances are for instance used in the proofs of equcomiv 2092 and equid 2090 respectively. (Contributed by BJ, 7-Dec-2020.) Use ax7 2094 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝑥 = 𝑧 → 𝑦 = 𝑧)) | ||
Theorem | ax7v1 2088* | First of two weakened versions of ax7v 2087, with an extra dv condition on 𝑥, 𝑧, see comments there. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑥 = 𝑧 → 𝑦 = 𝑧)) | ||
Theorem | ax7v2 2089* | Second of two weakened versions of ax7v 2087, with an extra dv condition on 𝑦, 𝑧, see comments there. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑥 = 𝑧 → 𝑦 = 𝑧)) | ||
Theorem | equid 2090 | Identity law for equality. Lemma 2 of [KalishMontague] p. 85. See also Lemma 6 of [Tarski] p. 68. (Contributed by NM, 1-Apr-2005.) (Revised by NM, 9-Apr-2017.) (Proof shortened by Wolf Lammen, 5-Feb-2018.) (Proof shortened by Wolf Lammen, 22-Aug-2020.) |
⊢ 𝑥 = 𝑥 | ||
Theorem | nfequid 2091 | 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 | equcomiv 2092* | Weaker form of equcomi 2095 with a dv condition on 𝑥, 𝑦. This is an intermediate step and equcomi 2095 is fully recovered later. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → 𝑦 = 𝑥) | ||
Theorem | ax6evr 2093* | A commuted form of ax6ev 2052. (Contributed by BJ, 7-Dec-2020.) |
⊢ ∃𝑥 𝑦 = 𝑥 | ||
Theorem | ax7 2094 |
Proof of ax-7 2086 from ax7v1 2088 and ax7v2 2089, proving sufficiency of the
conjunction of the latter two weakened versions of ax7v 2087,
which is
itself a weakened version of ax-7 2086.
Note that the weakened version of ax-7 2086 obtained by adding a dv condition on 𝑥, 𝑧 (resp. on 𝑦, 𝑧) does not permit, together with the other axioms, to prove reflexivity (resp. symmetry). (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑥 = 𝑧 → 𝑦 = 𝑧)) | ||
Theorem | equcomi 2095 | Commutative law for equality. Equality is a symmetric relation. Lemma 3 of [KalishMontague] p. 85. See also Lemma 7 of [Tarski] p. 69. (Contributed by NM, 10-Jan-1993.) (Revised by NM, 9-Apr-2017.) |
⊢ (𝑥 = 𝑦 → 𝑦 = 𝑥) | ||
Theorem | equcom 2096 | Commutative law for equality. Equality is a symmetric relation. (Contributed by NM, 20-Aug-1993.) |
⊢ (𝑥 = 𝑦 ↔ 𝑦 = 𝑥) | ||
Theorem | equcomd 2097 | Deduction form of equcom 2096, symmetry of equality. For the versions for classes, see eqcom 2763 and eqcomd 2762. (Contributed by BJ, 6-Oct-2019.) |
⊢ (𝜑 → 𝑥 = 𝑦) ⇒ ⊢ (𝜑 → 𝑦 = 𝑥) | ||
Theorem | equcoms 2098 | An inference commuting equality in antecedent. Used to eliminate the need for a syllogism. (Contributed by NM, 10-Jan-1993.) |
⊢ (𝑥 = 𝑦 → 𝜑) ⇒ ⊢ (𝑦 = 𝑥 → 𝜑) | ||
Theorem | equtr 2099 | A transitive law for equality. (Contributed by NM, 23-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝑦 = 𝑧 → 𝑥 = 𝑧)) | ||
Theorem | equtrr 2100 | A transitive law for equality. Lemma L17 in [Megill] p. 446 (p. 14 of the preprint). (Contributed by NM, 23-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝑧 = 𝑥 → 𝑧 = 𝑦)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |