![]() |
Intuitionistic Logic Explorer Theorem List (p. 15 of 153) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > ILE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | xor2dc 1401 | Two ways to express "exclusive or" between decidable propositions. (Contributed by Jim Kingdon, 17-Apr-2018.) |
⊢ (DECID 𝜑 → (DECID 𝜓 → (¬ (𝜑 ↔ 𝜓) ↔ ((𝜑 ∨ 𝜓) ∧ ¬ (𝜑 ∧ 𝜓))))) | ||
Theorem | xornbidc 1402 | Exclusive or is equivalent to negated biconditional for decidable propositions. (Contributed by Jim Kingdon, 27-Apr-2018.) |
⊢ (DECID 𝜑 → (DECID 𝜓 → ((𝜑 ⊻ 𝜓) ↔ ¬ (𝜑 ↔ 𝜓)))) | ||
Theorem | xordc 1403 | Two ways to express "exclusive or" between decidable propositions. Theorem *5.22 of [WhiteheadRussell] p. 124, but for decidable propositions. (Contributed by Jim Kingdon, 5-May-2018.) |
⊢ (DECID 𝜑 → (DECID 𝜓 → (¬ (𝜑 ↔ 𝜓) ↔ ((𝜑 ∧ ¬ 𝜓) ∨ (𝜓 ∧ ¬ 𝜑))))) | ||
Theorem | xordc1 1404 | Exclusive or implies the left proposition is decidable. (Contributed by Jim Kingdon, 12-Mar-2018.) |
⊢ ((𝜑 ⊻ 𝜓) → DECID 𝜑) | ||
Theorem | nbbndc 1405 | Move negation outside of biconditional, for decidable propositions. Compare Theorem *5.18 of [WhiteheadRussell] p. 124. (Contributed by Jim Kingdon, 18-Apr-2018.) |
⊢ (DECID 𝜑 → (DECID 𝜓 → ((¬ 𝜑 ↔ 𝜓) ↔ ¬ (𝜑 ↔ 𝜓)))) | ||
Theorem | biassdc 1406 |
Associative law for the biconditional, for decidable propositions.
The classical version (without the decidability conditions) is an axiom of system DS in Vladimir Lifschitz, "On calculational proofs", Annals of Pure and Applied Logic, 113:207-224, 2002, http://www.cs.utexas.edu/users/ai-lab/pub-view.php?PubID=26805, and, interestingly, was not included in Principia Mathematica but was apparently first noted by Jan Lukasiewicz circa 1923. (Contributed by Jim Kingdon, 4-May-2018.) |
⊢ (DECID 𝜑 → (DECID 𝜓 → (DECID 𝜒 → (((𝜑 ↔ 𝜓) ↔ 𝜒) ↔ (𝜑 ↔ (𝜓 ↔ 𝜒)))))) | ||
Theorem | bilukdc 1407 | Lukasiewicz's shortest axiom for equivalential calculus (but modified to require decidable propositions). Storrs McCall, ed., Polish Logic 1920-1939 (Oxford, 1967), p. 96. (Contributed by Jim Kingdon, 5-May-2018.) |
⊢ (((DECID 𝜑 ∧ DECID 𝜓) ∧ DECID 𝜒) → ((𝜑 ↔ 𝜓) ↔ ((𝜒 ↔ 𝜓) ↔ (𝜑 ↔ 𝜒)))) | ||
Theorem | dfbi3dc 1408 | An alternate definition of the biconditional for decidable propositions. Theorem *5.23 of [WhiteheadRussell] p. 124, but with decidability conditions. (Contributed by Jim Kingdon, 5-May-2018.) |
⊢ (DECID 𝜑 → (DECID 𝜓 → ((𝜑 ↔ 𝜓) ↔ ((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ ¬ 𝜓))))) | ||
Theorem | pm5.24dc 1409 | Theorem *5.24 of [WhiteheadRussell] p. 124, but for decidable propositions. (Contributed by Jim Kingdon, 5-May-2018.) |
⊢ (DECID 𝜑 → (DECID 𝜓 → (¬ ((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ ¬ 𝜓)) ↔ ((𝜑 ∧ ¬ 𝜓) ∨ (𝜓 ∧ ¬ 𝜑))))) | ||
Theorem | xordidc 1410 | Conjunction distributes over exclusive-or, for decidable propositions. This is one way to interpret the distributive law of multiplication over addition in modulo 2 arithmetic. (Contributed by Jim Kingdon, 14-Jul-2018.) |
⊢ (DECID 𝜑 → (DECID 𝜓 → (DECID 𝜒 → ((𝜑 ∧ (𝜓 ⊻ 𝜒)) ↔ ((𝜑 ∧ 𝜓) ⊻ (𝜑 ∧ 𝜒)))))) | ||
Theorem | anxordi 1411 | Conjunction distributes over exclusive-or. (Contributed by Mario Carneiro and Jim Kingdon, 7-Oct-2018.) |
⊢ ((𝜑 ∧ (𝜓 ⊻ 𝜒)) ↔ ((𝜑 ∧ 𝜓) ⊻ (𝜑 ∧ 𝜒))) | ||
For classical logic, truth tables can be used to define propositional logic operations, by showing the results of those operations for all possible combinations of true (⊤) and false (⊥). Although the intuitionistic logic connectives are not as simply defined, ⊤ and ⊥ do play similar roles as in classical logic and most theorems from classical logic continue to hold. Here we show that our definitions and axioms produce equivalent results for ⊤ and ⊥ as we would get from truth tables for ∧ (conjunction aka logical 'and') wa 104, ∨ (disjunction aka logical inclusive 'or') wo 709, → (implies) wi 4, ¬ (not) wn 3, ↔ (logical equivalence) df-bi 117, and ⊻ (exclusive or) df-xor 1387. | ||
Theorem | truantru 1412 | A ∧ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ ((⊤ ∧ ⊤) ↔ ⊤) | ||
Theorem | truanfal 1413 | A ∧ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ ((⊤ ∧ ⊥) ↔ ⊥) | ||
Theorem | falantru 1414 | A ∧ identity. (Contributed by David A. Wheeler, 23-Feb-2018.) |
⊢ ((⊥ ∧ ⊤) ↔ ⊥) | ||
Theorem | falanfal 1415 | A ∧ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ ((⊥ ∧ ⊥) ↔ ⊥) | ||
Theorem | truortru 1416 | A ∨ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ ((⊤ ∨ ⊤) ↔ ⊤) | ||
Theorem | truorfal 1417 | A ∨ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ ((⊤ ∨ ⊥) ↔ ⊤) | ||
Theorem | falortru 1418 | A ∨ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ ((⊥ ∨ ⊤) ↔ ⊤) | ||
Theorem | falorfal 1419 | A ∨ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ ((⊥ ∨ ⊥) ↔ ⊥) | ||
Theorem | truimtru 1420 | A → identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ ((⊤ → ⊤) ↔ ⊤) | ||
Theorem | truimfal 1421 | A → identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ ((⊤ → ⊥) ↔ ⊥) | ||
Theorem | falimtru 1422 | A → identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ ((⊥ → ⊤) ↔ ⊤) | ||
Theorem | falimfal 1423 | A → identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ ((⊥ → ⊥) ↔ ⊤) | ||
Theorem | nottru 1424 | A ¬ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (¬ ⊤ ↔ ⊥) | ||
Theorem | notfal 1425 | A ¬ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (¬ ⊥ ↔ ⊤) | ||
Theorem | trubitru 1426 | A ↔ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ ((⊤ ↔ ⊤) ↔ ⊤) | ||
Theorem | trubifal 1427 | A ↔ identity. (Contributed by David A. Wheeler, 23-Feb-2018.) |
⊢ ((⊤ ↔ ⊥) ↔ ⊥) | ||
Theorem | falbitru 1428 | A ↔ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ ((⊥ ↔ ⊤) ↔ ⊥) | ||
Theorem | falbifal 1429 | A ↔ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ ((⊥ ↔ ⊥) ↔ ⊤) | ||
Theorem | truxortru 1430 | A ⊻ identity. (Contributed by David A. Wheeler, 2-Mar-2018.) |
⊢ ((⊤ ⊻ ⊤) ↔ ⊥) | ||
Theorem | truxorfal 1431 | A ⊻ identity. (Contributed by David A. Wheeler, 2-Mar-2018.) |
⊢ ((⊤ ⊻ ⊥) ↔ ⊤) | ||
Theorem | falxortru 1432 | A ⊻ identity. (Contributed by David A. Wheeler, 2-Mar-2018.) |
⊢ ((⊥ ⊻ ⊤) ↔ ⊤) | ||
Theorem | falxorfal 1433 | A ⊻ identity. (Contributed by David A. Wheeler, 2-Mar-2018.) |
⊢ ((⊥ ⊻ ⊥) ↔ ⊥) | ||
The Greek Stoics developed a system of logic. The Stoic Chrysippus, in particular, was often considered one of the greatest logicians of antiquity. Stoic logic is different from Aristotle's system, since it focuses on propositional logic, though later thinkers did combine the systems of the Stoics with Aristotle. Jan Lukasiewicz reports, "For anybody familiar with mathematical logic it is self-evident that the Stoic dialectic is the ancient form of modern propositional logic" ( On the history of the logic of proposition by Jan Lukasiewicz (1934), translated in: Selected Works - Edited by Ludwik Borkowski - Amsterdam, North-Holland, 1970 pp. 197-217, referenced in "History of Logic" https://www.historyoflogic.com/logic-stoics.htm). For more about Aristotle's system, see barbara and related theorems. A key part of the Stoic logic system is a set of five "indemonstrables" assigned to Chrysippus of Soli by Diogenes Laertius, though in general it is difficult to assign specific ideas to specific thinkers. The indemonstrables are described in, for example, [Lopez-Astorga] p. 11 , [Sanford] p. 39, and [Hitchcock] p. 5. These indemonstrables are modus ponendo ponens (modus ponens) ax-mp 5, modus tollendo tollens (modus tollens) mto 663, modus ponendo tollens I mptnan 1434, modus ponendo tollens II mptxor 1435, and modus tollendo ponens (exclusive-or version) mtpxor 1437. The first is an axiom, the second is already proved; in this section we prove the other three. Since we assume or prove all of indemonstrables, the system of logic we use here is as at least as strong as the set of Stoic indemonstrables. Note that modus tollendo ponens mtpxor 1437 originally used exclusive-or, but over time the name modus tollendo ponens has increasingly referred to an inclusive-or variation, which is proved in mtpor 1436. This set of indemonstrables is not the entire system of Stoic logic. | ||
Theorem | mptnan 1434 | Modus ponendo tollens 1, one of the "indemonstrables" in Stoic logic. See rule 1 on [Lopez-Astorga] p. 12 , rule 1 on [Sanford] p. 40, and rule A3 in [Hitchcock] p. 5. Sanford describes this rule second (after mptxor 1435) as a "safer, and these days much more common" version of modus ponendo tollens because it avoids confusion between inclusive-or and exclusive-or. (Contributed by David A. Wheeler, 3-Jul-2016.) |
⊢ 𝜑 & ⊢ ¬ (𝜑 ∧ 𝜓) ⇒ ⊢ ¬ 𝜓 | ||
Theorem | mptxor 1435 | Modus ponendo tollens 2, one of the "indemonstrables" in Stoic logic. Note that this uses exclusive-or ⊻. See rule 2 on [Lopez-Astorga] p. 12 , rule 4 on [Sanford] p. 39 and rule A4 in [Hitchcock] p. 5 . (Contributed by David A. Wheeler, 2-Mar-2018.) |
⊢ 𝜑 & ⊢ (𝜑 ⊻ 𝜓) ⇒ ⊢ ¬ 𝜓 | ||
Theorem | mtpor 1436 | Modus tollendo ponens (inclusive-or version), aka disjunctive syllogism. This is similar to mtpxor 1437, one of the five original "indemonstrables" in Stoic logic. However, in Stoic logic this rule used exclusive-or, while the name modus tollendo ponens often refers to a variant of the rule that uses inclusive-or instead. The rule says, "if 𝜑 is not true, and 𝜑 or 𝜓 (or both) are true, then 𝜓 must be true". An alternate phrasing is, "Once you eliminate the impossible, whatever remains, no matter how improbable, must be the truth". -- Sherlock Holmes (Sir Arthur Conan Doyle, 1890: The Sign of the Four, ch. 6). (Contributed by David A. Wheeler, 3-Jul-2016.) (Proof shortened by Wolf Lammen, 11-Nov-2017.) |
⊢ ¬ 𝜑 & ⊢ (𝜑 ∨ 𝜓) ⇒ ⊢ 𝜓 | ||
Theorem | mtpxor 1437 | Modus tollendo ponens (original exclusive-or version), aka disjunctive syllogism, similar to mtpor 1436, one of the five "indemonstrables" in Stoic logic. The rule says, "if 𝜑 is not true, and either 𝜑 or 𝜓 (exclusively) are true, then 𝜓 must be true". Today the name "modus tollendo ponens" often refers to a variant, the inclusive-or version as defined in mtpor 1436. See rule 3 on [Lopez-Astorga] p. 12 (note that the "or" is the same as mptxor 1435, that is, it is exclusive-or df-xor 1387), rule 3 of [Sanford] p. 39 (where it is not as clearly stated which kind of "or" is used but it appears to be in the same sense as mptxor 1435), and rule A5 in [Hitchcock] p. 5 (exclusive-or is expressly used). (Contributed by David A. Wheeler, 4-Jul-2016.) (Proof shortened by Wolf Lammen, 11-Nov-2017.) (Proof shortened by BJ, 19-Apr-2019.) |
⊢ ¬ 𝜑 & ⊢ (𝜑 ⊻ 𝜓) ⇒ ⊢ 𝜓 | ||
Theorem | stoic1a 1438 |
Stoic logic Thema 1 (part a).
The first thema of the four Stoic logic themata, in its basic form, was: "When from two (assertibles) a third follows, then from either of them together with the contradictory of the conclusion the contradictory of the other follows." (Apuleius Int. 209.9-14), see [Bobzien] p. 117 and https://plato.stanford.edu/entries/logic-ancient/ We will represent thema 1 as two very similar rules stoic1a 1438 and stoic1b 1439 to represent each side. (Contributed by David A. Wheeler, 16-Feb-2019.) (Proof shortened by Wolf Lammen, 21-May-2020.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜃) ⇒ ⊢ ((𝜑 ∧ ¬ 𝜃) → ¬ 𝜓) | ||
Theorem | stoic1b 1439 | Stoic logic Thema 1 (part b). The other part of thema 1 of Stoic logic; see stoic1a 1438. (Contributed by David A. Wheeler, 16-Feb-2019.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜃) ⇒ ⊢ ((𝜓 ∧ ¬ 𝜃) → ¬ 𝜑) | ||
Theorem | stoic2a 1440 |
Stoic logic Thema 2 version a.
Statement T2 of [Bobzien] p. 117 shows a reconstructed version of Stoic logic thema 2 as follows: "When from two assertibles a third follows, and from the third and one (or both) of the two another follows, then this other follows from the first two." Bobzien uses constructs such as 𝜑, 𝜓⊢ 𝜒; in Metamath we will represent that construct as 𝜑 ∧ 𝜓 → 𝜒. This version a is without the phrase "or both"; see stoic2b 1441 for the version with the phrase "or both". We already have this rule as syldan 282, so here we show the equivalence and discourage its use. (New usage is discouraged.) (Contributed by David A. Wheeler, 17-Feb-2019.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) & ⊢ ((𝜑 ∧ 𝜒) → 𝜃) ⇒ ⊢ ((𝜑 ∧ 𝜓) → 𝜃) | ||
Theorem | stoic2b 1441 |
Stoic logic Thema 2 version b. See stoic2a 1440.
Version b is with the phrase "or both". We already have this rule as mpd3an3 1349, so here we prove the equivalence and discourage its use. (New usage is discouraged.) (Contributed by David A. Wheeler, 17-Feb-2019.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) & ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) ⇒ ⊢ ((𝜑 ∧ 𝜓) → 𝜃) | ||
Theorem | stoic3 1442 |
Stoic logic Thema 3.
Statement T3 of [Bobzien] p. 116-117 discusses Stoic logic thema 3. "When from two (assemblies) a third follows, and from the one that follows (i.e., the third) together with another, external external assumption, another follows, then other follows from the first two and the externally co-assumed one. (Simp. Cael. 237.2-4)" (Contributed by David A. Wheeler, 17-Feb-2019.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) & ⊢ ((𝜒 ∧ 𝜃) → 𝜏) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜃) → 𝜏) | ||
Theorem | stoic4a 1443 |
Stoic logic Thema 4 version a.
Statement T4 of [Bobzien] p. 117 shows a reconstructed version of Stoic logic thema 4: "When from two assertibles a third follows, and from the third and one (or both) of the two and one (or more) external assertible(s) another follows, then this other follows from the first two and the external(s)." We use 𝜃 to represent the "external" assertibles. This is version a, which is without the phrase "or both"; see stoic4b 1444 for the version with the phrase "or both". (Contributed by David A. Wheeler, 17-Feb-2019.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) & ⊢ ((𝜒 ∧ 𝜑 ∧ 𝜃) → 𝜏) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜃) → 𝜏) | ||
Theorem | stoic4b 1444 |
Stoic logic Thema 4 version b.
This is version b, which is with the phrase "or both". See stoic4a 1443 for more information. (Contributed by David A. Wheeler, 17-Feb-2019.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) & ⊢ (((𝜒 ∧ 𝜑 ∧ 𝜓) ∧ 𝜃) → 𝜏) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜃) → 𝜏) | ||
Theorem | syl6an 1445 | A syllogism deduction combined with conjoining antecedents. (Contributed by Alan Sare, 28-Oct-2011.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜑 → (𝜒 → 𝜃)) & ⊢ ((𝜓 ∧ 𝜃) → 𝜏) ⇒ ⊢ (𝜑 → (𝜒 → 𝜏)) | ||
Theorem | syl10 1446 | A nested syllogism inference. (Contributed by Alan Sare, 17-Jul-2011.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜏))) & ⊢ (𝜒 → (𝜏 → 𝜂)) ⇒ ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜂))) | ||
Theorem | exbir 1447 | Exportation implication also converting head from biconditional to conditional. (Contributed by Alan Sare, 31-Dec-2011.) |
⊢ (((𝜑 ∧ 𝜓) → (𝜒 ↔ 𝜃)) → (𝜑 → (𝜓 → (𝜃 → 𝜒)))) | ||
Theorem | 3impexp 1448 | impexp 263 with a 3-conjunct antecedent. (Contributed by Alan Sare, 31-Dec-2011.) |
⊢ (((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) ↔ (𝜑 → (𝜓 → (𝜒 → 𝜃)))) | ||
Theorem | 3impexpbicom 1449 | 3impexp 1448 with biconditional consequent of antecedent that is commuted in consequent. (Contributed by Alan Sare, 31-Dec-2011.) |
⊢ (((𝜑 ∧ 𝜓 ∧ 𝜒) → (𝜃 ↔ 𝜏)) ↔ (𝜑 → (𝜓 → (𝜒 → (𝜏 ↔ 𝜃))))) | ||
Theorem | 3impexpbicomi 1450 | Deduction form of 3impexpbicom 1449. (Contributed by Alan Sare, 31-Dec-2011.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → (𝜃 ↔ 𝜏)) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 → (𝜏 ↔ 𝜃)))) | ||
Theorem | ancomsimp 1451 | Closed form of ancoms 268. (Contributed by Alan Sare, 31-Dec-2011.) |
⊢ (((𝜑 ∧ 𝜓) → 𝜒) ↔ ((𝜓 ∧ 𝜑) → 𝜒)) | ||
Theorem | expcomd 1452 | Deduction form of expcom 116. (Contributed by Alan Sare, 22-Jul-2012.) |
⊢ (𝜑 → ((𝜓 ∧ 𝜒) → 𝜃)) ⇒ ⊢ (𝜑 → (𝜒 → (𝜓 → 𝜃))) | ||
Theorem | expdcom 1453 | Commuted form of expd 258. (Contributed by Alan Sare, 18-Mar-2012.) |
⊢ (𝜑 → ((𝜓 ∧ 𝜒) → 𝜃)) ⇒ ⊢ (𝜓 → (𝜒 → (𝜑 → 𝜃))) | ||
Theorem | simplbi2comg 1454 | Implication form of simplbi2com 1455. (Contributed by Alan Sare, 22-Jul-2012.) |
⊢ ((𝜑 ↔ (𝜓 ∧ 𝜒)) → (𝜒 → (𝜓 → 𝜑))) | ||
Theorem | simplbi2com 1455 | A deduction eliminating a conjunct, similar to simplbi2 385. (Contributed by Alan Sare, 22-Jul-2012.) (Proof shortened by Wolf Lammen, 10-Nov-2012.) |
⊢ (𝜑 ↔ (𝜓 ∧ 𝜒)) ⇒ ⊢ (𝜒 → (𝜓 → 𝜑)) | ||
Theorem | syl6ci 1456 | A syllogism inference combined with contraction. (Contributed by Alan Sare, 18-Mar-2012.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → 𝜃) & ⊢ (𝜒 → (𝜃 → 𝜏)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜏)) | ||
Theorem | mpisyl 1457 | A syllogism combined with a modus ponens inference. (Contributed by Alan Sare, 25-Jul-2011.) |
⊢ (𝜑 → 𝜓) & ⊢ 𝜒 & ⊢ (𝜓 → (𝜒 → 𝜃)) ⇒ ⊢ (𝜑 → 𝜃) | ||
The universal quantifier was introduced above in wal 1362 for use by df-tru 1367. See the comments in that section. In this section, we continue with the first "real" use of it. | ||
Axiom | ax-5 1458 | Axiom of Quantified Implication. Axiom C4 of [Monk2] p. 105. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥(𝜑 → 𝜓) → (∀𝑥𝜑 → ∀𝑥𝜓)) | ||
Axiom | ax-7 1459 | Axiom of Quantifier Commutation. This axiom says universal quantifiers can be swapped. One of the predicate logic axioms which do not involve equality. 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. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥∀𝑦𝜑 → ∀𝑦∀𝑥𝜑) | ||
Axiom | ax-gen 1460 | Rule of Generalization. The postulated inference rule of predicate calculus. See, e.g., Rule 2 of [Hamilton] p. 74. This rule says that if something is unconditionally true, then it is true for all values of a variable. For example, if we have proved 𝑥 = 𝑥, we can conclude ∀𝑥𝑥 = 𝑥 or even ∀𝑦𝑥 = 𝑥. Theorem spi 1547 shows we can go the other way also: in other words we can add or remove universal quantifiers from the beginning of any theorem as required. (Contributed by NM, 5-Aug-1993.) |
⊢ 𝜑 ⇒ ⊢ ∀𝑥𝜑 | ||
Theorem | gen2 1461 | Generalization applied twice. (Contributed by NM, 30-Apr-1998.) |
⊢ 𝜑 ⇒ ⊢ ∀𝑥∀𝑦𝜑 | ||
Theorem | mpg 1462 | Modus ponens combined with generalization. (Contributed by NM, 24-May-1994.) |
⊢ (∀𝑥𝜑 → 𝜓) & ⊢ 𝜑 ⇒ ⊢ 𝜓 | ||
Theorem | mpgbi 1463 | Modus ponens on biconditional combined with generalization. (Contributed by NM, 24-May-1994.) (Proof shortened by Stefan Allan, 28-Oct-2008.) |
⊢ (∀𝑥𝜑 ↔ 𝜓) & ⊢ 𝜑 ⇒ ⊢ 𝜓 | ||
Theorem | mpgbir 1464 | Modus ponens on biconditional combined with generalization. (Contributed by NM, 24-May-1994.) (Proof shortened by Stefan Allan, 28-Oct-2008.) |
⊢ (𝜑 ↔ ∀𝑥𝜓) & ⊢ 𝜓 ⇒ ⊢ 𝜑 | ||
Theorem | a7s 1465 | Swap quantifiers in an antecedent. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥∀𝑦𝜑 → 𝜓) ⇒ ⊢ (∀𝑦∀𝑥𝜑 → 𝜓) | ||
Theorem | alimi 1466 | Inference quantifying both antecedent and consequent. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (∀𝑥𝜑 → ∀𝑥𝜓) | ||
Theorem | 2alimi 1467 | Inference doubly quantifying both antecedent and consequent. (Contributed by NM, 3-Feb-2005.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (∀𝑥∀𝑦𝜑 → ∀𝑥∀𝑦𝜓) | ||
Theorem | alim 1468 | Theorem 19.20 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) (Proof shortened by O'Cat, 30-Mar-2008.) |
⊢ (∀𝑥(𝜑 → 𝜓) → (∀𝑥𝜑 → ∀𝑥𝜓)) | ||
Theorem | al2imi 1469 | Inference quantifying antecedent, nested antecedent, and consequent. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (∀𝑥𝜑 → (∀𝑥𝜓 → ∀𝑥𝜒)) | ||
Theorem | alanimi 1470 | Variant of al2imi 1469 with conjunctive antecedent. (Contributed by Andrew Salmon, 8-Jun-2011.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) ⇒ ⊢ ((∀𝑥𝜑 ∧ ∀𝑥𝜓) → ∀𝑥𝜒) | ||
Syntax | wnf 1471 | Extend wff definition to include the not-free predicate. |
wff Ⅎ𝑥𝜑 | ||
Definition | df-nf 1472 |
Define the not-free predicate for wffs. This is read "𝑥 is not
free
in 𝜑". Not-free means that the
value of 𝑥 cannot affect the
value of 𝜑, e.g., any occurrence of 𝑥 in
𝜑 is
effectively
bound by a "for all" or something that expands to one (such as
"there
exists"). In particular, substitution for a variable not free in a
wff
does not affect its value (sbf 1788). An example of where this is used is
stdpc5 1595. See nf2 1679 for an alternate definition which
does not involve
nested quantifiers on the same variable.
Nonfreeness is a commonly used condition, so it is useful to have a notation for it. Surprisingly, there is no common formal notation for it, so here we devise one. Our definition lets us work with the notion of nonfreeness within the logic itself rather than as a metalogical side condition. To be precise, our definition really means "effectively not free", because it is slightly less restrictive than the usual textbook definition for "not free" (which considers syntactic freedom). For example, 𝑥 is effectively not free in the expression 𝑥 = 𝑥 (even though 𝑥 is syntactically free in it, so would be considered "free" in the usual textbook definition) because the value of 𝑥 in the formula 𝑥 = 𝑥 does not affect the truth of that formula (and thus substitutions will not change the result), see nfequid 1713. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (Ⅎ𝑥𝜑 ↔ ∀𝑥(𝜑 → ∀𝑥𝜑)) | ||
Theorem | nfi 1473 | Deduce that 𝑥 is not free in 𝜑 from the definition. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ Ⅎ𝑥𝜑 | ||
Theorem | hbth 1474 |
No variable is (effectively) free in a theorem.
This and later "hypothesis-building" lemmas, with labels starting "hb...", allow us to construct proofs of formulas of the form ⊢ (𝜑 → ∀𝑥𝜑) from smaller formulas of this form. These are useful for constructing hypotheses that state "𝑥 is (effectively) not free in 𝜑". (Contributed by NM, 5-Aug-1993.) |
⊢ 𝜑 ⇒ ⊢ (𝜑 → ∀𝑥𝜑) | ||
Theorem | nfth 1475 | No variable is (effectively) free in a theorem. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ 𝜑 ⇒ ⊢ Ⅎ𝑥𝜑 | ||
Theorem | nfnth 1476 | No variable is (effectively) free in a non-theorem. (Contributed by Mario Carneiro, 6-Dec-2016.) |
⊢ ¬ 𝜑 ⇒ ⊢ Ⅎ𝑥𝜑 | ||
Theorem | nftru 1477 | The true constant has no free variables. (This can also be proven in one step with nfv 1539, but this proof does not use ax-17 1537.) (Contributed by Mario Carneiro, 6-Oct-2016.) |
⊢ Ⅎ𝑥⊤ | ||
Theorem | alimdh 1478 | Deduction from Theorem 19.20 of [Margaris] p. 90. (Contributed by NM, 4-Jan-2002.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 → ∀𝑥𝜒)) | ||
Theorem | albi 1479 | Theorem 19.15 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥(𝜑 ↔ 𝜓) → (∀𝑥𝜑 ↔ ∀𝑥𝜓)) | ||
Theorem | alrimih 1480 | Inference from Theorem 19.21 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) (New usage is discouraged.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → 𝜓) ⇒ ⊢ (𝜑 → ∀𝑥𝜓) | ||
Theorem | albii 1481 | Inference adding universal quantifier to both sides of an equivalence. (Contributed by NM, 7-Aug-1994.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (∀𝑥𝜑 ↔ ∀𝑥𝜓) | ||
Theorem | 2albii 1482 | Inference adding 2 universal quantifiers to both sides of an equivalence. (Contributed by NM, 9-Mar-1997.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (∀𝑥∀𝑦𝜑 ↔ ∀𝑥∀𝑦𝜓) | ||
Theorem | hbxfrbi 1483 | A utility lemma to transfer a bound-variable hypothesis builder into a definition. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 → ∀𝑥𝜓) ⇒ ⊢ (𝜑 → ∀𝑥𝜑) | ||
Theorem | nfbii 1484 | Equality theorem for not-free. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (Ⅎ𝑥𝜑 ↔ Ⅎ𝑥𝜓) | ||
Theorem | nfxfr 1485 | A utility lemma to transfer a bound-variable hypothesis builder into a definition. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ Ⅎ𝑥𝜓 ⇒ ⊢ Ⅎ𝑥𝜑 | ||
Theorem | nfxfrd 1486 | A utility lemma to transfer a bound-variable hypothesis builder into a definition. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜒 → Ⅎ𝑥𝜑) | ||
Theorem | alcoms 1487 | Swap quantifiers in an antecedent. (Contributed by NM, 11-May-1993.) |
⊢ (∀𝑥∀𝑦𝜑 → 𝜓) ⇒ ⊢ (∀𝑦∀𝑥𝜑 → 𝜓) | ||
Theorem | hbal 1488 | If 𝑥 is not free in 𝜑, it is not free in ∀𝑦𝜑. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ (∀𝑦𝜑 → ∀𝑥∀𝑦𝜑) | ||
Theorem | alcom 1489 | Theorem 19.5 of [Margaris] p. 89. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥∀𝑦𝜑 ↔ ∀𝑦∀𝑥𝜑) | ||
Theorem | alrimdh 1490 | Deduction from Theorem 19.21 of [Margaris] p. 90. (Contributed by NM, 10-Feb-1997.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜓 → ∀𝑥𝜓) & ⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → ∀𝑥𝜒)) | ||
Theorem | albidh 1491 | Formula-building rule for universal quantifier (deduction form). (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 ↔ ∀𝑥𝜒)) | ||
Theorem | 19.26 1492 | Theorem 19.26 of [Margaris] p. 90. Also Theorem *10.22 of [WhiteheadRussell] p. 119. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 4-Jul-2014.) |
⊢ (∀𝑥(𝜑 ∧ 𝜓) ↔ (∀𝑥𝜑 ∧ ∀𝑥𝜓)) | ||
Theorem | 19.26-2 1493 | Theorem 19.26 of [Margaris] p. 90 with two quantifiers. (Contributed by NM, 3-Feb-2005.) |
⊢ (∀𝑥∀𝑦(𝜑 ∧ 𝜓) ↔ (∀𝑥∀𝑦𝜑 ∧ ∀𝑥∀𝑦𝜓)) | ||
Theorem | 19.26-3an 1494 | Theorem 19.26 of [Margaris] p. 90 with triple conjunction. (Contributed by NM, 13-Sep-2011.) |
⊢ (∀𝑥(𝜑 ∧ 𝜓 ∧ 𝜒) ↔ (∀𝑥𝜑 ∧ ∀𝑥𝜓 ∧ ∀𝑥𝜒)) | ||
Theorem | 19.33 1495 | Theorem 19.33 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) |
⊢ ((∀𝑥𝜑 ∨ ∀𝑥𝜓) → ∀𝑥(𝜑 ∨ 𝜓)) | ||
Theorem | alrot3 1496 | Theorem *11.21 in [WhiteheadRussell] p. 160. (Contributed by Andrew Salmon, 24-May-2011.) |
⊢ (∀𝑥∀𝑦∀𝑧𝜑 ↔ ∀𝑦∀𝑧∀𝑥𝜑) | ||
Theorem | alrot4 1497 | Rotate 4 universal quantifiers twice. (Contributed by NM, 2-Feb-2005.) (Proof shortened by Wolf Lammen, 28-Jun-2014.) |
⊢ (∀𝑥∀𝑦∀𝑧∀𝑤𝜑 ↔ ∀𝑧∀𝑤∀𝑥∀𝑦𝜑) | ||
Theorem | albiim 1498 | Split a biconditional and distribute quantifier. (Contributed by NM, 18-Aug-1993.) |
⊢ (∀𝑥(𝜑 ↔ 𝜓) ↔ (∀𝑥(𝜑 → 𝜓) ∧ ∀𝑥(𝜓 → 𝜑))) | ||
Theorem | 2albiim 1499 | Split a biconditional and distribute 2 quantifiers. (Contributed by NM, 3-Feb-2005.) |
⊢ (∀𝑥∀𝑦(𝜑 ↔ 𝜓) ↔ (∀𝑥∀𝑦(𝜑 → 𝜓) ∧ ∀𝑥∀𝑦(𝜓 → 𝜑))) | ||
Theorem | hband 1500 | Deduction form of bound-variable hypothesis builder hban 1558. (Contributed by NM, 2-Jan-2002.) |
⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) & ⊢ (𝜑 → (𝜒 → ∀𝑥𝜒)) ⇒ ⊢ (𝜑 → ((𝜓 ∧ 𝜒) → ∀𝑥(𝜓 ∧ 𝜒))) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |