![]() |
Metamath
Proof Explorer Theorem List (p. 22 of 437) | < 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-28347) |
![]() (28348-29872) |
![]() (29873-43650) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | hbaev 2101* | Version of hbae 2396 with a disjoint variable condition, requiring fewer axioms. Instance of hbaevg 2100 and aev2 2104. (Contributed by Wolf Lammen, 22-Mar-2021.) |
⊢ (∀𝑥 𝑥 = 𝑦 → ∀𝑧∀𝑥 𝑥 = 𝑦) | ||
Theorem | hbnaevg 2102* | Generalization of hbnaev 2103. (Contributed by Wolf Lammen, 9-Apr-2021.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → ∀𝑧 ¬ ∀𝑡 𝑡 = 𝑢) | ||
Theorem | hbnaev 2103* | Any variable is free in ¬ ∀𝑥𝑥 = 𝑦, if 𝑥 and 𝑦 are distinct. This condition is dropped in hbnae 2398, at the expense of more axiom dependencies. Instance of hbnaevg 2102. (Contributed by Wolf Lammen, 9-Apr-2021.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → ∀𝑧 ¬ ∀𝑥 𝑥 = 𝑦) | ||
Theorem | aev2 2104* |
A version of aev 2099 with two universal quantifiers in the
consequent, and
a generalization of hbaevg 2100. One can prove similar statements with
arbitrary numbers of universal quantifiers in the consequent (the series
begins with aeveq 2098, aev 2099, aev2 2104).
Using aev 2099 and alrimiv 1970 (as in aev2ALT 2105), one can actually prove (with no more axioms) any scheme of the form (∀𝑥𝑥 = 𝑦 → PHI) , DV (𝑥, 𝑦) where PHI involves only setvar variables and the connectors →, ↔, ∧, ∨, ⊤, =, ∀, ∃, ∃*, ∃!, Ⅎ. An example is given by aevdemo 27892. This list cannot be extended to ¬ or ⊥ since the scheme ∀𝑥𝑥 = 𝑦 is consistent with ax-mp 5, ax-gen 1839, ax-1 6-- ax-13 2333 (as the one-element universe shows). (Contributed by BJ, 29-Mar-2021.) |
⊢ (∀𝑥 𝑥 = 𝑦 → ∀𝑧∀𝑡 𝑢 = 𝑣) | ||
Theorem | aev2ALT 2105* | Alternate proof of aev2 2104, bypassing hbaevg 2100. (Contributed by BJ, 23-Mar-2021.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (∀𝑥 𝑥 = 𝑦 → ∀𝑧∀𝑡 𝑢 = 𝑣) | ||
Syntax | wcel 2106 |
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 express, i.e., "prove", the wel 2107 of predicate calculus in terms of the wcel 2106 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 2763 for more information on the set theory usage of wcel 2106.) |
wff 𝐴 ∈ 𝐵 | ||
Theorem | wel 2107 |
Extend wff definition to include atomic formulas with the membership
predicate. This is read "𝑥 is an element of 𝑦",
"𝑥 is a
member of 𝑦", "𝑥 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 2107 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 2106. This lets us avoid overloading the ∈ connective, thus preventing ambiguity that would complicate certain Metamath parsers. However, logically wel 2107 is considered to be a primitive syntax, even though here it is artificially "derived" from wcel 2106. 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 2108 |
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 2112 that this axiom can be recovered from its weakened version ax8v 2109 where 𝑥 and 𝑦 are assumed to be disjoint variables. In particular, the only theorem referencing ax-8 2108 should be ax8v 2109. See the comment of ax8v 2109 for more details on these matters. (Contributed by NM, 30-Jun-1993.) (Revised by BJ, 7-Dec-2020.) Use ax8 2112 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Theorem | ax8v 2109* | Weakened version of ax-8 2108, with a disjoint variable condition on 𝑥, 𝑦. This should be the only proof referencing ax-8 2108, and it should be referenced only by its two weakened versions ax8v1 2110 and ax8v2 2111, from which ax-8 2108 is then rederived as ax8 2112, which shows that either ax8v 2109 or the conjunction of ax8v1 2110 and ax8v2 2111 is sufficient. (Contributed by BJ, 7-Dec-2020.) Use ax8 2112 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Theorem | ax8v1 2110* | First of two weakened versions of ax8v 2109, with an extra disjoint variable condition on 𝑥, 𝑧, see comments there. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Theorem | ax8v2 2111* | Second of two weakened versions of ax8v 2109, with an extra disjoint variable condition on 𝑦, 𝑧 see comments there. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Theorem | ax8 2112 | Proof of ax-8 2108 from ax8v1 2110 and ax8v2 2111, proving sufficiency of the conjunction of the latter two weakened versions of ax8v 2109, which is itself a weakened version of ax-8 2108. (Contributed by BJ, 7-Dec-2020.) (Proof shortened by Wolf Lammen, 11-Apr-2021.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Theorem | elequ1 2113 | An identity law for the non-logical predicate. (Contributed by NM, 30-Jun-1993.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 ↔ 𝑦 ∈ 𝑧)) | ||
Theorem | cleljust 2114* | When the class variables in definition df-clel 2773 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 2107 with the class variables in wcel 2106. (Contributed by NM, 28-Jan-2004.) Revised to use equsexvw 2051 in order to remove dependencies on ax-10 2134, ax-12 2162, ax-13 2333. 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 2115 |
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 2119 that this axiom can be recovered from its weakened version ax9v 2116 where 𝑥 and 𝑦 are assumed to be disjoint variables. In particular, the only theorem referencing ax-9 2115 should be ax9v 2116. See the comment of ax9v 2116 for more details on these matters. (Contributed by NM, 21-Jun-1993.) (Revised by BJ, 7-Dec-2020.) Use ax9 2119 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | ax9v 2116* | Weakened version of ax-9 2115, with a disjoint variable condition on 𝑥, 𝑦. This should be the only proof referencing ax-9 2115, and it should be referenced only by its two weakened versions ax9v1 2117 and ax9v2 2118, from which ax-9 2115 is then rederived as ax9 2119, which shows that either ax9v 2116 or the conjunction of ax9v1 2117 and ax9v2 2118 is sufficient. (Contributed by BJ, 7-Dec-2020.) Use ax9 2119 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | ax9v1 2117* | First of two weakened versions of ax9v 2116, with an extra disjoint variable condition on 𝑥, 𝑧, see comments there. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | ax9v2 2118* | Second of two weakened versions of ax9v 2116, with an extra disjoint variable condition on 𝑦, 𝑧 see comments there. (Contributed by BJ, 7-Dec-2020.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | ax9 2119 | Proof of ax-9 2115 from ax9v1 2117 and ax9v2 2118, proving sufficiency of the conjunction of the latter two weakened versions of ax9v 2116, which is itself a weakened version of ax-9 2115. (Contributed by BJ, 7-Dec-2020.) (Proof shortened by Wolf Lammen, 11-Apr-2021.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | elequ2 2120 | An identity law for the non-logical predicate. (Contributed by NM, 21-Jun-1993.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 ↔ 𝑧 ∈ 𝑦)) | ||
The original axiom schemes of Tarski's predicate calculus are ax-4 1853, ax-5 1953, ax6v 2022, ax-7 2054, ax-8 2108, and ax-9 2115, together with rule ax-gen 1839. See mmset.html#compare 1839. 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 2134, ax-11 2149, ax-12 2162, and ax-13 2333, 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 "metalogically 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 2162 has been proved to be required; see http://us.metamath.org/award2003.html#9a. Metalogical independence of the other three are open problems.) (There are additional predicate calculus axiom schemes included in set.mm such as ax-c5 35032, 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 2021 are bundled, but they are not in ax6v 2022. 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 2022 is the principal instance of ax-6 2021. 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 2021 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 2134, ax-11 2149, ax-12 2162, and ax-13 2333. "Translating" this to Metamath, it means that Tarski's axioms can prove any substitution instance of ax-10 2134, ax-11 2149, ax-12 2162, or ax-13 2333 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 2122, ax11w 2123, ax12w 2126, and ax13w 2129 are derived using only Tarski's axiom schemes, showing that Tarski's schemes can be used to derive all substitution instances of ax-10 2134, ax-11 2149, ax-12 2162, and ax-13 2333 meeting Conditions (1) and (2). (The "w" suffix stands for "weak version".) Each hypothesis of ax10w 2122, ax11w 2123, and ax12w 2126 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 2128 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 2124, ax12dgen 2127, ax13dgen1 2130, ax13dgen2 2131, ax13dgen3 2132, and ax13dgen4 2133. (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 2134, ax-11 2149, ax-12 2162, and ax-13 2333 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 2021 in an older system, so it seems the main purpose of his later ax6v 2022 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 2021 as our official axiom, we show that the degenerate instance holds in ax6dgen 2121. (Recall that in set.mm, the only statement referencing ax-6 2021 is ax6v 2022.) The case of sp 2166 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 2083, again requiring substitution instances of 𝜑 that meet Conditions (1) and (2) above. Note that our direct proof sp 2166 requires ax-12 2162, which is not part of Tarski's system. | ||
Theorem | ax6dgen 2121 | Tarski's system uses the weaker ax6v 2022 instead of the bundled ax-6 2021, so here we show that the degenerate case of ax-6 2021 can be derived. Even though ax-6 2021 is in the list of axioms used, recall that in set.mm, the only statement referencing ax-6 2021 is ax6v 2022. We later rederive from ax6v 2022 the bundled form as ax6 2347 with the help of the auxiliary axiom schemes. (Contributed by NM, 23-Apr-2017.) |
⊢ ¬ ∀𝑥 ¬ 𝑥 = 𝑥 | ||
Theorem | ax10w 2122* | Weak version of ax-10 2134 from which we can prove any ax-10 2134 instance not involving wff variables or bundling. Uses only Tarski's FOL axiom schemes. It is an alias of hbn1w 2090 introduced for labeling consistency. (Contributed by NM, 9-Apr-2017.) Use hbn1w 2090 instead. (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (¬ ∀𝑥𝜑 → ∀𝑥 ¬ ∀𝑥𝜑) | ||
Theorem | ax11w 2123* | Weak version of ax-11 2149 from which we can prove any ax-11 2149 instance not involving wff variables or bundling. Uses only Tarski's FOL axiom schemes. Unlike ax-11 2149, this theorem requires that 𝑥 and 𝑦 be distinct i.e. are not bundled. It is an alias of alcomiw 2087 introduced for labeling consistency. (Contributed by NM, 10-Apr-2017.) Use alcomiw 2087 instead. (New usage is discouraged.) |
⊢ (𝑦 = 𝑧 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥∀𝑦𝜑 → ∀𝑦∀𝑥𝜑) | ||
Theorem | ax11dgen 2124 | Degenerate instance of ax-11 2149 where bundled variables 𝑥 and 𝑦 have a common substitution. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) |
⊢ (∀𝑥∀𝑥𝜑 → ∀𝑥∀𝑥𝜑) | ||
Theorem | ax12wlem 2125* | Lemma for weak version of ax-12 2162. Uses only Tarski's FOL axiom schemes. In some cases, this lemma may lead to shorter proofs than ax12w 2126. (Contributed by NM, 10-Apr-2017.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | ax12w 2126* | Weak version of ax-12 2162 from which we can prove any ax-12 2162 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 2128. (Contributed by NM, 10-Apr-2017.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝑦 = 𝑧 → (𝜑 ↔ 𝜒)) ⇒ ⊢ (𝑥 = 𝑦 → (∀𝑦𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | ax12dgen 2127 | Degenerate instance of ax-12 2162 where bundled variables 𝑥 and 𝑦 have a common substitution. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) |
⊢ (𝑥 = 𝑥 → (∀𝑥𝜑 → ∀𝑥(𝑥 = 𝑥 → 𝜑))) | ||
Theorem | ax12wdemo 2128* | Example of an application of ax12w 2126 that results in an instance of ax-12 2162 for a contrived formula with mixed free and bound variables, (𝑥 ∈ 𝑦 ∧ ∀𝑥𝑧 ∈ 𝑥 ∧ ∀𝑦∀𝑧𝑦 ∈ 𝑥), in place of 𝜑. The proof illustrates bound variable renaming with cbvalvw 2085 to obtain fresh variables to avoid distinct variable clashes. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 14-Apr-2017.) |
⊢ (𝑥 = 𝑦 → (∀𝑦(𝑥 ∈ 𝑦 ∧ ∀𝑥 𝑧 ∈ 𝑥 ∧ ∀𝑦∀𝑧 𝑦 ∈ 𝑥) → ∀𝑥(𝑥 = 𝑦 → (𝑥 ∈ 𝑦 ∧ ∀𝑥 𝑧 ∈ 𝑥 ∧ ∀𝑦∀𝑧 𝑦 ∈ 𝑥)))) | ||
Theorem | ax13w 2129* | Weak version (principal instance) of ax-13 2333. (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 2122, ax11w 2123, and ax12w 2126. (Contributed by NM, 10-Apr-2017.) |
⊢ (¬ 𝑥 = 𝑦 → (𝑦 = 𝑧 → ∀𝑥 𝑦 = 𝑧)) | ||
Theorem | ax13dgen1 2130 | Degenerate instance of ax-13 2333 where bundled variables 𝑥 and 𝑦 have a common substitution. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) |
⊢ (¬ 𝑥 = 𝑥 → (𝑥 = 𝑧 → ∀𝑥 𝑥 = 𝑧)) | ||
Theorem | ax13dgen2 2131 | Degenerate instance of ax-13 2333 where bundled variables 𝑥 and 𝑧 have a common substitution. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) |
⊢ (¬ 𝑥 = 𝑦 → (𝑦 = 𝑥 → ∀𝑥 𝑦 = 𝑥)) | ||
Theorem | ax13dgen3 2132 | Degenerate instance of ax-13 2333 where bundled variables 𝑦 and 𝑧 have a common substitution. Uses only Tarski's FOL axiom schemes. (Contributed by NM, 13-Apr-2017.) |
⊢ (¬ 𝑥 = 𝑦 → (𝑦 = 𝑦 → ∀𝑥 𝑦 = 𝑦)) | ||
Theorem | ax13dgen4 2133 | Degenerate instance of ax-13 2333 where bundled variables 𝑥, 𝑦, and 𝑧 have a common substitution. Therefore, also a degenerate instance of ax13dgen1 2130, ax13dgen2 2131, and ax13dgen3 2132. Also an instance of the intuitionistic tautology pm2.21 121. 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 2134, ax-11 2149, ax-12 2162, and ax-13 2333 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 2122, ax11w 2123, ax12w 2126, and ax13w 2129, 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 2162 from all others has been shown, and independence of Tarski's ax-6 2021 from all others has been shown; see items 9a and 11 on http://us.metamath.org/award2003.html. | ||
Axiom | ax-10 2134 | Axiom of Quantified Negation. Axiom C5-2 of [Monk2] p. 113. This axiom scheme is logically redundant (see ax10w 2122) 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 2135 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 2134 through ax-13 2333, by invoking ax10w 2122 through ax13w 2129. We encourage proving theorems *without* ax-10 2134 through ax-13 2333 and moving them up to the ax-4 1853 through ax-9 2115 section. (New usage is discouraged.) |
⊢ (¬ ∀𝑥𝜑 → ∀𝑥 ¬ ∀𝑥𝜑) | ||
Theorem | hbn1 2135 | Alias for ax-10 2134 to be used instead of it. (Contributed by NM, 24-Jan-1993.) (Proof shortened by Wolf Lammen, 18-Aug-2014.) |
⊢ (¬ ∀𝑥𝜑 → ∀𝑥 ¬ ∀𝑥𝜑) | ||
Theorem | hbe1 2136 | The setvar 𝑥 is not free in ∃𝑥𝜑. Corresponds to the axiom (5) of modal logic (see also modal5 2148). (Contributed by NM, 24-Jan-1993.) |
⊢ (∃𝑥𝜑 → ∀𝑥∃𝑥𝜑) | ||
Theorem | hbe1a 2137 | Dual statement of hbe1 2136. Modified version of axc7e 2292 with a universally quantified consequent. (Contributed by Wolf Lammen, 15-Sep-2021.) |
⊢ (∃𝑥∀𝑥𝜑 → ∀𝑥𝜑) | ||
Theorem | nf5-1 2138 | One direction of nf5 2255 can be proved with a smaller footprint on axiom usage. (Contributed by Wolf Lammen, 16-Sep-2021.) |
⊢ (∀𝑥(𝜑 → ∀𝑥𝜑) → Ⅎ𝑥𝜑) | ||
Theorem | nf5i 2139 | Deduce that 𝑥 is not free in 𝜑 from the definition. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ Ⅎ𝑥𝜑 | ||
Theorem | nf5dh 2140 | Deduce that 𝑥 is not free in 𝜓 in a context. (Contributed by Mario Carneiro, 24-Sep-2016.) df-nf 1828 changed. (Revised by Wolf Lammen, 11-Oct-2021.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
Theorem | nf5dv 2141* | Apply the definition of not-free in a context. (Contributed by Mario Carneiro, 11-Aug-2016.) df-nf 1828 changed. (Revised by Wolf Lammen, 18-Sep-2021.) (Proof shortened by Wolf Lammen, 13-Jul-2022.) |
⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
Theorem | nf5dvOLD 2142* | Obsolete version of nf5dv 2141 as of 13-Jul-2022. (Contributed by Mario Carneiro, 11-Aug-2016.) df-nf 1828 changed. (Revised by Wolf Lammen, 18-Sep-2021.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
Theorem | nfe1 2143 | The setvar 𝑥 is not free in ∃𝑥𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ Ⅎ𝑥∃𝑥𝜑 | ||
Theorem | nfa1 2144 | The setvar 𝑥 is not free in ∀𝑥𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.) df-nf 1828 changed. (Revised by Wolf Lammen, 11-Sep-2021.) Remove dependency on ax-12 2162. (Revised by Wolf Lammen, 12-Oct-2021.) |
⊢ Ⅎ𝑥∀𝑥𝜑 | ||
Theorem | nfna1 2145 | A convenience theorem particularly designed to remove dependencies on ax-11 2149 in conjunction with distinctors. (Contributed by Wolf Lammen, 2-Sep-2018.) |
⊢ Ⅎ𝑥 ¬ ∀𝑥𝜑 | ||
Theorem | nfia1 2146 | Lemma 23 of [Monk2] p. 114. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥(∀𝑥𝜑 → ∀𝑥𝜓) | ||
Theorem | nfnf1 2147 | The setvar 𝑥 is not free in Ⅎ𝑥𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.) Remove dependency on ax-12 2162. (Revised by Wolf Lammen, 12-Oct-2021.) |
⊢ Ⅎ𝑥Ⅎ𝑥𝜑 | ||
Theorem | modal5 2148 | The analogue in our predicate calculus of axiom (5) of modal logic S5. See also hbe1 2136. (Contributed by NM, 5-Oct-2005.) |
⊢ (¬ ∀𝑥 ¬ 𝜑 → ∀𝑥 ¬ ∀𝑥 ¬ 𝜑) | ||
Axiom | ax-11 2149 | 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 2123) but is used as an auxiliary axiom scheme to achieve metalogical completeness. (Contributed by NM, 12-Mar-1993.) |
⊢ (∀𝑥∀𝑦𝜑 → ∀𝑦∀𝑥𝜑) | ||
Theorem | alcoms 2150 | Swap quantifiers in an antecedent. (Contributed by NM, 11-May-1993.) |
⊢ (∀𝑥∀𝑦𝜑 → 𝜓) ⇒ ⊢ (∀𝑦∀𝑥𝜑 → 𝜓) | ||
Theorem | alcom 2151 | Theorem 19.5 of [Margaris] p. 89. (Contributed by NM, 30-Jun-1993.) |
⊢ (∀𝑥∀𝑦𝜑 ↔ ∀𝑦∀𝑥𝜑) | ||
Theorem | alrot3 2152 | Theorem *11.21 in [WhiteheadRussell] p. 160. (Contributed by Andrew Salmon, 24-May-2011.) |
⊢ (∀𝑥∀𝑦∀𝑧𝜑 ↔ ∀𝑦∀𝑧∀𝑥𝜑) | ||
Theorem | alrot4 2153 | Rotate four universal quantifiers twice. (Contributed by NM, 2-Feb-2005.) (Proof shortened by Fan Zheng, 6-Jun-2016.) |
⊢ (∀𝑥∀𝑦∀𝑧∀𝑤𝜑 ↔ ∀𝑧∀𝑤∀𝑥∀𝑦𝜑) | ||
Theorem | excom 2154 | Theorem 19.11 of [Margaris] p. 89. (Contributed by NM, 5-Aug-1993.) Remove dependencies on ax-5 1953, ax-6 2021, ax-7 2054, ax-10 2134, ax-12 2162. (Revised by Wolf Lammen, 8-Jan-2018.) (Proof shortened by Wolf Lammen, 22-Aug-2020.) |
⊢ (∃𝑥∃𝑦𝜑 ↔ ∃𝑦∃𝑥𝜑) | ||
Theorem | excomim 2155 | 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 1953, ax-6 2021, ax-7 2054, ax-10 2134, ax-12 2162. (Revised by Wolf Lammen, 8-Jan-2018.) |
⊢ (∃𝑥∃𝑦𝜑 → ∃𝑦∃𝑥𝜑) | ||
Theorem | excom13 2156 | Swap 1st and 3rd existential quantifiers. (Contributed by NM, 9-Mar-1995.) |
⊢ (∃𝑥∃𝑦∃𝑧𝜑 ↔ ∃𝑧∃𝑦∃𝑥𝜑) | ||
Theorem | exrot3 2157 | Rotate existential quantifiers. (Contributed by NM, 17-Mar-1995.) |
⊢ (∃𝑥∃𝑦∃𝑧𝜑 ↔ ∃𝑦∃𝑧∃𝑥𝜑) | ||
Theorem | exrot4 2158 | Rotate existential quantifiers twice. (Contributed by NM, 9-Mar-1995.) |
⊢ (∃𝑥∃𝑦∃𝑧∃𝑤𝜑 ↔ ∃𝑧∃𝑤∃𝑥∃𝑦𝜑) | ||
Theorem | hbal 2159 | If 𝑥 is not free in 𝜑, it is not free in ∀𝑦𝜑. (Contributed by NM, 12-Mar-1993.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ (∀𝑦𝜑 → ∀𝑥∀𝑦𝜑) | ||
Theorem | hbald 2160 | Deduction form of bound-variable hypothesis builder hbal 2159. (Contributed by NM, 2-Jan-2002.) |
⊢ (𝜑 → ∀𝑦𝜑) & ⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → (∀𝑦𝜓 → ∀𝑥∀𝑦𝜓)) | ||
Theorem | nfa2 2161 | Lemma 24 of [Monk2] p. 114. (Contributed by Mario Carneiro, 24-Sep-2016.) Remove dependency on ax-12 2162. (Revised by Wolf Lammen, 18-Oct-2021.) |
⊢ Ⅎ𝑥∀𝑦∀𝑥𝜑 | ||
Axiom | ax-12 2162 |
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 2249). 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 35038 and was replaced with this shorter ax-12 2162 in Jan. 2007. The old axiom is proved from this one as theorem axc15 2386. Conversely, this axiom is proved from ax-c15 35038 as theorem ax12 2388. Juha Arpiainen proved the metalogical independence of this axiom (in the form of the older axiom ax-c15 35038) from the others on 19-Jan-2006. See item 9a at http://us.metamath.org/award2003.html. See ax12v 2163 and ax12v2 2164 for other equivalents of this axiom that (unlike this axiom) have distinct variable restrictions. This axiom scheme is logically redundant (see ax12w 2126) but is used as an auxiliary axiom scheme to achieve scheme completeness. (Contributed by NM, 22-Jan-2007.) (New usage is discouraged.) |
⊢ (𝑥 = 𝑦 → (∀𝑦𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | ax12v 2163* |
This is essentially axiom ax-12 2162 weakened by additional restrictions on
variables. Besides axc11r 2332, this theorem should be the only one
referencing ax-12 2162 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 1953, 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 2164* | It is possible to remove any restriction on 𝜑 in ax12v 2163. Same as Axiom C8 of [Monk2] p. 105. Use ax12v 2163 instead when sufficient. (Contributed by NM, 5-Aug-1993.) Remove dependencies on ax-10 2134 and ax-13 2333. (Revised by Jim Kingdon, 15-Dec-2017.) (Proof shortened by Wolf Lammen, 8-Dec-2019.) |
⊢ (𝑥 = 𝑦 → (𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | 19.8a 2165 | 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 2029 for a version with a disjoint variable condition requiring fewer axioms. (Contributed by NM, 9-Jan-1993.) Allow a shortening of sp 2166. (Revised by Wolf Lammen, 13-Jan-2018.) (Proof shortened by Wolf Lammen, 8-Dec-2019.) |
⊢ (𝜑 → ∃𝑥𝜑) | ||
Theorem | sp 2166 |
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 2427. This theorem shows that our obsolete axiom ax-c5 35032 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 2162. It is thought the best we can do using only Tarski's axioms is spw 2083. (Contributed by NM, 21-May-2008.) (Proof shortened by Scott Fenton, 24-Jan-2011.) (Proof shortened by Wolf Lammen, 13-Jan-2018.) |
⊢ (∀𝑥𝜑 → 𝜑) | ||
Theorem | spi 2167 | Inference rule of universal instantiation, or universal specialization. Converse of the inference rule of (universal) generalization ax-gen 1839. Contrary to the rule of generalization, its closed form is valid, see sp 2166. (Contributed by NM, 5-Aug-1993.) |
⊢ ∀𝑥𝜑 ⇒ ⊢ 𝜑 | ||
Theorem | sps 2168 | Generalization of antecedent. (Contributed by NM, 5-Jan-1993.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (∀𝑥𝜑 → 𝜓) | ||
Theorem | 2sp 2169 | A double specialization (see sp 2166). Another double specialization, closer to PM*11.1, is 2stdpc4 2428. (Contributed by BJ, 15-Sep-2018.) |
⊢ (∀𝑥∀𝑦𝜑 → 𝜑) | ||
Theorem | spsd 2170 | Deduction generalizing antecedent. (Contributed by NM, 17-Aug-1994.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 → 𝜒)) | ||
Theorem | 19.2g 2171 | Theorem 19.2 of [Margaris] p. 89, generalized to use two setvar variables. Use 19.2 2026 when sufficient. (Contributed by Mel L. O'Cat, 31-Mar-2008.) |
⊢ (∀𝑥𝜑 → ∃𝑦𝜑) | ||
Theorem | 19.21bi 2172 | Inference form of 19.21 2191 and also deduction form of sp 2166. (Contributed by NM, 26-May-1993.) |
⊢ (𝜑 → ∀𝑥𝜓) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | 19.21bbi 2173 | Inference removing two universal quantifiers. Version of 19.21bi 2172 with two quantifiers. (Contributed by NM, 20-Apr-1994.) |
⊢ (𝜑 → ∀𝑥∀𝑦𝜓) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | 19.23bi 2174 | Inference form of Theorem 19.23 of [Margaris] p. 90, see 19.23 2196. (Contributed by NM, 12-Mar-1993.) |
⊢ (∃𝑥𝜑 → 𝜓) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | nexr 2175 | Inference associated with the contrapositive of 19.8a 2165. (Contributed by Jeff Hankins, 26-Jul-2009.) |
⊢ ¬ ∃𝑥𝜑 ⇒ ⊢ ¬ 𝜑 | ||
Theorem | qexmid 2176 | Quantified excluded middle (see exmid 881). 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 2177 | Consequence of the definition of not-free. (Contributed by Mario Carneiro, 26-Sep-2016.) df-nf 1828 changed. (Revised by Wolf Lammen, 11-Sep-2021.) |
⊢ (Ⅎ𝑥𝜑 → (𝜑 → ∀𝑥𝜑)) | ||
Theorem | nf5ri 2178 | Consequence of the definition of not-free. (Contributed by Mario Carneiro, 11-Aug-2016.) (Proof shortened by Wolf Lammen, 15-Mar-2023.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (𝜑 → ∀𝑥𝜑) | ||
Theorem | nf5riOLD 2179 | Obsolete proof of nf5ri 2178 as of 15-Mar-2023. (Contributed by Mario Carneiro, 11-Aug-2016.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (𝜑 → ∀𝑥𝜑) | ||
Theorem | nf5rd 2180 | Consequence of the definition of not-free in a context. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) | ||
Theorem | nfim1 2181 | A closed form of nfim 1943. (Contributed by NM, 2-Jun-1993.) (Revised by Mario Carneiro, 24-Sep-2016.) (Proof shortened by Wolf Lammen, 2-Jan-2018.) df-nf 1828 changed. (Revised by Wolf Lammen, 18-Sep-2021.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ Ⅎ𝑥(𝜑 → 𝜓) | ||
Theorem | nfan1 2182 | A closed form of nfan 1946. (Contributed by Mario Carneiro, 3-Oct-2016.) df-nf 1828 changed. (Revised by Wolf Lammen, 18-Sep-2021.) (Proof shortened by Wolf Lammen, 7-Jul-2022.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ Ⅎ𝑥(𝜑 ∧ 𝜓) | ||
Theorem | nfan1OLDOLD 2183 | Obsolete version of nfan1 2182 as of 7-Jul-2022. (Contributed by Mario Carneiro, 3-Oct-2016.) df-nf 1828 changed. (Revised by Wolf Lammen, 18-Sep-2021.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ Ⅎ𝑥(𝜑 ∧ 𝜓) | ||
Theorem | 19.3t 2184 | Closed form of 19.3 2185 and version of 19.9t 2188 with a universal quantifier. (Contributed by NM, 9-Nov-2020.) (Proof shortened by BJ, 9-Oct-2022.) |
⊢ (Ⅎ𝑥𝜑 → (∀𝑥𝜑 ↔ 𝜑)) | ||
Theorem | 19.3 2185 | A wff may be quantified with a variable not free in it. Version of 19.9 2189 with a universal quantifier. Theorem 19.3 of [Margaris] p. 89. See 19.3v 2031 for a version requiring fewer axioms. (Contributed by NM, 12-Mar-1993.) (Revised by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (∀𝑥𝜑 ↔ 𝜑) | ||
Theorem | 19.9d 2186 | A deduction version of one direction of 19.9 2189. (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 1828 changed. (Revised by Wolf Lammen, 11-Sep-2021.) (Proof shortened by Wolf Lammen, 8-Jul-2022.) |
⊢ (𝜓 → Ⅎ𝑥𝜑) ⇒ ⊢ (𝜓 → (∃𝑥𝜑 → 𝜑)) | ||
Theorem | 19.9dOLDOLD 2187 | Obsolete version of 19.9d 2186 as of 8-Jul-2022. (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 1828 changed. (Revised by Wolf Lammen, 11-Sep-2021.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝜓 → Ⅎ𝑥𝜑) ⇒ ⊢ (𝜓 → (∃𝑥𝜑 → 𝜑)) | ||
Theorem | 19.9t 2188 | Closed form of 19.9 2189 and version of 19.3t 2184 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.) |
⊢ (Ⅎ𝑥𝜑 → (∃𝑥𝜑 ↔ 𝜑)) | ||
Theorem | 19.9 2189 | A wff may be existentially quantified with a variable not free in it. Version of 19.3 2185 with an existential quantifier. Theorem 19.9 of [Margaris] p. 89. See 19.9v 2030 for a version requiring fewer axioms. (Contributed by FL, 24-Mar-2007.) (Revised by Mario Carneiro, 24-Sep-2016.) (Proof shortened by Wolf Lammen, 30-Dec-2017.) Revised to shorten other proofs. (Revised by Wolf Lammen, 14-Jul-2020.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (∃𝑥𝜑 ↔ 𝜑) | ||
Theorem | 19.21t 2190 | Closed form of Theorem 19.21 of [Margaris] p. 90, see 19.21 2191. (Contributed by NM, 27-May-1997.) (Revised by Mario Carneiro, 24-Sep-2016.) (Proof shortened by Wolf Lammen, 3-Jan-2018.) df-nf 1828 changed. (Revised by Wolf Lammen, 11-Sep-2021.) (Proof shortened by BJ, 3-Nov-2021.) |
⊢ (Ⅎ𝑥𝜑 → (∀𝑥(𝜑 → 𝜓) ↔ (𝜑 → ∀𝑥𝜓))) | ||
Theorem | 19.21 2191 | Theorem 19.21 of [Margaris] p. 90. The hypothesis can be thought of as "𝑥 is not free in 𝜑". See 19.21v 1982 for a version requiring fewer axioms. See also 19.21h 2260. (Contributed by NM, 14-May-1993.) (Revised by Mario Carneiro, 24-Sep-2016.) df-nf 1828 changed. (Revised by Wolf Lammen, 18-Sep-2021.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (∀𝑥(𝜑 → 𝜓) ↔ (𝜑 → ∀𝑥𝜓)) | ||
Theorem | stdpc5 2192 | An axiom scheme of standard predicate calculus that emulates Axiom 5 of [Mendelson] p. 69. The hypothesis Ⅎ𝑥𝜑 can be thought of as emulating "𝑥 is not free in 𝜑". With this definition, the meaning of "not free" is less restrictive than the usual textbook definition; for example 𝑥 would not (for us) be free in 𝑥 = 𝑥 by nfequid 2059. This theorem scheme can be proved as a metatheorem of Mendelson's axiom system, even though it is slightly stronger than his Axiom 5. See stdpc5v 1981 for a version requiring fewer axioms. (Contributed by NM, 22-Sep-1993.) (Revised by Mario Carneiro, 12-Oct-2016.) (Proof shortened by Wolf Lammen, 1-Jan-2018.) Remove dependency on ax-10 2134. (Revised by Wolf Lammen, 4-Jul-2021.) (Proof shortened by Wolf Lammen, 11-Oct-2021.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (∀𝑥(𝜑 → 𝜓) → (𝜑 → ∀𝑥𝜓)) | ||
Theorem | 19.21-2 2193 | Version of 19.21 2191 with two quantifiers. (Contributed by NM, 4-Feb-2005.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑦𝜑 ⇒ ⊢ (∀𝑥∀𝑦(𝜑 → 𝜓) ↔ (𝜑 → ∀𝑥∀𝑦𝜓)) | ||
Theorem | 19.23t 2194 | Closed form of Theorem 19.23 of [Margaris] p. 90. See 19.23 2196. (Contributed by NM, 7-Nov-2005.) (Proof shortened by Wolf Lammen, 13-Aug-2020.) df-nf 1828 changed. (Revised by Wolf Lammen, 11-Sep-2021.) (Proof shortened by BJ, 8-Oct-2022.) |
⊢ (Ⅎ𝑥𝜓 → (∀𝑥(𝜑 → 𝜓) ↔ (∃𝑥𝜑 → 𝜓))) | ||
Theorem | 19.23tOLD 2195 | Obsolete proof of 19.23t as of 8-Oct-2022. (Contributed by NM, 7-Nov-2005.) (Proof shortened by Wolf Lammen, 13-Aug-2020.) df-nf 1828 changed. (Revised by Wolf Lammen, 11-Sep-2021.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (Ⅎ𝑥𝜓 → (∀𝑥(𝜑 → 𝜓) ↔ (∃𝑥𝜑 → 𝜓))) | ||
Theorem | 19.23 2196 | Theorem 19.23 of [Margaris] p. 90. See 19.23v 1985 for a version requiring fewer axioms. (Contributed by NM, 24-Jan-1993.) (Revised by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥𝜓 ⇒ ⊢ (∀𝑥(𝜑 → 𝜓) ↔ (∃𝑥𝜑 → 𝜓)) | ||
Theorem | alimd 2197 | Deduction form of Theorem 19.20 of [Margaris] p. 90, see alim 1854. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 → ∀𝑥𝜒)) | ||
Theorem | alrimi 2198 | Inference form of Theorem 19.21 of [Margaris] p. 90, see 19.21 2191. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → 𝜓) ⇒ ⊢ (𝜑 → ∀𝑥𝜓) | ||
Theorem | alrimdd 2199 | Deduction form of Theorem 19.21 of [Margaris] p. 90, see 19.21 2191. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) & ⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → ∀𝑥𝜒)) | ||
Theorem | alrimd 2200 | Deduction form of Theorem 19.21 of [Margaris] p. 90, see 19.21 2191. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → ∀𝑥𝜒)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |