![]() |
Metamath
Proof Explorer Theorem List (p. 335 of 429) | < 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-27903) |
![]() (27904-29428) |
![]() (29429-42879) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | wl-impchain-mp-1 33401 | This theorem is in fact a copy of wl-syl 33376, and repeated here to demonstrate a recursive proof scheme. The number '1' in the theorem name indicates that a chain of length 1 is modified. (Contributed by Wolf Lammen, 6-Jul-2019.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝜒 → 𝜓) & ⊢ (𝜓 → 𝜑) ⇒ ⊢ (𝜒 → 𝜑) | ||
Theorem | wl-impchain-mp-2 33402 | This theorem is in fact a copy of wl-syl6 33384, and repeated here to demonstrate a recursive proof scheme. The number '2' in the theorem name indicates that a chain of length 2 is modified. (Contributed by Wolf Lammen, 6-Jul-2019.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝜃 → (𝜒 → 𝜓)) & ⊢ (𝜓 → 𝜑) ⇒ ⊢ (𝜃 → (𝜒 → 𝜑)) | ||
Theorem | wl-impchain-com-1.x 33403 |
It is often convenient to have the antecedent under focus in first
position, so we can apply immediate theorem forms (as opposed to
deduction, tautology form). This series of theorems swaps the first with
the last antecedent in an implication chain. This kind of swapping is
self-inverse, whence we prefer it over, say, rotating theorems. A
consequent can hide a tail of a longer chain, so theorems of this series
appear as swapping a pair of antecedents with fixed offsets. This form of
swapping antecedents is flexible enough to allow for any permutation of
antecedents in an implication chain.
The first elements of this series correspond to com12 32, com13 88, com14 96 and com15 101 in the main part. The proofs of this series aim at automated proving using a simple recursive scheme. It employs the previous theorem in the series along with a sample from the wl-impchain-mp-x 33399 series developed before. (Contributed by Wolf Lammen, 17-Nov-2019.) |
⊢ ⊤ | ||
Theorem | wl-impchain-com-1.1 33404 |
A degenerate form of antecedent swapping. The number '1' in the theorem
name indicates that it handles a chain of length 1.
Since there is just one antecedent in the chain, there is nothing to swap. Non-degenerated forms begin with wl-impchain-com-1.2 33405, for more see there. (Contributed by Wolf Lammen, 7-Jul-2019.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝜓 → 𝜑) ⇒ ⊢ (𝜓 → 𝜑) | ||
Theorem | wl-impchain-com-1.2 33405 |
This theorem is in fact a copy of wl-com12 33388, and repeated here to
demonstrate a simple proof scheme. The number '2' in the theorem name
indicates that a chain of length 2 is modified.
See wl-impchain-com-1.x 33403 for more information how this proof is generated. (Contributed by Wolf Lammen, 7-Jul-2019.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝜒 → (𝜓 → 𝜑)) ⇒ ⊢ (𝜓 → (𝜒 → 𝜑)) | ||
Theorem | wl-impchain-com-1.3 33406 |
This theorem is in fact a copy of com13 88, and repeated here to
demonstrate a simple proof scheme. The number '3' in the theorem name
indicates that a chain of length 3 is modified.
See wl-impchain-com-1.x 33403 for more information how this proof is generated. (Contributed by Wolf Lammen, 7-Jul-2019.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝜃 → (𝜒 → (𝜓 → 𝜑))) ⇒ ⊢ (𝜓 → (𝜒 → (𝜃 → 𝜑))) | ||
Theorem | wl-impchain-com-1.4 33407 |
This theorem is in fact a copy of com14 96, and repeated here to
demonstrate a simple proof scheme. The number '4' in the theorem name
indicates that a chain of length 4 is modified.
See wl-impchain-com-1.x 33403 for more information how this proof is generated. (Contributed by Wolf Lammen, 7-Jul-2019.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝜂 → (𝜃 → (𝜒 → (𝜓 → 𝜑)))) ⇒ ⊢ (𝜓 → (𝜃 → (𝜒 → (𝜂 → 𝜑)))) | ||
Theorem | wl-impchain-com-n.m 33408 |
This series of theorems allow swapping any two antecedents in an
implication chain. The theorem names follow a pattern wl-impchain-com-n.m
with integral numbers n < m, that swaps the m-th antecedent with n-th
one
in an implication chain. It is sufficient to restrict the length of the
chain to m, too, since the consequent can be assumed to be the tail right
of the m-th antecedent of any arbitrary sized implication chain. We
further assume n > 1, since the wl-impchain-com-1.x 33403 series already
covers the special case n = 1.
Being able to swap any two antecedents in an implication chain lays the foundation of permuting its antecedents arbitrarily. The proofs of this series aim at automated proofing using a simple scheme. Any instance of this series is a triple step of swapping the first and n-th antecedent, then the first and the m-th, then the first and the n-th antecedent again. Each of these steps is an instance of the wl-impchain-com-1.x 33403 series. (Contributed by Wolf Lammen, 17-Nov-2019.) |
⊢ ⊤ | ||
Theorem | wl-impchain-com-2.3 33409 | This theorem is in fact a copy of com23 86. It starts a series of theorems named after wl-impchain-com-n.m 33408. For more information see there. (Contributed by Wolf Lammen, 12-Nov-2019.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝜃 → (𝜒 → (𝜓 → 𝜑))) ⇒ ⊢ (𝜃 → (𝜓 → (𝜒 → 𝜑))) | ||
Theorem | wl-impchain-com-2.4 33410 | This theorem is in fact a copy of com24 95. It is another instantiation of theorems named after wl-impchain-com-n.m 33408. For more information see there. (Contributed by Wolf Lammen, 17-Nov-2019.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝜂 → (𝜃 → (𝜒 → (𝜓 → 𝜑)))) ⇒ ⊢ (𝜂 → (𝜓 → (𝜒 → (𝜃 → 𝜑)))) | ||
Theorem | wl-impchain-com-3.2.1 33411 | This theorem is in fact a copy of com3r 87. The proof is an example of how to arrive at arbitrary permutations of antecedents, using only swapping theorems. The recursion principle is to first swap the correct antecedent to the position just before the consequent, and then employ a theorem handling an implication chain of length one less to reorder the others. (Contributed by Wolf Lammen, 17-Nov-2019.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝜃 → (𝜒 → (𝜓 → 𝜑))) ⇒ ⊢ (𝜓 → (𝜃 → (𝜒 → 𝜑))) | ||
Theorem | wl-impchain-a1-x 33412 |
If an implication chain is assumed (hypothesis) or proven (theorem) to
hold, then we may add any extra antecedent to it, without changing its
truth. This is expressed in its simplest form in wl-a1i 33381, that allows
us prepending an arbitrary antecedent to an implication chain. Using our
antecedent swapping theorems described in wl-impchain-com-n.m 33408, we may
then move such a prepended antecedent to any desired location within all
antecedents. The first series of theorems of this kind adds a single
antecedent somewhere to an implication chain. The appended number in the
theorem name indicates its position within all antecedents, 1 denoting the
head position. A second theorem series extends this idea to multiple
additions (TODO).
Adding antecedents to an implication chain usually weakens their universality. The consequent afterwards dependends on more conditions than before, which renders the implication chain less versatile. So you find this proof technique mostly when you adjust a chain to a hypothesis of a rule. A common case are syllogisms merging two implication chains into one. The first elements of the first series correspond to a1i 11, a1d 25 and a1dd 50 in the main part. The proofs of this series aim at automated proving using a simple recursive scheme. It employs the previous theorem in the series along with a sample from the wl-impchain-com-1.x 33403 series developed before. (Contributed by Wolf Lammen, 20-Jun-2020.) |
⊢ ⊤ | ||
Theorem | wl-impchain-a1-1 33413 | Inference rule, a copy of a1i 11. Head start of a recursive proof pattern. (Contributed by Wolf Lammen, 20-Jun-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ 𝜑 ⇒ ⊢ (𝜓 → 𝜑) | ||
Theorem | wl-impchain-a1-2 33414 | Inference rule, a copy of a1d 25. First recursive proof based on the previous instance. (Contributed by Wolf Lammen, 20-Jun-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (𝜑 → (𝜒 → 𝜓)) | ||
Theorem | wl-impchain-a1-3 33415 | Inference rule, a copy of a1dd 50. A recursive proof depending on previous instances, and demonstrating the proof pattern. (Contributed by Wolf Lammen, 20-Jun-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜒))) | ||
Axiom | ax-wl-13v 33416* |
A version of ax13v 2283 with a distinctor instead of a distinct
variable
expression.
Had we additionally required 𝑥 and 𝑦 be distinct, too, this theorem would have been a direct consequence of ax-5 1879. So essentially this theorem states, that a distinct variable condition between set variables can be replaced with a distinctor expression. (Contributed by Wolf Lammen, 23-Jul-2021.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → (𝑦 = 𝑧 → ∀𝑥 𝑦 = 𝑧)) | ||
Theorem | wl-ax13lem1 33417* | A version of ax-wl-13v 33416 with one distinct variable restriction dropped. For convenience, 𝑦 is kept on the right side of equations. This proof bases on ideas from NM, 24-Dec-2015. (Contributed by Wolf Lammen, 23-Jul-2021.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → (𝑧 = 𝑦 → ∀𝑥 𝑧 = 𝑦)) | ||
Theorem | wl-jarri 33418 | Dropping a nested antecedent. This theorem is one of two reversions of ja 173. Since ja 173 is reversible, a nested (chain of) implication(s) is just a packed notation of two or more theorems/hypotheses with a common consequent. axc5c7 34515 is an instance of this idea. (Contributed by Wolf Lammen, 20-Sep-2013.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ ((𝜑 → 𝜓) → 𝜒) ⇒ ⊢ (𝜓 → 𝜒) | ||
Theorem | wl-jarli 33419 | Dropping a nested consequent. This theorem is one of two reversions of ja 173. Since ja 173 is reversible, one can conclude, that a nested (chain of) implication(s) is just a packed notation of two or more theorems/ hypotheses with a common consequent. axc5c7 34515 is an instance of this idea. (Contributed by Wolf Lammen, 4-Oct-2013.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ ((𝜑 → 𝜓) → 𝜒) ⇒ ⊢ (¬ 𝜑 → 𝜒) | ||
Theorem | wl-mps 33420 | Replacing a nested consequent. A sort of modus ponens in antecedent position. (Contributed by Wolf Lammen, 20-Sep-2013.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ ((𝜑 → 𝜒) → 𝜃) ⇒ ⊢ ((𝜑 → 𝜓) → 𝜃) | ||
Theorem | wl-syls1 33421 | Replacing a nested consequent. A sort of syllogism in antecedent position. (Contributed by Wolf Lammen, 20-Sep-2013.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝜓 → 𝜒) & ⊢ ((𝜑 → 𝜒) → 𝜃) ⇒ ⊢ ((𝜑 → 𝜓) → 𝜃) | ||
Theorem | wl-syls2 33422 | Replacing a nested antecedent. A sort of syllogism in antecedent position. (Contributed by Wolf Lammen, 20-Sep-2013.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝜑 → 𝜓) & ⊢ ((𝜑 → 𝜒) → 𝜃) ⇒ ⊢ ((𝜓 → 𝜒) → 𝜃) | ||
Theorem | wl-embant 33423 | A true wff can always be added as a nested antecedent to an antecedent. Note: this theorem is intuitionistically valid. (Contributed by Wolf Lammen, 4-Oct-2013.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ 𝜑 & ⊢ (𝜓 → 𝜒) ⇒ ⊢ ((𝜑 → 𝜓) → 𝜒) | ||
Theorem | wl-orel12 33424 | In a conjunctive normal form a pair of nodes like (𝜑 ∨ 𝜓) ∧ (¬ 𝜑 ∨ 𝜒) eliminates the need of a node (𝜓 ∨ 𝜒). This theorem allows simplifications in that respect. (Contributed by Wolf Lammen, 20-Jun-2020.) |
⊢ (((𝜑 ∨ 𝜓) ∧ (¬ 𝜑 ∨ 𝜒)) → (𝜓 ∨ 𝜒)) | ||
Theorem | wl-cases2-dnf 33425 | A particular instance of orddi 931 and anddi 932 converting between disjunctive and conjunctive normal forms, when both 𝜑 and ¬ 𝜑 appear. This theorem in fact rephrases cases2 1016, and is related to consensus 1023. I restate it here in DNF and CNF. The proof deliberately does not use df-ifp 1033 and dfifp4 1036, by which it can be shortened. (Contributed by Wolf Lammen, 21-Jun-2020.) (Proof modification is discouraged.) |
⊢ (((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ 𝜒)) ↔ ((¬ 𝜑 ∨ 𝜓) ∧ (𝜑 ∨ 𝜒))) | ||
Theorem | wl-dfnan2 33426 | An alternative definition of "nand" based on imnan 437. See df-nan 1488 for the original definition. This theorem allows various shortenings. (Contributed by Wolf Lammen, 26-Jun-2020.) |
⊢ ((𝜑 ⊼ 𝜓) ↔ (𝜑 → ¬ 𝜓)) | ||
Theorem | wl-nancom 33427 | The 'nand' operator commutes. (Contributed by Mario Carneiro, 9-May-2015.) (Revised by Wolf Lammen, 26-Jun-2020.) |
⊢ ((𝜑 ⊼ 𝜓) ↔ (𝜓 ⊼ 𝜑)) | ||
Theorem | wl-nannan 33428 | Lemma for handling nested 'nand's. (Contributed by Jeff Hoffman, 19-Nov-2007.) (Revised by Wolf Lammen, 26-Jun-2020.) |
⊢ ((𝜑 ⊼ (𝜓 ⊼ 𝜒)) ↔ (𝜑 → (𝜓 ∧ 𝜒))) | ||
Theorem | wl-nannot 33429 | Show equivalence between negation and the Nicod version. To derive nic-dfneg 1635, apply nanbi 1494. (Contributed by Jeff Hoffman, 19-Nov-2007.) (Revised by Wolf Lammen, 26-Jun-2020.) |
⊢ (¬ 𝜑 ↔ (𝜑 ⊼ 𝜑)) | ||
Theorem | wl-nanbi1 33430 | Introduce a right anti-conjunct to both sides of a logical equivalence. (Contributed by SF, 2-Jan-2018.) (Revised by Wolf Lammen, 27-Jun-2020.) |
⊢ ((𝜑 ↔ 𝜓) → ((𝜑 ⊼ 𝜒) ↔ (𝜓 ⊼ 𝜒))) | ||
Theorem | wl-nanbi2 33431 | Introduce a left anti-conjunct to both sides of a logical equivalence. (Contributed by SF, 2-Jan-2018.) (Revised by Wolf Lammen, 27-Jun-2020.) |
⊢ ((𝜑 ↔ 𝜓) → ((𝜑 ⊼ 𝜒) ↔ (𝜓 ⊼ 𝜒))) | ||
Theorem | wl-naev 33432* | If some set variables can assume different values, then any two distinct set variables cannot always be the same. (Contributed by Wolf Lammen, 10-Aug-2019.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → ¬ ∀𝑢 𝑢 = 𝑣) | ||
Theorem | wl-hbae1 33433 | This specialization of hbae 2348 does not depend on ax-11 2074. (Contributed by Wolf Lammen, 8-Aug-2021.) |
⊢ (∀𝑥 𝑥 = 𝑦 → ∀𝑦∀𝑥 𝑥 = 𝑦) | ||
Theorem | wl-naevhba1v 33434* | An instance of hbn1w 2015 applied to equality. (Contributed by Wolf Lammen, 7-Apr-2021.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → ∀𝑥 ¬ ∀𝑥 𝑥 = 𝑦) | ||
Theorem | wl-hbnaev 33435* | Any variable is free in ¬ ∀𝑥𝑥 = 𝑦, if 𝑥 and 𝑦 are distinct. The latter condition can actually be lifted, but this version is easier to prove. The proof does not use ax-10 2059. (Contributed by Wolf Lammen, 9-Apr-2021.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → ∀𝑧 ¬ ∀𝑥 𝑥 = 𝑦) | ||
Theorem | wl-spae 33436 |
Prove an instance of sp 2091 from ax-13 2282 and Tarski's FOL only, without
distinct variable conditions. The antecedent ∀𝑥𝑥 = 𝑦 holds in a
multi-object universe only if 𝑦 is substituted for 𝑥, or
vice
versa, i.e. both variables are effectively the same. The converse
¬ ∀𝑥𝑥 = 𝑦 indicates that both variables are
distinct, and it so
provides a simple translation of a distinct variable condition to a
logical term. In case studies ∀𝑥𝑥 = 𝑦 and ¬
∀𝑥𝑥 = 𝑦 can
help eliminating distinct variable conditions.
The antecedent ∀𝑥𝑥 = 𝑦 is expressed in the theorem's name by the abbreviation ae standing for 'all equal'. Note that we cannot provide a logical predicate telling us directly whether a logical expression contains a particular variable, as such a construct would usually contradict ax-12 2087. Note that this theorem is also provable from ax-12 2087 alone, so you can pick the axiom it is based on. Compare this result to 19.3v 1954 and spaev 2020 having distinct variable conditions, but a smaller footprint on axiom usage. (Contributed by Wolf Lammen, 5-Apr-2021.) |
⊢ (∀𝑥 𝑥 = 𝑦 → 𝑥 = 𝑦) | ||
Theorem | wl-cbv3vv 33437* | Avoiding ax-11 2074. (Contributed by Wolf Lammen, 30-Aug-2021.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 → 𝜓)) ⇒ ⊢ (∀𝑥𝜑 → ∀𝑦𝜓) | ||
Theorem | wl-speqv 33438* | Under the assumption ¬ 𝑥 = 𝑦 a specialized version of sp 2091 is provable from Tarski's FOL and ax13v 2283 only. Note that this reverts the implication in ax13lem1 2284, so in fact (¬ 𝑥 = 𝑦 → (∀𝑥𝑧 = 𝑦 ↔ 𝑧 = 𝑦)) holds. (Contributed by Wolf Lammen, 17-Apr-2021.) |
⊢ (¬ 𝑥 = 𝑦 → (∀𝑥 𝑧 = 𝑦 → 𝑧 = 𝑦)) | ||
Theorem | wl-19.8eqv 33439* | Under the assumption ¬ 𝑥 = 𝑦 a specialized version of 19.8a 2090 is provable from Tarski's FOL and ax13v 2283 only. Note that this reverts the implication in ax13lem2 2332, so in fact (¬ 𝑥 = 𝑦 → (∃𝑥𝑧 = 𝑦 ↔ 𝑧 = 𝑦)) holds. (Contributed by Wolf Lammen, 17-Apr-2021.) |
⊢ (¬ 𝑥 = 𝑦 → (𝑧 = 𝑦 → ∃𝑥 𝑧 = 𝑦)) | ||
Theorem | wl-19.2reqv 33440* | Under the assumption ¬ 𝑥 = 𝑦 the reverse direction of 19.2 1949 is provable from Tarski's FOL and ax13v 2283 only. Note that in conjunction with 19.2 1949 in fact (¬ 𝑥 = 𝑦 → (∀𝑥𝑧 = 𝑦 ↔ ∃𝑥𝑧 = 𝑦)) holds. (Contributed by Wolf Lammen, 17-Apr-2021.) |
⊢ (¬ 𝑥 = 𝑦 → (∃𝑥 𝑧 = 𝑦 → ∀𝑥 𝑧 = 𝑦)) | ||
Theorem | wl-dveeq12 33441* | The current form of ax-13 2282 has a particular disadvantage: The condition ¬ 𝑥 = 𝑦 is less versatile than the general form ¬ ∀𝑥𝑥 = 𝑦. You need ax-10 2059 to arrive at the more general form presented here. You need 19.8a 2090 (or ax-12 2087) to restore 𝑦 = 𝑧 from ∃𝑥𝑦 = 𝑧 again. (Contributed by Wolf Lammen, 9-Jun-2021.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → (∃𝑥 𝑧 = 𝑦 → ∀𝑥 𝑧 = 𝑦)) | ||
Theorem | wl-nfalv 33442* | If 𝑥 is not present in 𝜑, it is not free in ∀𝑦𝜑. (Contributed by Wolf Lammen, 11-Jan-2020.) |
⊢ Ⅎ𝑥∀𝑦𝜑 | ||
Theorem | wl-nfimf1 33443 | An antecedent is irrelevant to a not-free property, if it always holds. I used this variant of nfim 1865 in dvelimdf 2366 to simplify the proof. (Contributed by Wolf Lammen, 14-Oct-2018.) |
⊢ (∀𝑥𝜑 → (Ⅎ𝑥(𝜑 → 𝜓) ↔ Ⅎ𝑥𝜓)) | ||
Theorem | wl-nfnbi 33444 | Being free does not depend on an outside negation in an expression. This theorem is slightly more general than nfn 1824 or nfnd 1825. (Contributed by Wolf Lammen, 5-May-2018.) |
⊢ (Ⅎ𝑥𝜑 ↔ Ⅎ𝑥 ¬ 𝜑) | ||
Theorem | wl-nfae1 33445 | Unlike nfae 2349, this specialized theorem avoids ax-11 2074. (Contributed by Wolf Lammen, 26-Jun-2019.) |
⊢ Ⅎ𝑥∀𝑦 𝑦 = 𝑥 | ||
Theorem | wl-nfnae1 33446 | Unlike nfnae 2351, this specialized theorem avoids ax-11 2074. (Contributed by Wolf Lammen, 27-Jun-2019.) |
⊢ Ⅎ𝑥 ¬ ∀𝑦 𝑦 = 𝑥 | ||
Theorem | wl-aetr 33447 | A transitive law for variable identifying expressions. (Contributed by Wolf Lammen, 30-Jun-2019.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (∀𝑥 𝑥 = 𝑧 → ∀𝑦 𝑦 = 𝑧)) | ||
Theorem | wl-dral1d 33448 | A version of dral1 2356 with a context. Note: At first glance one might be tempted to generalize this (or a similar) theorem by weakening the first two hypotheses adding a 𝑥 = 𝑦, ∀𝑥𝑥 = 𝑦 or 𝜑 antecedent. wl-equsal1i 33459 and nf5di 2157 show that this is in fact pointless. (Contributed by Wolf Lammen, 28-Jul-2019.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → (∀𝑥 𝑥 = 𝑦 → (∀𝑥𝜓 ↔ ∀𝑦𝜒))) | ||
Theorem | wl-cbvalnaed 33449 | wl-cbvalnae 33450 with a context. (Contributed by Wolf Lammen, 28-Jul-2019.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → (¬ ∀𝑥 𝑥 = 𝑦 → Ⅎ𝑦𝜓)) & ⊢ (𝜑 → (¬ ∀𝑥 𝑥 = 𝑦 → Ⅎ𝑥𝜒)) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 ↔ ∀𝑦𝜒)) | ||
Theorem | wl-cbvalnae 33450 | A more general version of cbval 2307 when non-free properties depend on a distinctor. Such expressions arise in proofs aiming at the elimination of distinct variable constraints, specifically in application of dvelimf 2365, nfsb2 2388 or dveeq1 2336. (Contributed by Wolf Lammen, 4-Jun-2019.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → Ⅎ𝑦𝜑) & ⊢ (¬ ∀𝑥 𝑥 = 𝑦 → Ⅎ𝑥𝜓) & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥𝜑 ↔ ∀𝑦𝜓) | ||
Theorem | wl-exeq 33451 | The semantics of ∃𝑥𝑦 = 𝑧. (Contributed by Wolf Lammen, 27-Apr-2018.) |
⊢ (∃𝑥 𝑦 = 𝑧 ↔ (𝑦 = 𝑧 ∨ ∀𝑥 𝑥 = 𝑦 ∨ ∀𝑥 𝑥 = 𝑧)) | ||
Theorem | wl-aleq 33452 | The semantics of ∀𝑥𝑦 = 𝑧. (Contributed by Wolf Lammen, 27-Apr-2018.) |
⊢ (∀𝑥 𝑦 = 𝑧 ↔ (𝑦 = 𝑧 ∧ (∀𝑥 𝑥 = 𝑦 ↔ ∀𝑥 𝑥 = 𝑧))) | ||
Theorem | wl-nfeqfb 33453 | Extend nfeqf 2337 to an equivalence. (Contributed by Wolf Lammen, 31-Jul-2019.) |
⊢ (Ⅎ𝑥 𝑦 = 𝑧 ↔ (∀𝑥 𝑥 = 𝑦 ↔ ∀𝑥 𝑥 = 𝑧)) | ||
Theorem | wl-nfs1t 33454 | If 𝑦 is not free in 𝜑, 𝑥 is not free in [𝑦 / 𝑥]𝜑. Closed form of nfs1 2393. (Contributed by Wolf Lammen, 27-Jul-2019.) |
⊢ (Ⅎ𝑦𝜑 → Ⅎ𝑥[𝑦 / 𝑥]𝜑) | ||
Theorem | wl-equsald 33455 | Deduction version of equsal 2327. (Contributed by Wolf Lammen, 27-Jul-2019.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜒) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → (∀𝑥(𝑥 = 𝑦 → 𝜓) ↔ 𝜒)) | ||
Theorem | wl-equsal 33456 | A useful equivalence related to substitution. (Contributed by NM, 2-Jun-1993.) (Proof shortened by Andrew Salmon, 12-Aug-2011.) (Revised by Mario Carneiro, 3-Oct-2016.) It seems proving wl-equsald 33455 first, and then deriving more specialized versions wl-equsal 33456 and wl-equsal1t 33457 then is more efficient than the other way round, which is possible, too. See also equsal 2327. (Revised by Wolf Lammen, 27-Jul-2019.) (Proof modification is discouraged.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥(𝑥 = 𝑦 → 𝜑) ↔ 𝜓) | ||
Theorem | wl-equsal1t 33457 |
The expression 𝑥 = 𝑦 in antecedent position plays an
important role in
predicate logic, namely in implicit substitution. However, occasionally
it is irrelevant, and can safely be dropped. A sufficient condition for
this is when 𝑥 (or 𝑦 or both) is not free in
𝜑.
This theorem is more fundamental than equsal 2327, spimt 2289 or sbft 2407, to which it is related. (Contributed by Wolf Lammen, 19-Aug-2018.) |
⊢ (Ⅎ𝑥𝜑 → (∀𝑥(𝑥 = 𝑦 → 𝜑) ↔ 𝜑)) | ||
Theorem | wl-equsalcom 33458 | This simple equivalence eases substitution of one expression for the other. (Contributed by Wolf Lammen, 1-Sep-2018.) |
⊢ (∀𝑥(𝑥 = 𝑦 → 𝜑) ↔ ∀𝑥(𝑦 = 𝑥 → 𝜑)) | ||
Theorem | wl-equsal1i 33459 | The antecedent 𝑥 = 𝑦 is irrelevant, if one or both setvar variables are not free in 𝜑. (Contributed by Wolf Lammen, 1-Sep-2018.) |
⊢ (Ⅎ𝑥𝜑 ∨ Ⅎ𝑦𝜑) & ⊢ (𝑥 = 𝑦 → 𝜑) ⇒ ⊢ 𝜑 | ||
Theorem | wl-sb6rft 33460 | A specialization of wl-equsal1t 33457. Closed form of sb6rf 2451. (Contributed by Wolf Lammen, 27-Jul-2019.) |
⊢ (Ⅎ𝑥𝜑 → (𝜑 ↔ ∀𝑥(𝑥 = 𝑦 → [𝑥 / 𝑦]𝜑))) | ||
Theorem | wl-sbrimt 33461 | Substitution with a variable not free in antecedent affects only the consequent. Closed form of sbrim 2424. (Contributed by Wolf Lammen, 26-Jul-2019.) |
⊢ (Ⅎ𝑥𝜑 → ([𝑦 / 𝑥](𝜑 → 𝜓) ↔ (𝜑 → [𝑦 / 𝑥]𝜓))) | ||
Theorem | wl-sblimt 33462 | Substitution with a variable not free in antecedent affects only the consequent. Closed form of sbrim 2424. (Contributed by Wolf Lammen, 26-Jul-2019.) |
⊢ (Ⅎ𝑥𝜓 → ([𝑦 / 𝑥](𝜑 → 𝜓) ↔ ([𝑦 / 𝑥]𝜑 → 𝜓))) | ||
Theorem | wl-sb8t 33463 | Substitution of variable in universal quantifier. Closed form of sb8 2452. (Contributed by Wolf Lammen, 27-Jul-2019.) |
⊢ (∀𝑥Ⅎ𝑦𝜑 → (∀𝑥𝜑 ↔ ∀𝑦[𝑦 / 𝑥]𝜑)) | ||
Theorem | wl-sb8et 33464 | Substitution of variable in universal quantifier. Closed form of sb8e 2453. (Contributed by Wolf Lammen, 27-Jul-2019.) |
⊢ (∀𝑥Ⅎ𝑦𝜑 → (∃𝑥𝜑 ↔ ∃𝑦[𝑦 / 𝑥]𝜑)) | ||
Theorem | wl-sbhbt 33465 | Closed form of sbhb 2466. Characterizing the expression 𝜑 → ∀𝑥𝜑 using a substitution expression. (Contributed by Wolf Lammen, 28-Jul-2019.) |
⊢ (∀𝑥Ⅎ𝑦𝜑 → ((𝜑 → ∀𝑥𝜑) ↔ ∀𝑦(𝜑 → [𝑦 / 𝑥]𝜑))) | ||
Theorem | wl-sbnf1 33466 | Two ways expressing that 𝑥 is effectively not free in 𝜑. Simplified version of sbnf2 2467. Note: This theorem shows that sbnf2 2467 has unnecessary distinct variable constraints. (Contributed by Wolf Lammen, 28-Jul-2019.) |
⊢ (∀𝑥Ⅎ𝑦𝜑 → (Ⅎ𝑥𝜑 ↔ ∀𝑥∀𝑦(𝜑 → [𝑦 / 𝑥]𝜑))) | ||
Theorem | wl-equsb3 33467 | equsb3 2460 with a distinctor. (Contributed by Wolf Lammen, 27-Jun-2019.) |
⊢ (¬ ∀𝑦 𝑦 = 𝑧 → ([𝑥 / 𝑦]𝑦 = 𝑧 ↔ 𝑥 = 𝑧)) | ||
Theorem | wl-equsb4 33468 | Substitution applied to an atomic wff. The distinctor antecedent is more general than a distinct variable constraint. (Contributed by Wolf Lammen, 26-Jun-2019.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑧 → ([𝑦 / 𝑥]𝑦 = 𝑧 ↔ 𝑦 = 𝑧)) | ||
Theorem | wl-sb6nae 33469 | Version of sb6 2457 suitable for elimination of unnecessary dv restrictions. (Contributed by Wolf Lammen, 28-Jul-2019.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → ([𝑦 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | wl-sb5nae 33470 | Version of sb5 2458 suitable for elimination of unnecessary dv restrictions. (Contributed by Wolf Lammen, 28-Jul-2019.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → ([𝑦 / 𝑥]𝜑 ↔ ∃𝑥(𝑥 = 𝑦 ∧ 𝜑))) | ||
Theorem | wl-2sb6d 33471 | Version of 2sb6 2472 with a context, and distinct variable conditions replaced with distinctors. (Contributed by Wolf Lammen, 4-Aug-2019.) |
⊢ (𝜑 → ¬ ∀𝑦 𝑦 = 𝑥) & ⊢ (𝜑 → ¬ ∀𝑦 𝑦 = 𝑤) & ⊢ (𝜑 → ¬ ∀𝑦 𝑦 = 𝑧) & ⊢ (𝜑 → ¬ ∀𝑥 𝑥 = 𝑧) ⇒ ⊢ (𝜑 → ([𝑧 / 𝑥][𝑤 / 𝑦]𝜓 ↔ ∀𝑥∀𝑦((𝑥 = 𝑧 ∧ 𝑦 = 𝑤) → 𝜓))) | ||
Theorem | wl-sbcom2d-lem1 33472* | Lemma used to prove wl-sbcom2d 33474. (Contributed by Wolf Lammen, 10-Aug-2019.) (New usage is discouraged.) |
⊢ ((𝑢 = 𝑦 ∧ 𝑣 = 𝑤) → (¬ ∀𝑥 𝑥 = 𝑤 → ([𝑢 / 𝑥][𝑣 / 𝑧]𝜑 ↔ [𝑦 / 𝑥][𝑤 / 𝑧]𝜑))) | ||
Theorem | wl-sbcom2d-lem2 33473* | Lemma used to prove wl-sbcom2d 33474. (Contributed by Wolf Lammen, 10-Aug-2019.) (New usage is discouraged.) |
⊢ (¬ ∀𝑦 𝑦 = 𝑥 → ([𝑢 / 𝑥][𝑣 / 𝑦]𝜑 ↔ ∀𝑥∀𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) → 𝜑))) | ||
Theorem | wl-sbcom2d 33474 | Version of sbcom2 2473 with a context, and distinct variable conditions replaced with distinctors. (Contributed by Wolf Lammen, 4-Aug-2019.) |
⊢ (𝜑 → ¬ ∀𝑥 𝑥 = 𝑤) & ⊢ (𝜑 → ¬ ∀𝑥 𝑥 = 𝑧) & ⊢ (𝜑 → ¬ ∀𝑧 𝑧 = 𝑦) ⇒ ⊢ (𝜑 → ([𝑤 / 𝑧][𝑦 / 𝑥]𝜓 ↔ [𝑦 / 𝑥][𝑤 / 𝑧]𝜓)) | ||
Theorem | wl-sbalnae 33475 | A theorem used in elimination of disjoint variable restrictions by replacing them with distinctors. (Contributed by Wolf Lammen, 25-Jul-2019.) |
⊢ ((¬ ∀𝑥 𝑥 = 𝑦 ∧ ¬ ∀𝑥 𝑥 = 𝑧) → ([𝑧 / 𝑦]∀𝑥𝜑 ↔ ∀𝑥[𝑧 / 𝑦]𝜑)) | ||
Theorem | wl-sbal1 33476* | A theorem used in elimination of disjoint variable restriction on 𝑥 and 𝑦 by replacing it with a distinctor ¬ ∀𝑥𝑥 = 𝑧. (Contributed by NM, 15-May-1993.) Proof is based on wl-sbalnae 33475 now. See also sbal1 2488. (Revised by Wolf Lammen, 25-Jul-2019.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑧 → ([𝑧 / 𝑦]∀𝑥𝜑 ↔ ∀𝑥[𝑧 / 𝑦]𝜑)) | ||
Theorem | wl-sbal2 33477* | Move quantifier in and out of substitution. Revised to remove a distinct variable constraint. (Contributed by NM, 2-Jan-2002.) Proof is based on wl-sbalnae 33475 now. See also sbal2 2489. (Revised by Wolf Lammen, 25-Jul-2019.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → ([𝑧 / 𝑦]∀𝑥𝜑 ↔ ∀𝑥[𝑧 / 𝑦]𝜑)) | ||
Theorem | wl-lem-exsb 33478* | This theorem provides a basic working step in proving theorems about ∃* or ∃!. (Contributed by Wolf Lammen, 3-Oct-2019.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Theorem | wl-lem-nexmo 33479 | This theorem provides a basic working step in proving theorems about ∃* or ∃!. (Contributed by Wolf Lammen, 3-Oct-2019.) |
⊢ (¬ ∃𝑥𝜑 → ∀𝑥(𝜑 → 𝑥 = 𝑧)) | ||
Theorem | wl-lem-moexsb 33480* |
The antecedent ∀𝑥(𝜑 → 𝑥 = 𝑧) relates to ∃*𝑥𝜑, but is
better suited for usage in proofs. Note that no distinct variable
restriction is placed on 𝜑.
This theorem provides a basic working step in proving theorems about ∃* or ∃!. (Contributed by Wolf Lammen, 3-Oct-2019.) |
⊢ (∀𝑥(𝜑 → 𝑥 = 𝑧) → (∃𝑥𝜑 ↔ [𝑧 / 𝑥]𝜑)) | ||
Theorem | wl-alanbii 33481 | This theorem extends alanimi 1784 to a biconditional. Recurrent usage stacks up more quantifiers. (Contributed by Wolf Lammen, 4-Oct-2019.) |
⊢ (𝜑 ↔ (𝜓 ∧ 𝜒)) ⇒ ⊢ (∀𝑥𝜑 ↔ (∀𝑥𝜓 ∧ ∀𝑥𝜒)) | ||
Theorem | wl-mo2df 33482 | Version of mo2 2507 with a context and a distinctor replacing a distinct variable condition. This version should be used only to eliminate dv conditions. (Contributed by Wolf Lammen, 11-Aug-2019.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → ¬ ∀𝑥 𝑥 = 𝑦) & ⊢ (𝜑 → Ⅎ𝑦𝜓) ⇒ ⊢ (𝜑 → (∃*𝑥𝜓 ↔ ∃𝑦∀𝑥(𝜓 → 𝑥 = 𝑦))) | ||
Theorem | wl-mo2tf 33483 | Closed form of mo2 2507 with a distinctor avoiding distinct variable conditions. (Contributed by Wolf Lammen, 20-Sep-2020.) |
⊢ ((¬ ∀𝑥 𝑥 = 𝑦 ∧ ∀𝑥Ⅎ𝑦𝜑) → (∃*𝑥𝜑 ↔ ∃𝑦∀𝑥(𝜑 → 𝑥 = 𝑦))) | ||
Theorem | wl-eudf 33484 | Version of df-eu 2502 with a context and a distinctor replacing a distinct variable condition. This version should be used only to eliminate dv conditions. (Contributed by Wolf Lammen, 23-Sep-2020.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → ¬ ∀𝑥 𝑥 = 𝑦) & ⊢ (𝜑 → Ⅎ𝑦𝜓) ⇒ ⊢ (𝜑 → (∃!𝑥𝜓 ↔ ∃𝑦∀𝑥(𝜓 ↔ 𝑥 = 𝑦))) | ||
Theorem | wl-eutf 33485 | Closed form of df-eu 2502 with a distinctor avoiding distinct variable conditions. (Contributed by Wolf Lammen, 23-Sep-2020.) |
⊢ ((¬ ∀𝑥 𝑥 = 𝑦 ∧ ∀𝑥Ⅎ𝑦𝜑) → (∃!𝑥𝜑 ↔ ∃𝑦∀𝑥(𝜑 ↔ 𝑥 = 𝑦))) | ||
Theorem | wl-euequ1f 33486 | euequ1 2504 proved with a distinctor. (Contributed by Wolf Lammen, 23-Sep-2020.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → ∃!𝑥 𝑥 = 𝑦) | ||
Theorem | wl-mo2t 33487* | Closed form of mo2 2507. (Contributed by Wolf Lammen, 18-Aug-2019.) |
⊢ (∀𝑥Ⅎ𝑦𝜑 → (∃*𝑥𝜑 ↔ ∃𝑦∀𝑥(𝜑 → 𝑥 = 𝑦))) | ||
Theorem | wl-mo3t 33488* | Closed form of mo3 2536. (Contributed by Wolf Lammen, 18-Aug-2019.) |
⊢ (∀𝑥Ⅎ𝑦𝜑 → (∃*𝑥𝜑 ↔ ∀𝑥∀𝑦((𝜑 ∧ [𝑦 / 𝑥]𝜑) → 𝑥 = 𝑦))) | ||
Theorem | wl-sb8eut 33489 | Substitution of variable in universal quantifier. Closed form of sb8eu 2532. (Contributed by Wolf Lammen, 11-Aug-2019.) |
⊢ (∀𝑥Ⅎ𝑦𝜑 → (∃!𝑥𝜑 ↔ ∃!𝑦[𝑦 / 𝑥]𝜑)) | ||
Theorem | wl-sb8mot 33490 |
Substitution of variable in universal quantifier. Closed form of
sb8mo 2533.
This theorem relates to wl-mo3t 33488, since replacing 𝜑 with [𝑦 / 𝑥]𝜑 in the latter yields subexpressions like [𝑥 / 𝑦][𝑦 / 𝑥]𝜑, which can be reduced to 𝜑 via sbft 2407 and sbco 2440. So ∃*𝑥𝜑 ↔ ∃*𝑦[𝑦 / 𝑥]𝜑 is provable from wl-mo3t 33488 in a simple fashion, unfortunately only if 𝑥 and 𝑦 are known to be distinct. To avoid any hassle with distinctors, we prefer to derive this theorem independently, ignoring the close connection between both theorems. From an educational standpoint, one would assume wl-mo3t 33488 to be more fundamental, as it hints how the "at most one" objects on both sides of the biconditional correlate (they are the same), if they exist at all, and then prove this theorem from it. (Contributed by Wolf Lammen, 11-Aug-2019.) |
⊢ (∀𝑥Ⅎ𝑦𝜑 → (∃*𝑥𝜑 ↔ ∃*𝑦[𝑦 / 𝑥]𝜑)) | ||
Axiom | ax-wl-11v 33491* | Version of ax-11 2074 with distinct variable conditions. Currently implemented as an axiom to detect unintended references to the foundational axiom ax-11 2074. It will later be converted into a theorem directly based on ax-11 2074. (Contributed by Wolf Lammen, 28-Jun-2019.) |
⊢ (∀𝑥∀𝑦𝜑 → ∀𝑦∀𝑥𝜑) | ||
Theorem | wl-ax11-lem1 33492 | A transitive law for variable identifying expressions. (Contributed by Wolf Lammen, 30-Jun-2019.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (∀𝑥 𝑥 = 𝑧 ↔ ∀𝑦 𝑦 = 𝑧)) | ||
Theorem | wl-ax11-lem2 33493* | Lemma. (Contributed by Wolf Lammen, 30-Jun-2019.) |
⊢ ((∀𝑢 𝑢 = 𝑦 ∧ ¬ ∀𝑥 𝑥 = 𝑦) → ∀𝑥 𝑢 = 𝑦) | ||
Theorem | wl-ax11-lem3 33494* | Lemma. (Contributed by Wolf Lammen, 30-Jun-2019.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → Ⅎ𝑥∀𝑢 𝑢 = 𝑦) | ||
Theorem | wl-ax11-lem4 33495* | Lemma. (Contributed by Wolf Lammen, 30-Jun-2019.) |
⊢ Ⅎ𝑥(∀𝑢 𝑢 = 𝑦 ∧ ¬ ∀𝑥 𝑥 = 𝑦) | ||
Theorem | wl-ax11-lem5 33496 | Lemma. (Contributed by Wolf Lammen, 30-Jun-2019.) |
⊢ (∀𝑢 𝑢 = 𝑦 → (∀𝑢[𝑢 / 𝑦]𝜑 ↔ ∀𝑦𝜑)) | ||
Theorem | wl-ax11-lem6 33497* | Lemma. (Contributed by Wolf Lammen, 30-Jun-2019.) |
⊢ ((∀𝑢 𝑢 = 𝑦 ∧ ¬ ∀𝑥 𝑥 = 𝑦) → (∀𝑢∀𝑥[𝑢 / 𝑦]𝜑 ↔ ∀𝑥∀𝑦𝜑)) | ||
Theorem | wl-ax11-lem7 33498 | Lemma. (Contributed by Wolf Lammen, 30-Jun-2019.) |
⊢ (∀𝑥(¬ ∀𝑥 𝑥 = 𝑦 ∧ 𝜑) ↔ (¬ ∀𝑥 𝑥 = 𝑦 ∧ ∀𝑥𝜑)) | ||
Theorem | wl-ax11-lem8 33499* | Lemma. (Contributed by Wolf Lammen, 30-Jun-2019.) |
⊢ ((∀𝑢 𝑢 = 𝑦 ∧ ¬ ∀𝑥 𝑥 = 𝑦) → (∀𝑢∀𝑥[𝑢 / 𝑦]𝜑 ↔ ∀𝑦∀𝑥𝜑)) | ||
Theorem | wl-ax11-lem9 33500 | The easy part when 𝑥 coincides with 𝑦. (Contributed by Wolf Lammen, 30-Jun-2019.) |
⊢ (∀𝑥 𝑥 = 𝑦 → (∀𝑦∀𝑥𝜑 ↔ ∀𝑥∀𝑦𝜑)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |