Home | Metamath
Proof Explorer Theorem List (p. 22 of 464) | < 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: | Metamath Proof Explorer
(1-29181) |
Hilbert Space Explorer
(29182-30704) |
Users' Mathboxes
(30705-46395) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | sbievw2 2101* | sbievw 2097 applied twice, avoiding a DV condition on 𝑥, 𝑦. Based on proofs by Wolf Lammen. (Contributed by Steven Nguyen, 29-Jul-2023.) |
⊢ (𝑥 = 𝑤 → (𝜑 ↔ 𝜒)) & ⊢ (𝑤 = 𝑦 → (𝜒 ↔ 𝜓)) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 ↔ 𝜓) | ||
Theorem | sbco2vv 2102* | A composition law for substitution. Version of sbco2 2515 with disjoint variable conditions and fewer axioms. (Contributed by NM, 30-Jun-1994.) (Revised by BJ, 22-Dec-2020.) (Proof shortened by Wolf Lammen, 29-Apr-2023.) |
⊢ ([𝑦 / 𝑧][𝑧 / 𝑥]𝜑 ↔ [𝑦 / 𝑥]𝜑) | ||
Theorem | equsb3 2103* | Substitution in an equality. (Contributed by Raph Levien and FL, 4-Dec-2005.) Reduce axiom usage. (Revised by Wolf Lammen, 23-Jul-2023.) |
⊢ ([𝑦 / 𝑥]𝑥 = 𝑧 ↔ 𝑦 = 𝑧) | ||
Theorem | equsb3r 2104* | Substitution applied to the atomic wff with equality. Variant of equsb3 2103. (Contributed by AV, 29-Jul-2023.) (Proof shortened by Wolf Lammen, 2-Sep-2023.) |
⊢ ([𝑦 / 𝑥]𝑧 = 𝑥 ↔ 𝑧 = 𝑦) | ||
Theorem | equsb1v 2105* | Substitution applied to an atomic wff. Version of equsb1 2495 with a disjoint variable condition, which neither requires ax-12 2173 nor ax-13 2372. (Contributed by NM, 10-May-1993.) (Revised by BJ, 11-Sep-2019.) Remove dependencies on axioms. (Revised by Wolf Lammen, 30-May-2023.) (Proof shortened by Steven Nguyen, 19-Jun-2023.) Revise df-sb 2069. (Revised by Steven Nguyen, 11-Jul-2023.) (Proof shortened by Steven Nguyen, 22-Jul-2023.) |
⊢ [𝑦 / 𝑥]𝑥 = 𝑦 | ||
Theorem | nsb 2106 | Any substitution in an always false formula is false. (Contributed by Steven Nguyen, 3-May-2023.) |
⊢ (∀𝑥 ¬ 𝜑 → ¬ [𝑡 / 𝑥]𝜑) | ||
Theorem | sbn1 2107 | One direction of sbn 2280, using fewer axioms. Compare 19.2 1981. (Contributed by Steven Nguyen, 18-Aug-2023.) |
⊢ ([𝑡 / 𝑥] ¬ 𝜑 → ¬ [𝑡 / 𝑥]𝜑) | ||
Syntax | wcel 2108 |
Extend wff definition to include the membership connective between
classes.
For a general discussion of the theory of classes, see mmset.html#class. The purpose of introducing wff 𝐴 ∈ 𝐵 here is to allow us to prove the wel 2109 of predicate calculus in terms of the wcel 2108 of set theory, so that we do not overload the ∈ connective with two syntax definitions. This is done to prevent ambiguity that would complicate some Metamath parsers. The class variables 𝐴 and 𝐵 are introduced temporarily for the purpose of this definition but otherwise not used in predicate calculus. See df-clab 2716 for more information on the set theory usage of wcel 2108. |
wff 𝐴 ∈ 𝐵 | ||
Theorem | wel 2109 |
Extend wff definition to include atomic formulas with the membership
predicate. This is read either "𝑥 is an element of 𝑦",
or "𝑥 is a member of 𝑦", or "𝑥 belongs
to 𝑦",
or "𝑦 contains 𝑥". Note: The
phrase "𝑦 includes
𝑥 " means "𝑥 is a
subset of 𝑦"; to use it also for
𝑥
∈ 𝑦, as some
authors occasionally do, is poor form and causes
confusion, according to George Boolos (1992 lecture at MIT).
This syntactic construction introduces a binary non-logical predicate symbol ∈ (stylized lowercase epsilon) into our predicate calculus. We will eventually use it for the membership predicate of set theory, but that is irrelevant at this point: the predicate calculus axioms for ∈ apply to any arbitrary binary predicate symbol. "Non-logical" means that the predicate is presumed to have additional properties beyond the realm of predicate calculus, although these additional properties are not specified by predicate calculus itself but rather by the axioms of a theory (in our case set theory) added to predicate calculus. "Binary" means that the predicate has two arguments. Instead of introducing wel 2109 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 wcel 2108. This lets us avoid overloading the ∈ connective, thus preventing ambiguity that would complicate certain Metamath parsers. However, logically wel 2109 is considered to be a primitive syntax, even though here it is artificially "derived" from wcel 2108. Note: To see the proof steps of this syntax proof, type "MM> SHOW PROOF wel / ALL" in the Metamath program. (Contributed by NM, 24-Jan-2006.) |
wff 𝑥 ∈ 𝑦 | ||
Axiom | ax-8 2110 |
Axiom of Left Equality for Binary Predicate. One of the equality and
substitution axioms for a non-logical predicate in our predicate calculus
with equality. It substitutes equal variables into the left-hand side of
an arbitrary binary predicate ∈, which we
will use for the set
membership relation when set theory is introduced. 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 scheme C12' in [Megill] p. 448 (p.
16 of the preprint).
"Non-logical" means that the predicate is not a primitive of
predicate
calculus proper but instead is an extension to it. "Binary"
means that
the predicate has two arguments. In a system of predicate calculus with
equality, like ours, equality is not usually considered to be a
non-logical predicate. In systems of predicate calculus without equality,
it typically would be.
We prove in ax8 2114 that this axiom can be recovered from its weakened version ax8v 2111 where 𝑥 and 𝑦 are assumed to be disjoint variables. In particular, the only theorem referencing ax-8 2110 should be ax8v 2111. See the comment of ax8v 2111 for more details on these matters. (Contributed by NM, 30-Jun-1993.) (Revised by BJ, 7-Dec-2020.) Use ax8 2114 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Theorem | ax8v 2111* | Weakened version of ax-8 2110, with a disjoint variable condition on 𝑥, 𝑦. This should be the only proof referencing ax-8 2110, and it should be referenced only by its two weakened versions ax8v1 2112 and ax8v2 2113, from which ax-8 2110 is then rederived as ax8 2114, which shows that either ax8v 2111 or the conjunction of ax8v1 2112 and ax8v2 2113 is sufficient. (Contributed by BJ, 7-Dec-2020.) Use ax8 2114 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Theorem | ax8v1 2112* | First of two weakened versions of ax8v 2111, with an extra disjoint variable condition on 𝑥, 𝑧, see comments there. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Theorem | ax8v2 2113* | Second of two weakened versions of ax8v 2111, with an extra disjoint variable condition on 𝑦, 𝑧 see comments there. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Theorem | ax8 2114 | Proof of ax-8 2110 from ax8v1 2112 and ax8v2 2113, proving sufficiency of the conjunction of the latter two weakened versions of ax8v 2111, which is itself a weakened version of ax-8 2110. (Contributed by BJ, 7-Dec-2020.) (Proof shortened by Wolf Lammen, 11-Apr-2021.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Theorem | elequ1 2115 | An identity law for the non-logical predicate. (Contributed by NM, 30-Jun-1993.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 ↔ 𝑦 ∈ 𝑧)) | ||
Theorem | elsb1 2116* | Substitution for the first argument of the non-logical predicate in an atomic formula. See elsb2 2125 for substitution for the second argument. (Contributed by NM, 7-Nov-2006.) (Proof shortened by Andrew Salmon, 14-Jun-2011.) Reduce axiom usage. (Revised by Wolf Lammen, 24-Jul-2023.) |
⊢ ([𝑦 / 𝑥]𝑥 ∈ 𝑧 ↔ 𝑦 ∈ 𝑧) | ||
Theorem | cleljust 2117* | When the class variables in Definition df-clel 2817 are replaced with setvar variables, this theorem of predicate calculus is the result. This theorem provides part of the justification for the consistency of that definition, which "overloads" the setvar variables in wel 2109 with the class variables in wcel 2108. (Contributed by NM, 28-Jan-2004.) Revised to use equsexvw 2009 in order to remove dependencies on ax-10 2139, ax-12 2173, ax-13 2372. Note that there is no disjoint variable condition on 𝑥, 𝑦, that is, on the variables of the left-hand side, as should be the case for definitions. (Revised by BJ, 29-Dec-2020.) |
⊢ (𝑥 ∈ 𝑦 ↔ ∃𝑧(𝑧 = 𝑥 ∧ 𝑧 ∈ 𝑦)) | ||
Axiom | ax-9 2118 |
Axiom of Right Equality for Binary Predicate. One of the equality and
substitution axioms for a non-logical predicate in our predicate calculus
with equality. It substitutes equal variables into the right-hand side of
an arbitrary binary predicate ∈, which we
will use for the set
membership relation when set theory is introduced. 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 scheme C13' in [Megill] p. 448 (p.
16 of the preprint).
We prove in ax9 2122 that this axiom can be recovered from its weakened version ax9v 2119 where 𝑥 and 𝑦 are assumed to be disjoint variables. In particular, the only theorem referencing ax-9 2118 should be ax9v 2119. See the comment of ax9v 2119 for more details on these matters. (Contributed by NM, 21-Jun-1993.) (Revised by BJ, 7-Dec-2020.) Use ax9 2122 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | ax9v 2119* | Weakened version of ax-9 2118, with a disjoint variable condition on 𝑥, 𝑦. This should be the only proof referencing ax-9 2118, and it should be referenced only by its two weakened versions ax9v1 2120 and ax9v2 2121, from which ax-9 2118 is then rederived as ax9 2122, which shows that either ax9v 2119 or the conjunction of ax9v1 2120 and ax9v2 2121 is sufficient. (Contributed by BJ, 7-Dec-2020.) Use ax9 2122 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | ax9v1 2120* | First of two weakened versions of ax9v 2119, with an extra disjoint variable condition on 𝑥, 𝑧, see comments there. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | ax9v2 2121* | Second of two weakened versions of ax9v 2119, with an extra disjoint variable condition on 𝑦, 𝑧 see comments there. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | ax9 2122 | Proof of ax-9 2118 from ax9v1 2120 and ax9v2 2121, proving sufficiency of the conjunction of the latter two weakened versions of ax9v 2119, which is itself a weakened version of ax-9 2118. (Contributed by BJ, 7-Dec-2020.) (Proof shortened by Wolf Lammen, 11-Apr-2021.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | elequ2 2123 | An identity law for the non-logical predicate. (Contributed by NM, 21-Jun-1993.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 ↔ 𝑧 ∈ 𝑦)) | ||
Theorem | elequ2g 2124* | A form of elequ2 2123 with a universal quantifier. Its converse is the axiom of extensionality ax-ext 2709. (Contributed by BJ, 3-Oct-2019.) |
⊢ (𝑥 = 𝑦 → ∀𝑧(𝑧 ∈ 𝑥 ↔ 𝑧 ∈ 𝑦)) | ||
Theorem | elsb2 2125* | Substitution for the second argument of the non-logical predicate in an atomic formula. See elsb1 2116 for substitution for the first argument. (Contributed by Rodolfo Medina, 3-Apr-2010.) (Proof shortened by Andrew Salmon, 14-Jun-2011.) Reduce axiom usage. (Revised by Wolf Lammen, 24-Jul-2023.) |
⊢ ([𝑦 / 𝑥]𝑧 ∈ 𝑥 ↔ 𝑧 ∈ 𝑦) | ||
The original axiom schemes of Tarski's predicate calculus are ax-4 1813, ax-5 1914, ax6v 1973, ax-7 2012, ax-8 2110, and ax-9 2118, together with rule ax-gen 1799. See mmset.html#compare 1799. They are given as axiom schemes B4 through B8 in [KalishMontague] p. 81. These are shown to be logically complete by Theorem 1 of [KalishMontague] p. 85. The axiom system of set.mm includes the auxiliary axiom schemes ax-10 2139, ax-11 2156, ax-12 2173, and ax-13 2372, which are not part of Tarski's axiom schemes. Each object-language instance of them is provable from Tarski's axioms, so they are logically redundant. However, they are conjectured not to be provable directly as schemes from Tarski's axiom schemes using only Metamath's direct substitution rule. They are used to make our system "scheme complete", i.e., able to prove directly all possible schemes with wff and setvar variables, bundled or not, whose object-language instances are valid. (ax-12 2173 has been proved to be required; see https://us.metamath.org/award2003.html#9a 2173. Metalogical independence of the other three are open problems.) (There are additional predicate calculus axiom schemes included in set.mm such as ax-c5 36824, but they can all be proved as theorems from the above.) Terminology: Two setvar (individual) metavariables are "bundled" in an axiom or theorem scheme when there is no distinct variable constraint ($d) imposed on them. (The term "bundled" is due to Raph Levien.) For example, the 𝑥 and 𝑦 in ax-6 1972 are bundled, but they are not in ax6v 1973. We also say that a scheme is bundled when it has at least one pair of bundled setvar variables. If distinct variable conditions are added to all setvar variable pairs in a bundled scheme, we call that the "principal" instance of the bundled scheme. For example, ax6v 1973 is the principal instance of ax-6 1972. Whenever a common variable is substituted for two or more bundled variables in an axiom or theorem scheme, we call the substitution instance "degenerate". For example, the instance ¬ ∀𝑥¬ 𝑥 = 𝑥 of ax-6 1972 is degenerate. An advantage of bundling is ease of use since there are fewer distinct variable restrictions ($d) to be concerned with, and theorems are more general. There may be some economy in being able to prove facts about principal and degenerate instances simultaneously. A disadvantage is that bundling may present difficulties in translations to other proof languages, which typically lack the concept (in part because their variables often represent the variables of the object language rather than metavariables ranging over them). Because Tarski's axiom schemes are logically complete, they can be used to prove any object-language instance of ax-10 2139, ax-11 2156, ax-12 2173, and ax-13 2372. "Translating" this to Metamath, it means that Tarski's axioms can prove any substitution instance of ax-10 2139, ax-11 2156, ax-12 2173, or ax-13 2372 in which (1) there are no wff metavariables and (2) all setvar variables are mutually distinct i.e. are not bundled. In effect this is mimicking the object language by pretending that each setvar variable is an object-language variable. (There may also be specific instances with wff metavariables and/or bundling that are directly provable from Tarski's axiom schemes, but it isn't guaranteed. Whether all of them are possible is part of the still open metalogical independence problem for our additional axiom schemes.) It can be useful to see how this can be done, both to show that our additional schemes are valid metatheorems of Tarski's system and to be able to translate object-language instances of our proofs into proofs that would work with a system using only Tarski's original schemes. In addition, it may (or may not) provide insight into the conjectured metalogical independence of our additional schemes. The theorem schemes ax10w 2127, ax11w 2128, ax12w 2131, and ax13w 2134 are derived using only Tarski's axiom schemes, showing that Tarski's schemes can be used to derive all substitution instances of ax-10 2139, ax-11 2156, ax-12 2173, and ax-13 2372 meeting Conditions (1) and (2). (The "w" suffix stands for "weak version".) Each hypothesis of ax10w 2127, ax11w 2128, and ax12w 2131 is of the form (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) where 𝜓 is an auxiliary or "dummy" wff metavariable in which 𝑥 doesn't occur. We can show by induction on formula length that the hypotheses can be eliminated in all cases meeting Conditions (1) and (2). The example ax12wdemo 2133 illustrates the techniques (equality theorems and bound variable renaming) used to achieve this. We also show the degenerate instances for axioms with bundled variables in ax11dgen 2129, ax12dgen 2132, ax13dgen1 2135, ax13dgen2 2136, ax13dgen3 2137, and ax13dgen4 2138. (Their proofs are trivial, but we include them to be thorough.) Combining the principal and degenerate cases outside of Metamath, we show that the bundled schemes ax-10 2139, ax-11 2156, ax-12 2173, and ax-13 2372 are schemes of Tarski's system, meaning that all object-language instances they generate are theorems of Tarski's system. It is interesting that Tarski used the bundled scheme ax-6 1972 in an older system, so it seems the main purpose of his later ax6v 1973 was just to show that the weaker unbundled form is sufficient rather than an aesthetic objection to bundled free and bound variables. Since we adopt the bundled ax-6 1972 as our official axiom, we show that the degenerate instance holds in ax6dgen 2126. (Recall that in set.mm, the only statement referencing ax-6 1972 is ax6v 1973.) The case of sp 2178 is curious: originally an axiom scheme of Tarski's system, it was proved logically redundant by Lemma 9 of [KalishMontague] p. 86. However, the proof is by induction on formula length, and the scheme form ∀𝑥𝜑 → 𝜑 apparently cannot be proved directly from Tarski's other axiom schemes. The best we can do seems to be spw 2038, again requiring substitution instances of 𝜑 that meet Conditions (1) and (2) above. Note that our direct proof sp 2178 requires ax-12 2173, which is not part of Tarski's system. | ||
Theorem | ax6dgen 2126 | Tarski's system uses the weaker ax6v 1973 instead of the bundled ax-6 1972, so here we show that the degenerate case of ax-6 1972 can be derived. Even though ax-6 1972 is in the list of axioms used, recall that in set.mm, the only statement referencing ax-6 1972 is ax6v 1973. We later rederive from ax6v 1973 the bundled form as ax6 2384 with the help of the auxiliary axiom schemes. (Contributed by NM, 23-Apr-2017.) |
⊢ ¬ ∀𝑥 ¬ 𝑥 = 𝑥 | ||
Theorem | ax10w 2127* | Weak version of ax-10 2139 from which we can prove any ax-10 2139 instance not involving wff variables or bundling. Uses only Tarski's FOL axiom schemes. It is an alias of hbn1w 2050 introduced for labeling consistency. (Contributed by NM, 9-Apr-2017.) Use hbn1w 2050 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (¬ ∀𝑥𝜑 → ∀𝑥 ¬ ∀𝑥𝜑) | ||
Theorem | ax11w 2128* | Weak version of ax-11 2156 from which we can prove any ax-11 2156 instance not involving wff variables or bundling. Uses only Tarski's FOL axiom schemes. Unlike ax-11 2156, this theorem requires that 𝑥 and 𝑦 be distinct i.e. are not bundled. It is an alias of alcomiw 2047 introduced for labeling consistency. (Contributed by NM, 10-Apr-2017.) Use alcomiw 2047 instead. (New usage is discouraged.) |
⊢ (𝑦 = 𝑧 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥∀𝑦𝜑 → ∀𝑦∀𝑥𝜑) | ||
Theorem | ax11dgen 2129 | Degenerate instance of ax-11 2156 where bundled variables 𝑥 and 𝑦 have a common substitution. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) |
⊢ (∀𝑥∀𝑥𝜑 → ∀𝑥∀𝑥𝜑) | ||
Theorem | ax12wlem 2130* | Lemma for weak version of ax-12 2173. Uses only Tarski's FOL axiom schemes. In some cases, this lemma may lead to shorter proofs than ax12w 2131. (Contributed by NM, 10-Apr-2017.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | ax12w 2131* | Weak version of ax-12 2173 from which we can prove any ax-12 2173 instance not involving wff variables or bundling. Uses only Tarski's FOL axiom schemes. An instance of the first hypothesis will normally require that 𝑥 and 𝑦 be distinct (unless 𝑥 does not occur in 𝜑). For an example of how the hypotheses can be eliminated when we substitute an expression without wff variables for 𝜑, see ax12wdemo 2133. (Contributed by NM, 10-Apr-2017.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝑦 = 𝑧 → (𝜑 ↔ 𝜒)) ⇒ ⊢ (𝑥 = 𝑦 → (∀𝑦𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | ax12dgen 2132 | Degenerate instance of ax-12 2173 where bundled variables 𝑥 and 𝑦 have a common substitution. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) |
⊢ (𝑥 = 𝑥 → (∀𝑥𝜑 → ∀𝑥(𝑥 = 𝑥 → 𝜑))) | ||
Theorem | ax12wdemo 2133* | Example of an application of ax12w 2131 that results in an instance of ax-12 2173 for a contrived formula with mixed free and bound variables, (𝑥 ∈ 𝑦 ∧ ∀𝑥𝑧 ∈ 𝑥 ∧ ∀𝑦∀𝑧𝑦 ∈ 𝑥), in place of 𝜑. The proof illustrates bound variable renaming with cbvalvw 2040 to obtain fresh variables to avoid distinct variable clashes. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 14-Apr-2017.) |
⊢ (𝑥 = 𝑦 → (∀𝑦(𝑥 ∈ 𝑦 ∧ ∀𝑥 𝑧 ∈ 𝑥 ∧ ∀𝑦∀𝑧 𝑦 ∈ 𝑥) → ∀𝑥(𝑥 = 𝑦 → (𝑥 ∈ 𝑦 ∧ ∀𝑥 𝑧 ∈ 𝑥 ∧ ∀𝑦∀𝑧 𝑦 ∈ 𝑥)))) | ||
Theorem | ax13w 2134* | Weak version (principal instance) of ax-13 2372. (Because 𝑦 and 𝑧 don't need to be distinct, this actually bundles the principal instance and the degenerate instance (¬ 𝑥 = 𝑦 → (𝑦 = 𝑦 → ∀𝑥𝑦 = 𝑦)).) Uses only Tarski's FOL axiom schemes. The proof is trivial but is included to complete the set ax10w 2127, ax11w 2128, and ax12w 2131. (Contributed by NM, 10-Apr-2017.) |
⊢ (¬ 𝑥 = 𝑦 → (𝑦 = 𝑧 → ∀𝑥 𝑦 = 𝑧)) | ||
Theorem | ax13dgen1 2135 | Degenerate instance of ax-13 2372 where bundled variables 𝑥 and 𝑦 have a common substitution. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) |
⊢ (¬ 𝑥 = 𝑥 → (𝑥 = 𝑧 → ∀𝑥 𝑥 = 𝑧)) | ||
Theorem | ax13dgen2 2136 | Degenerate instance of ax-13 2372 where bundled variables 𝑥 and 𝑧 have a common substitution. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) |
⊢ (¬ 𝑥 = 𝑦 → (𝑦 = 𝑥 → ∀𝑥 𝑦 = 𝑥)) | ||
Theorem | ax13dgen3 2137 | Degenerate instance of ax-13 2372 where bundled variables 𝑦 and 𝑧 have a common substitution. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) |
⊢ (¬ 𝑥 = 𝑦 → (𝑦 = 𝑦 → ∀𝑥 𝑦 = 𝑦)) | ||
Theorem | ax13dgen4 2138 | Degenerate instance of ax-13 2372 where bundled variables 𝑥, 𝑦, and 𝑧 have a common substitution. Therefore, also a degenerate instance of ax13dgen1 2135, ax13dgen2 2136, and ax13dgen3 2137. Also an instance of the intuitionistic tautology pm2.21 123. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) Reduce axiom usage. (Revised by Wolf Lammen, 10-Oct-2021.) |
⊢ (¬ 𝑥 = 𝑥 → (𝑥 = 𝑥 → ∀𝑥 𝑥 = 𝑥)) | ||
In this section we introduce four additional schemes ax-10 2139, ax-11 2156, ax-12 2173, and ax-13 2372 that are not part of Tarski's system but can be proved (outside of Metamath) as theorem schemes of Tarski's system. These are needed to give our system the property of "scheme completeness", which means that we can prove (with Metamath) all possible theorem schemes expressible in our language of wff metavariables ranging over object-language wffs, and setvar variables ranging over object-language individual variables. To show that these schemes are valid metatheorems of Tarski's system S2, above we proved from Tarski's system theorems ax10w 2127, ax11w 2128, ax12w 2131, and ax13w 2134, which show that any object-language instance of these schemes (emulated by having no wff metavariables and requiring all setvar variables to be mutually distinct) can be proved using only the schemes in Tarski's system S2. An open problem is to show that these four additional schemes are mutually metalogically independent and metalogically independent from Tarski's. So far, independence of ax-12 2173 from all others has been shown, and independence of Tarski's ax-6 1972 from all others has been shown; see items 9a and 11 on https://us.metamath.org/award2003.html 1972. | ||
Axiom | ax-10 2139 | Axiom of Quantified Negation. Axiom C5-2 of [Monk2] p. 113. This axiom scheme is logically redundant (see ax10w 2127) but is used as an auxiliary axiom scheme to achieve scheme completeness. It means that 𝑥 is not free in ¬ ∀𝑥𝜑. (Contributed by NM, 21-May-2008.) Use its alias hbn1 2140 instead if you must use it. Any theorem in first-order logic (FOL) that contains only set variables that are all mutually distinct, and has no wff variables, can be proved *without* using ax-10 2139 through ax-13 2372, by invoking ax10w 2127 through ax13w 2134. We encourage proving theorems *without* ax-10 2139 through ax-13 2372 and moving them up to the ax-4 1813 through ax-9 2118 section. (New usage is discouraged.) |
⊢ (¬ ∀𝑥𝜑 → ∀𝑥 ¬ ∀𝑥𝜑) | ||
Theorem | hbn1 2140 | Alias for ax-10 2139 to be used instead of it. (Contributed by NM, 24-Jan-1993.) (Proof shortened by Wolf Lammen, 18-Aug-2014.) |
⊢ (¬ ∀𝑥𝜑 → ∀𝑥 ¬ ∀𝑥𝜑) | ||
Theorem | hbe1 2141 | The setvar 𝑥 is not free in ∃𝑥𝜑. Corresponds to the axiom (5) of modal logic (see also modal5 2154). (Contributed by NM, 24-Jan-1993.) |
⊢ (∃𝑥𝜑 → ∀𝑥∃𝑥𝜑) | ||
Theorem | hbe1a 2142 | Dual statement of hbe1 2141. Modified version of axc7e 2316 with a universally quantified consequent. (Contributed by Wolf Lammen, 15-Sep-2021.) |
⊢ (∃𝑥∀𝑥𝜑 → ∀𝑥𝜑) | ||
Theorem | nf5-1 2143 | One direction of nf5 2282 can be proved with a smaller footprint on axiom usage. (Contributed by Wolf Lammen, 16-Sep-2021.) |
⊢ (∀𝑥(𝜑 → ∀𝑥𝜑) → Ⅎ𝑥𝜑) | ||
Theorem | nf5i 2144 | Deduce that 𝑥 is not free in 𝜑 from the definition. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ Ⅎ𝑥𝜑 | ||
Theorem | nf5dh 2145 | Deduce that 𝑥 is not free in 𝜓 in a context. (Contributed by Mario Carneiro, 24-Sep-2016.) df-nf 1788 changed. (Revised by Wolf Lammen, 11-Oct-2021.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
Theorem | nf5dv 2146* | Apply the definition of not-free in a context. (Contributed by Mario Carneiro, 11-Aug-2016.) df-nf 1788 changed. (Revised by Wolf Lammen, 18-Sep-2021.) (Proof shortened by Wolf Lammen, 13-Jul-2022.) |
⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
Theorem | nfnaew 2147* | All variables are effectively bound in a distinct variable specifier. Version of nfnae 2434 with a disjoint variable condition, which does not require ax-13 2372. (Contributed by Mario Carneiro, 11-Aug-2016.) (Revised by Gino Giotto, 10-Jan-2024.) (Proof shortened by Wolf Lammen, 25-Sep-2024.) |
⊢ Ⅎ𝑧 ¬ ∀𝑥 𝑥 = 𝑦 | ||
Theorem | nfnaewOLD 2148* | Obsolete version of nfnaew 2147 as of 25-Sep-2024. (Contributed by Mario Carneiro, 11-Aug-2016.) (Revised by Gino Giotto, 10-Jan-2024.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ Ⅎ𝑧 ¬ ∀𝑥 𝑥 = 𝑦 | ||
Theorem | nfe1 2149 | The setvar 𝑥 is not free in ∃𝑥𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ Ⅎ𝑥∃𝑥𝜑 | ||
Theorem | nfa1 2150 | The setvar 𝑥 is not free in ∀𝑥𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.) df-nf 1788 changed. (Revised by Wolf Lammen, 11-Sep-2021.) Remove dependency on ax-12 2173. (Revised by Wolf Lammen, 12-Oct-2021.) |
⊢ Ⅎ𝑥∀𝑥𝜑 | ||
Theorem | nfna1 2151 | A convenience theorem particularly designed to remove dependencies on ax-11 2156 in conjunction with distinctors. (Contributed by Wolf Lammen, 2-Sep-2018.) |
⊢ Ⅎ𝑥 ¬ ∀𝑥𝜑 | ||
Theorem | nfia1 2152 | Lemma 23 of [Monk2] p. 114. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥(∀𝑥𝜑 → ∀𝑥𝜓) | ||
Theorem | nfnf1 2153 | The setvar 𝑥 is not free in Ⅎ𝑥𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.) Remove dependency on ax-12 2173. (Revised by Wolf Lammen, 12-Oct-2021.) |
⊢ Ⅎ𝑥Ⅎ𝑥𝜑 | ||
Theorem | modal5 2154 | The analogue in our predicate calculus of axiom (5) of modal logic S5. See also hbe1 2141. (Contributed by NM, 5-Oct-2005.) |
⊢ (¬ ∀𝑥 ¬ 𝜑 → ∀𝑥 ¬ ∀𝑥 ¬ 𝜑) | ||
Theorem | nfs1v 2155* | The setvar 𝑥 is not free in [𝑦 / 𝑥]𝜑 when 𝑥 and 𝑦 are distinct. (Contributed by Mario Carneiro, 11-Aug-2016.) Shorten nfs1v 2155 and hbs1 2269 combined. (Revised by Wolf Lammen, 28-Jul-2022.) |
⊢ Ⅎ𝑥[𝑦 / 𝑥]𝜑 | ||
Axiom | ax-11 2156 | Axiom of Quantifier Commutation. This axiom says universal quantifiers can be swapped. Axiom scheme C6' in [Megill] p. 448 (p. 16 of the preprint). Also appears as Lemma 12 of [Monk2] p. 109 and Axiom C5-3 of [Monk2] p. 113. This axiom scheme is logically redundant (see ax11w 2128) but is used as an auxiliary axiom scheme to achieve metalogical completeness. (Contributed by NM, 12-Mar-1993.) |
⊢ (∀𝑥∀𝑦𝜑 → ∀𝑦∀𝑥𝜑) | ||
Theorem | alcoms 2157 | Swap quantifiers in an antecedent. (Contributed by NM, 11-May-1993.) |
⊢ (∀𝑥∀𝑦𝜑 → 𝜓) ⇒ ⊢ (∀𝑦∀𝑥𝜑 → 𝜓) | ||
Theorem | alcom 2158 | Theorem 19.5 of [Margaris] p. 89. (Contributed by NM, 30-Jun-1993.) |
⊢ (∀𝑥∀𝑦𝜑 ↔ ∀𝑦∀𝑥𝜑) | ||
Theorem | alrot3 2159 | Theorem *11.21 in [WhiteheadRussell] p. 160. (Contributed by Andrew Salmon, 24-May-2011.) |
⊢ (∀𝑥∀𝑦∀𝑧𝜑 ↔ ∀𝑦∀𝑧∀𝑥𝜑) | ||
Theorem | alrot4 2160 | Rotate four universal quantifiers twice. (Contributed by NM, 2-Feb-2005.) (Proof shortened by Fan Zheng, 6-Jun-2016.) |
⊢ (∀𝑥∀𝑦∀𝑧∀𝑤𝜑 ↔ ∀𝑧∀𝑤∀𝑥∀𝑦𝜑) | ||
Theorem | sbal 2161* | Move universal quantifier in and out of substitution. (Contributed by NM, 16-May-1993.) (Proof shortened by Wolf Lammen, 29-Sep-2018.) Reduce dependencies on axioms. (Revised by Steven Nguyen, 13-Aug-2023.) |
⊢ ([𝑧 / 𝑦]∀𝑥𝜑 ↔ ∀𝑥[𝑧 / 𝑦]𝜑) | ||
Theorem | sbalv 2162* | Quantify with new variable inside substitution. (Contributed by NM, 18-Aug-1993.) |
⊢ ([𝑦 / 𝑥]𝜑 ↔ 𝜓) ⇒ ⊢ ([𝑦 / 𝑥]∀𝑧𝜑 ↔ ∀𝑧𝜓) | ||
Theorem | sbcom2 2163* | Commutativity law for substitution. Used in proof of Theorem 9.7 of [Megill] p. 449 (p. 16 of the preprint). (Contributed by NM, 27-May-1997.) (Proof shortened by Wolf Lammen, 23-Dec-2022.) |
⊢ ([𝑤 / 𝑧][𝑦 / 𝑥]𝜑 ↔ [𝑦 / 𝑥][𝑤 / 𝑧]𝜑) | ||
Theorem | excom 2164 | Theorem 19.11 of [Margaris] p. 89. (Contributed by NM, 5-Aug-1993.) Remove dependencies on ax-5 1914, ax-6 1972, ax-7 2012, ax-10 2139, ax-12 2173. (Revised by Wolf Lammen, 8-Jan-2018.) (Proof shortened by Wolf Lammen, 22-Aug-2020.) |
⊢ (∃𝑥∃𝑦𝜑 ↔ ∃𝑦∃𝑥𝜑) | ||
Theorem | excomim 2165 | One direction of Theorem 19.11 of [Margaris] p. 89. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 24-Sep-2016.) Remove dependencies on ax-5 1914, ax-6 1972, ax-7 2012, ax-10 2139, ax-12 2173. (Revised by Wolf Lammen, 8-Jan-2018.) |
⊢ (∃𝑥∃𝑦𝜑 → ∃𝑦∃𝑥𝜑) | ||
Theorem | excom13 2166 | Swap 1st and 3rd existential quantifiers. (Contributed by NM, 9-Mar-1995.) |
⊢ (∃𝑥∃𝑦∃𝑧𝜑 ↔ ∃𝑧∃𝑦∃𝑥𝜑) | ||
Theorem | exrot3 2167 | Rotate existential quantifiers. (Contributed by NM, 17-Mar-1995.) |
⊢ (∃𝑥∃𝑦∃𝑧𝜑 ↔ ∃𝑦∃𝑧∃𝑥𝜑) | ||
Theorem | exrot4 2168 | Rotate existential quantifiers twice. (Contributed by NM, 9-Mar-1995.) |
⊢ (∃𝑥∃𝑦∃𝑧∃𝑤𝜑 ↔ ∃𝑧∃𝑤∃𝑥∃𝑦𝜑) | ||
Theorem | hbal 2169 | If 𝑥 is not free in 𝜑, it is not free in ∀𝑦𝜑. (Contributed by NM, 12-Mar-1993.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ (∀𝑦𝜑 → ∀𝑥∀𝑦𝜑) | ||
Theorem | hbald 2170 | Deduction form of bound-variable hypothesis builder hbal 2169. (Contributed by NM, 2-Jan-2002.) |
⊢ (𝜑 → ∀𝑦𝜑) & ⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → (∀𝑦𝜓 → ∀𝑥∀𝑦𝜓)) | ||
Theorem | hbsbw 2171* | If 𝑧 is not free in 𝜑, it is not free in [𝑦 / 𝑥]𝜑 when 𝑦 and 𝑧 are distinct. Version of hbsb 2529 with a disjoint variable condition, which requires fewer axioms. (Contributed by NM, 12-Aug-1993.) (Revised by Gino Giotto, 23-May-2024.) |
⊢ (𝜑 → ∀𝑧𝜑) ⇒ ⊢ ([𝑦 / 𝑥]𝜑 → ∀𝑧[𝑦 / 𝑥]𝜑) | ||
Theorem | nfa2 2172 | Lemma 24 of [Monk2] p. 114. (Contributed by Mario Carneiro, 24-Sep-2016.) Remove dependency on ax-12 2173. (Revised by Wolf Lammen, 18-Oct-2021.) |
⊢ Ⅎ𝑥∀𝑦∀𝑥𝜑 | ||
Axiom | ax-12 2173 |
Axiom of Substitution. One of the 5 equality axioms of predicate
calculus. The final consequent ∀𝑥(𝑥 = 𝑦 → 𝜑) is a way of
expressing "𝑦 substituted for 𝑥 in wff
𝜑
" (cf. sb6 2089). 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.
The original version of this axiom was ax-c15 36830 and was replaced with this shorter ax-12 2173 in Jan. 2007. The old axiom is proved from this one as Theorem axc15 2422. Conversely, this axiom is proved from ax-c15 36830 as Theorem ax12 2423. Juha Arpiainen proved the metalogical independence of this axiom (in the form of the older axiom ax-c15 36830) from the others on 19-Jan-2006. See item 9a at https://us.metamath.org/award2003.html 36830. See ax12v 2174 and ax12v2 2175 for other equivalents of this axiom that (unlike this axiom) have distinct variable restrictions. This axiom scheme is logically redundant (see ax12w 2131) but is used as an auxiliary axiom scheme to achieve scheme completeness. (Contributed by NM, 22-Jan-2007.) (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (∀𝑦𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | ax12v 2174* |
This is essentially Axiom ax-12 2173 weakened by additional restrictions on
variables. Besides axc11r 2366, this theorem should be the only one
referencing ax-12 2173 directly.
Both restrictions on variables have their own value. If for a moment we assume 𝑥 could be set to 𝑦, then, after elimination of the tautology 𝑦 = 𝑦, immediately we have 𝜑 → ∀𝑦𝜑 for all 𝜑 and 𝑦, that is ax-5 1914, a degenerate result. The second restriction is not necessary, but a simplification that makes the following interpretation easier to see. Since 𝜑 textually at most depends on 𝑥, we can look at it at some given 'fixed' 𝑦. This theorem now states that the truth value of 𝜑 will stay constant, as long as we 'vary 𝑥 around 𝑦' only such that 𝑥 = 𝑦 still holds. Or in other words, equality is the finest grained logical expression. If you cannot differ two sets by =, you won't find a whatever sophisticated expression that does. One might wonder how the described variation of 𝑥 is possible at all. Note that Metamath is a text processor that easily sees a difference between text chunks {𝑥 ∣ ¬ 𝑥 = 𝑥} and {𝑦 ∣ ¬ 𝑦 = 𝑦}. Our usual interpretation is to abstract from textual variations of the same set, but we are free to interpret Metamath's formalism differently, and in fact let 𝑥 run through all textual representations of sets. Had we allowed 𝜑 to depend also on 𝑦, this idea is both harder to see, and it is less clear that this extra freedom introduces effects not covered by other axioms. (Contributed by Wolf Lammen, 8-Aug-2020.) |
⊢ (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | ax12v2 2175* | It is possible to remove any restriction on 𝜑 in ax12v 2174. Same as Axiom C8 of [Monk2] p. 105. Use ax12v 2174 instead when sufficient. (Contributed by NM, 5-Aug-1993.) Remove dependencies on ax-10 2139 and ax-13 2372. (Revised by Jim Kingdon, 15-Dec-2017.) (Proof shortened by Wolf Lammen, 8-Dec-2019.) |
⊢ (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | 19.8a 2176 | If a wff is true, it is true for at least one instance. Special case of Theorem 19.8 of [Margaris] p. 89. See 19.8v 1987 for a version with a disjoint variable condition requiring fewer axioms. (Contributed by NM, 9-Jan-1993.) Allow a shortening of sp 2178. (Revised by Wolf Lammen, 13-Jan-2018.) (Proof shortened by Wolf Lammen, 8-Dec-2019.) |
⊢ (𝜑 → ∃𝑥𝜑) | ||
Theorem | 19.8ad 2177 | If a wff is true, it is true for at least one instance. Deduction form of 19.8a 2176. (Contributed by DAW, 13-Feb-2017.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (𝜑 → ∃𝑥𝜓) | ||
Theorem | sp 2178 |
Specialization. A universally quantified wff implies the wff without a
quantifier. Axiom scheme B5 of [Tarski]
p. 67 (under his system S2,
defined in the last paragraph on p. 77). Also appears as Axiom scheme C5'
in [Megill] p. 448 (p. 16 of the
preprint). This corresponds to the axiom
(T) of modal logic.
For the axiom of specialization presented in many logic textbooks, see Theorem stdpc4 2072. This theorem shows that our obsolete axiom ax-c5 36824 can be derived from the others. The proof uses ideas from the proof of Lemma 21 of [Monk2] p. 114. It appears that this scheme cannot be derived directly from Tarski's axioms without auxiliary axiom scheme ax-12 2173. It is thought the best we can do using only Tarski's axioms is spw 2038. Also see spvw 1985 where 𝑥 and 𝜑 are disjoint, using fewer axioms. (Contributed by NM, 21-May-2008.) (Proof shortened by Scott Fenton, 24-Jan-2011.) (Proof shortened by Wolf Lammen, 13-Jan-2018.) |
⊢ (∀𝑥𝜑 → 𝜑) | ||
Theorem | spi 2179 | Inference rule of universal instantiation, or universal specialization. Converse of the inference rule of (universal) generalization ax-gen 1799. Contrary to the rule of generalization, its closed form is valid, see sp 2178. (Contributed by NM, 5-Aug-1993.) |
⊢ ∀𝑥𝜑 ⇒ ⊢ 𝜑 | ||
Theorem | sps 2180 | Generalization of antecedent. (Contributed by NM, 5-Jan-1993.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (∀𝑥𝜑 → 𝜓) | ||
Theorem | 2sp 2181 | A double specialization (see sp 2178). Another double specialization, closer to PM*11.1, is 2stdpc4 2074. (Contributed by BJ, 15-Sep-2018.) |
⊢ (∀𝑥∀𝑦𝜑 → 𝜑) | ||
Theorem | spsd 2182 | Deduction generalizing antecedent. (Contributed by NM, 17-Aug-1994.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 → 𝜒)) | ||
Theorem | 19.2g 2183 | Theorem 19.2 of [Margaris] p. 89, generalized to use two setvar variables. Use 19.2 1981 when sufficient. (Contributed by Mel L. O'Cat, 31-Mar-2008.) |
⊢ (∀𝑥𝜑 → ∃𝑦𝜑) | ||
Theorem | 19.21bi 2184 | Inference form of 19.21 2203 and also deduction form of sp 2178. (Contributed by NM, 26-May-1993.) |
⊢ (𝜑 → ∀𝑥𝜓) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | 19.21bbi 2185 | Inference removing two universal quantifiers. Version of 19.21bi 2184 with two quantifiers. (Contributed by NM, 20-Apr-1994.) |
⊢ (𝜑 → ∀𝑥∀𝑦𝜓) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | 19.23bi 2186 | Inference form of Theorem 19.23 of [Margaris] p. 90, see 19.23 2207. (Contributed by NM, 12-Mar-1993.) |
⊢ (∃𝑥𝜑 → 𝜓) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | nexr 2187 | Inference associated with the contrapositive of 19.8a 2176. (Contributed by Jeff Hankins, 26-Jul-2009.) |
⊢ ¬ ∃𝑥𝜑 ⇒ ⊢ ¬ 𝜑 | ||
Theorem | qexmid 2188 | Quantified excluded middle (see exmid 891). Also known as the drinker paradox (if 𝜑(𝑥) is interpreted as "𝑥 drinks", then this theorem tells that there exists a person such that, if this person drinks, then everyone drinks). Exercise 9.2a of Boolos, p. 111, Computability and Logic. (Contributed by NM, 10-Dec-2000.) |
⊢ ∃𝑥(𝜑 → ∀𝑥𝜑) | ||
Theorem | nf5r 2189 | Consequence of the definition of not-free. (Contributed by Mario Carneiro, 26-Sep-2016.) df-nf 1788 changed. (Revised by Wolf Lammen, 11-Sep-2021.) (Proof shortened by Wolf Lammen, 23-Nov-2023.) |
⊢ (Ⅎ𝑥𝜑 → (𝜑 → ∀𝑥𝜑)) | ||
Theorem | nf5rOLD 2190 | Obsolete version of nfrd 1795 as of 23-Nov-2023. (Contributed by Mario Carneiro, 26-Sep-2016.) df-nf 1788 changed. (Revised by Wolf Lammen, 11-Sep-2021.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (Ⅎ𝑥𝜑 → (𝜑 → ∀𝑥𝜑)) | ||
Theorem | nf5ri 2191 | Consequence of the definition of not-free. (Contributed by Mario Carneiro, 11-Aug-2016.) (Proof shortened by Wolf Lammen, 15-Mar-2023.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (𝜑 → ∀𝑥𝜑) | ||
Theorem | nf5rd 2192 | Consequence of the definition of not-free in a context. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) | ||
Theorem | spimedv 2193* | Deduction version of spimev 2392. Version of spimed 2388 with a disjoint variable condition, which does not require ax-13 2372. See spime 2389 for a non-deduction version. (Contributed by NM, 14-May-1993.) (Revised by BJ, 31-May-2019.) |
⊢ (𝜒 → Ⅎ𝑥𝜑) & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (𝜒 → (𝜑 → ∃𝑥𝜓)) | ||
Theorem | spimefv 2194* | Version of spime 2389 with a disjoint variable condition, which does not require ax-13 2372. (Contributed by BJ, 31-May-2019.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (𝜑 → ∃𝑥𝜓) | ||
Theorem | nfim1 2195 | A closed form of nfim 1900. (Contributed by NM, 2-Jun-1993.) (Revised by Mario Carneiro, 24-Sep-2016.) (Proof shortened by Wolf Lammen, 2-Jan-2018.) df-nf 1788 changed. (Revised by Wolf Lammen, 18-Sep-2021.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ Ⅎ𝑥(𝜑 → 𝜓) | ||
Theorem | nfan1 2196 | A closed form of nfan 1903. (Contributed by Mario Carneiro, 3-Oct-2016.) df-nf 1788 changed. (Revised by Wolf Lammen, 18-Sep-2021.) (Proof shortened by Wolf Lammen, 7-Jul-2022.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ Ⅎ𝑥(𝜑 ∧ 𝜓) | ||
Theorem | 19.3t 2197 | Closed form of 19.3 2198 and version of 19.9t 2200 with a universal quantifier. (Contributed by NM, 9-Nov-2020.) (Proof shortened by BJ, 9-Oct-2022.) |
⊢ (Ⅎ𝑥𝜑 → (∀𝑥𝜑 ↔ 𝜑)) | ||
Theorem | 19.3 2198 | A wff may be quantified with a variable not free in it. Version of 19.9 2201 with a universal quantifier. Theorem 19.3 of [Margaris] p. 89. See 19.3v 1986 for a version requiring fewer axioms. (Contributed by NM, 12-Mar-1993.) (Revised by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (∀𝑥𝜑 ↔ 𝜑) | ||
Theorem | 19.9d 2199 | A deduction version of one direction of 19.9 2201. (Contributed by NM, 14-May-1993.) (Revised by Mario Carneiro, 24-Sep-2016.) Revised to shorten other proofs. (Revised by Wolf Lammen, 14-Jul-2020.) df-nf 1788 changed. (Revised by Wolf Lammen, 11-Sep-2021.) (Proof shortened by Wolf Lammen, 8-Jul-2022.) |
⊢ (𝜓 → Ⅎ𝑥𝜑) ⇒ ⊢ (𝜓 → (∃𝑥𝜑 → 𝜑)) | ||
Theorem | 19.9t 2200 | Closed form of 19.9 2201 and version of 19.3t 2197 with an existential quantifier. (Contributed by NM, 13-May-1993.) (Revised by Mario Carneiro, 24-Sep-2016.) (Proof shortened by Wolf Lammen, 14-Jul-2020.) |
⊢ (Ⅎ𝑥𝜑 → (∃𝑥𝜑 ↔ 𝜑)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |