![]() |
Intuitionistic Logic Explorer Theorem List (p. 15 of 130) | < 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 | expdcom 1401 | Commuted form of expd 256. (Contributed by Alan Sare, 18-Mar-2012.) |
⊢ (𝜑 → ((𝜓 ∧ 𝜒) → 𝜃)) ⇒ ⊢ (𝜓 → (𝜒 → (𝜑 → 𝜃))) | ||
Theorem | simplbi2comg 1402 | Implication form of simplbi2com 1403. (Contributed by Alan Sare, 22-Jul-2012.) |
⊢ ((𝜑 ↔ (𝜓 ∧ 𝜒)) → (𝜒 → (𝜓 → 𝜑))) | ||
Theorem | simplbi2com 1403 | A deduction eliminating a conjunct, similar to simplbi2 380. (Contributed by Alan Sare, 22-Jul-2012.) (Proof shortened by Wolf Lammen, 10-Nov-2012.) |
⊢ (𝜑 ↔ (𝜓 ∧ 𝜒)) ⇒ ⊢ (𝜒 → (𝜓 → 𝜑)) | ||
Theorem | syl6ci 1404 | A syllogism inference combined with contraction. (Contributed by Alan Sare, 18-Mar-2012.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → 𝜃) & ⊢ (𝜒 → (𝜃 → 𝜏)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜏)) | ||
Theorem | mpisyl 1405 | A syllogism combined with a modus ponens inference. (Contributed by Alan Sare, 25-Jul-2011.) |
⊢ (𝜑 → 𝜓) & ⊢ 𝜒 & ⊢ (𝜓 → (𝜒 → 𝜃)) ⇒ ⊢ (𝜑 → 𝜃) | ||
The universal quantifier was introduced above in wal 1312 for use by df-tru 1317. See the comments in that section. In this section, we continue with the first "real" use of it. | ||
Axiom | ax-5 1406 | Axiom of Quantified Implication. Axiom C4 of [Monk2] p. 105. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥(𝜑 → 𝜓) → (∀𝑥𝜑 → ∀𝑥𝜓)) | ||
Axiom | ax-7 1407 | 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 1408 | 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 1499 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 1409 | Generalization applied twice. (Contributed by NM, 30-Apr-1998.) |
⊢ 𝜑 ⇒ ⊢ ∀𝑥∀𝑦𝜑 | ||
Theorem | mpg 1410 | Modus ponens combined with generalization. (Contributed by NM, 24-May-1994.) |
⊢ (∀𝑥𝜑 → 𝜓) & ⊢ 𝜑 ⇒ ⊢ 𝜓 | ||
Theorem | mpgbi 1411 | Modus ponens on biconditional combined with generalization. (Contributed by NM, 24-May-1994.) (Proof shortened by Stefan Allan, 28-Oct-2008.) |
⊢ (∀𝑥𝜑 ↔ 𝜓) & ⊢ 𝜑 ⇒ ⊢ 𝜓 | ||
Theorem | mpgbir 1412 | Modus ponens on biconditional combined with generalization. (Contributed by NM, 24-May-1994.) (Proof shortened by Stefan Allan, 28-Oct-2008.) |
⊢ (𝜑 ↔ ∀𝑥𝜓) & ⊢ 𝜓 ⇒ ⊢ 𝜑 | ||
Theorem | a7s 1413 | Swap quantifiers in an antecedent. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥∀𝑦𝜑 → 𝜓) ⇒ ⊢ (∀𝑦∀𝑥𝜑 → 𝜓) | ||
Theorem | alimi 1414 | Inference quantifying both antecedent and consequent. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (∀𝑥𝜑 → ∀𝑥𝜓) | ||
Theorem | 2alimi 1415 | Inference doubly quantifying both antecedent and consequent. (Contributed by NM, 3-Feb-2005.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (∀𝑥∀𝑦𝜑 → ∀𝑥∀𝑦𝜓) | ||
Theorem | alim 1416 | Theorem 19.20 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) (Proof shortened by O'Cat, 30-Mar-2008.) |
⊢ (∀𝑥(𝜑 → 𝜓) → (∀𝑥𝜑 → ∀𝑥𝜓)) | ||
Theorem | al2imi 1417 | Inference quantifying antecedent, nested antecedent, and consequent. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (∀𝑥𝜑 → (∀𝑥𝜓 → ∀𝑥𝜒)) | ||
Theorem | alanimi 1418 | Variant of al2imi 1417 with conjunctive antecedent. (Contributed by Andrew Salmon, 8-Jun-2011.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) ⇒ ⊢ ((∀𝑥𝜑 ∧ ∀𝑥𝜓) → ∀𝑥𝜒) | ||
Syntax | wnf 1419 | Extend wff definition to include the not-free predicate. |
wff Ⅎ𝑥𝜑 | ||
Definition | df-nf 1420 |
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 1733). An example of where this is used is
stdpc5 1546. See nf2 1629 for an alternate definition which
does not involve
nested quantifiers on the same variable.
Not-free is a commonly used constraint, 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 not-free notion 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 only considers syntactic freedom). For example, 𝑥 is effectively not free in the bare expression 𝑥 = 𝑥, even though 𝑥 would be considered free in the usual textbook definition, because the value of 𝑥 in the expression 𝑥 = 𝑥 cannot affect the truth of the expression (and thus substitution will not change the result). (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (Ⅎ𝑥𝜑 ↔ ∀𝑥(𝜑 → ∀𝑥𝜑)) | ||
Theorem | nfi 1421 | Deduce that 𝑥 is not free in 𝜑 from the definition. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ Ⅎ𝑥𝜑 | ||
Theorem | hbth 1422 |
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 1423 | No variable is (effectively) free in a theorem. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ 𝜑 ⇒ ⊢ Ⅎ𝑥𝜑 | ||
Theorem | nfnth 1424 | No variable is (effectively) free in a non-theorem. (Contributed by Mario Carneiro, 6-Dec-2016.) |
⊢ ¬ 𝜑 ⇒ ⊢ Ⅎ𝑥𝜑 | ||
Theorem | nftru 1425 | The true constant has no free variables. (This can also be proven in one step with nfv 1491, but this proof does not use ax-17 1489.) (Contributed by Mario Carneiro, 6-Oct-2016.) |
⊢ Ⅎ𝑥⊤ | ||
Theorem | alimdh 1426 | Deduction from Theorem 19.20 of [Margaris] p. 90. (Contributed by NM, 4-Jan-2002.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 → ∀𝑥𝜒)) | ||
Theorem | albi 1427 | Theorem 19.15 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥(𝜑 ↔ 𝜓) → (∀𝑥𝜑 ↔ ∀𝑥𝜓)) | ||
Theorem | alrimih 1428 | Inference from Theorem 19.21 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) (New usage is discouraged.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → 𝜓) ⇒ ⊢ (𝜑 → ∀𝑥𝜓) | ||
Theorem | albii 1429 | Inference adding universal quantifier to both sides of an equivalence. (Contributed by NM, 7-Aug-1994.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (∀𝑥𝜑 ↔ ∀𝑥𝜓) | ||
Theorem | 2albii 1430 | Inference adding 2 universal quantifiers to both sides of an equivalence. (Contributed by NM, 9-Mar-1997.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (∀𝑥∀𝑦𝜑 ↔ ∀𝑥∀𝑦𝜓) | ||
Theorem | hbxfrbi 1431 | A utility lemma to transfer a bound-variable hypothesis builder into a definition. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 → ∀𝑥𝜓) ⇒ ⊢ (𝜑 → ∀𝑥𝜑) | ||
Theorem | nfbii 1432 | Equality theorem for not-free. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (Ⅎ𝑥𝜑 ↔ Ⅎ𝑥𝜓) | ||
Theorem | nfxfr 1433 | A utility lemma to transfer a bound-variable hypothesis builder into a definition. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ Ⅎ𝑥𝜓 ⇒ ⊢ Ⅎ𝑥𝜑 | ||
Theorem | nfxfrd 1434 | A utility lemma to transfer a bound-variable hypothesis builder into a definition. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜒 → Ⅎ𝑥𝜑) | ||
Theorem | alcoms 1435 | Swap quantifiers in an antecedent. (Contributed by NM, 11-May-1993.) |
⊢ (∀𝑥∀𝑦𝜑 → 𝜓) ⇒ ⊢ (∀𝑦∀𝑥𝜑 → 𝜓) | ||
Theorem | hbal 1436 | If 𝑥 is not free in 𝜑, it is not free in ∀𝑦𝜑. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → ∀𝑥𝜑) ⇒ ⊢ (∀𝑦𝜑 → ∀𝑥∀𝑦𝜑) | ||
Theorem | alcom 1437 | Theorem 19.5 of [Margaris] p. 89. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥∀𝑦𝜑 ↔ ∀𝑦∀𝑥𝜑) | ||
Theorem | alrimdh 1438 | 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 1439 | Formula-building rule for universal quantifier (deduction form). (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 ↔ ∀𝑥𝜒)) | ||
Theorem | 19.26 1440 | 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 1441 | Theorem 19.26 of [Margaris] p. 90 with two quantifiers. (Contributed by NM, 3-Feb-2005.) |
⊢ (∀𝑥∀𝑦(𝜑 ∧ 𝜓) ↔ (∀𝑥∀𝑦𝜑 ∧ ∀𝑥∀𝑦𝜓)) | ||
Theorem | 19.26-3an 1442 | Theorem 19.26 of [Margaris] p. 90 with triple conjunction. (Contributed by NM, 13-Sep-2011.) |
⊢ (∀𝑥(𝜑 ∧ 𝜓 ∧ 𝜒) ↔ (∀𝑥𝜑 ∧ ∀𝑥𝜓 ∧ ∀𝑥𝜒)) | ||
Theorem | 19.33 1443 | Theorem 19.33 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) |
⊢ ((∀𝑥𝜑 ∨ ∀𝑥𝜓) → ∀𝑥(𝜑 ∨ 𝜓)) | ||
Theorem | alrot3 1444 | Theorem *11.21 in [WhiteheadRussell] p. 160. (Contributed by Andrew Salmon, 24-May-2011.) |
⊢ (∀𝑥∀𝑦∀𝑧𝜑 ↔ ∀𝑦∀𝑧∀𝑥𝜑) | ||
Theorem | alrot4 1445 | Rotate 4 universal quantifiers twice. (Contributed by NM, 2-Feb-2005.) (Proof shortened by Wolf Lammen, 28-Jun-2014.) |
⊢ (∀𝑥∀𝑦∀𝑧∀𝑤𝜑 ↔ ∀𝑧∀𝑤∀𝑥∀𝑦𝜑) | ||
Theorem | albiim 1446 | Split a biconditional and distribute quantifier. (Contributed by NM, 18-Aug-1993.) |
⊢ (∀𝑥(𝜑 ↔ 𝜓) ↔ (∀𝑥(𝜑 → 𝜓) ∧ ∀𝑥(𝜓 → 𝜑))) | ||
Theorem | 2albiim 1447 | Split a biconditional and distribute 2 quantifiers. (Contributed by NM, 3-Feb-2005.) |
⊢ (∀𝑥∀𝑦(𝜑 ↔ 𝜓) ↔ (∀𝑥∀𝑦(𝜑 → 𝜓) ∧ ∀𝑥∀𝑦(𝜓 → 𝜑))) | ||
Theorem | hband 1448 | Deduction form of bound-variable hypothesis builder hban 1509. (Contributed by NM, 2-Jan-2002.) |
⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) & ⊢ (𝜑 → (𝜒 → ∀𝑥𝜒)) ⇒ ⊢ (𝜑 → ((𝜓 ∧ 𝜒) → ∀𝑥(𝜓 ∧ 𝜒))) | ||
Theorem | hb3and 1449 | Deduction form of bound-variable hypothesis builder hb3an 1512. (Contributed by NM, 17-Feb-2013.) |
⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) & ⊢ (𝜑 → (𝜒 → ∀𝑥𝜒)) & ⊢ (𝜑 → (𝜃 → ∀𝑥𝜃)) ⇒ ⊢ (𝜑 → ((𝜓 ∧ 𝜒 ∧ 𝜃) → ∀𝑥(𝜓 ∧ 𝜒 ∧ 𝜃))) | ||
Theorem | hbald 1450 | Deduction form of bound-variable hypothesis builder hbal 1436. (Contributed by NM, 2-Jan-2002.) |
⊢ (𝜑 → ∀𝑦𝜑) & ⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → (∀𝑦𝜓 → ∀𝑥∀𝑦𝜓)) | ||
Syntax | wex 1451 | Extend wff definition to include the existential quantifier ("there exists"). |
wff ∃𝑥𝜑 | ||
Axiom | ax-ie1 1452 | 𝑥 is bound in ∃𝑥𝜑. One of the axioms of predicate logic. (Contributed by Mario Carneiro, 31-Jan-2015.) |
⊢ (∃𝑥𝜑 → ∀𝑥∃𝑥𝜑) | ||
Axiom | ax-ie2 1453 | Define existential quantification. ∃𝑥𝜑 means "there exists at least one set 𝑥 such that 𝜑 is true." One of the axioms of predicate logic. (Contributed by Mario Carneiro, 31-Jan-2015.) |
⊢ (∀𝑥(𝜓 → ∀𝑥𝜓) → (∀𝑥(𝜑 → 𝜓) ↔ (∃𝑥𝜑 → 𝜓))) | ||
Theorem | hbe1 1454 | 𝑥 is not free in ∃𝑥𝜑. (Contributed by NM, 5-Aug-1993.) |
⊢ (∃𝑥𝜑 → ∀𝑥∃𝑥𝜑) | ||
Theorem | nfe1 1455 | 𝑥 is not free in ∃𝑥𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ Ⅎ𝑥∃𝑥𝜑 | ||
Theorem | 19.23ht 1456 | Closed form of Theorem 19.23 of [Margaris] p. 90. (Contributed by NM, 7-Nov-2005.) (Revised by Mario Carneiro, 1-Feb-2015.) |
⊢ (∀𝑥(𝜓 → ∀𝑥𝜓) → (∀𝑥(𝜑 → 𝜓) ↔ (∃𝑥𝜑 → 𝜓))) | ||
Theorem | 19.23h 1457 | Theorem 19.23 of [Margaris] p. 90. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 1-Feb-2015.) |
⊢ (𝜓 → ∀𝑥𝜓) ⇒ ⊢ (∀𝑥(𝜑 → 𝜓) ↔ (∃𝑥𝜑 → 𝜓)) | ||
Theorem | alnex 1458 | Theorem 19.7 of [Margaris] p. 89. To read this intuitionistically, think of it as "if 𝜑 can be refuted for all 𝑥, then it is not possible to find an 𝑥 for which 𝜑 holds" (and likewise for the converse). Comparing this with dfexdc 1460 illustrates that statements which look similar (to someone used to classical logic) can be different intuitionistically due to different placement of negations. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 1-Feb-2015.) (Revised by Mario Carneiro, 12-May-2015.) |
⊢ (∀𝑥 ¬ 𝜑 ↔ ¬ ∃𝑥𝜑) | ||
Theorem | nex 1459 | Generalization rule for negated wff. (Contributed by NM, 18-May-1994.) |
⊢ ¬ 𝜑 ⇒ ⊢ ¬ ∃𝑥𝜑 | ||
Theorem | dfexdc 1460 | Defining ∃𝑥𝜑 given decidability. It is common in classical logic to define ∃𝑥𝜑 as ¬ ∀𝑥¬ 𝜑 but in intuitionistic logic without a decidability condition, that is only an implication not an equivalence, as seen at exalim 1461. (Contributed by Jim Kingdon, 15-Mar-2018.) |
⊢ (DECID ∃𝑥𝜑 → (∃𝑥𝜑 ↔ ¬ ∀𝑥 ¬ 𝜑)) | ||
Theorem | exalim 1461 | One direction of a classical definition of existential quantification. One direction of Definition of [Margaris] p. 49. For a decidable proposition, this is an equivalence, as seen as dfexdc 1460. (Contributed by Jim Kingdon, 29-Jul-2018.) |
⊢ (∃𝑥𝜑 → ¬ ∀𝑥 ¬ 𝜑) | ||
The equality predicate was introduced above in wceq 1314 for use by df-tru 1317. See the comments in that section. In this section, we continue with the first "real" use of it. | ||
Theorem | weq 1462 |
Extend wff definition to include atomic formulas using the equality
predicate.
(Instead of introducing weq 1462 as an axiomatic statement, as was done in an older version of this database, we introduce it by "proving" a special case of set theory's more general wceq 1314. This lets us avoid overloading the = connective, thus preventing ambiguity that would complicate certain Metamath parsers. However, logically weq 1462 is considered to be a primitive syntax, even though here it is artificially "derived" from wceq 1314. Note: To see the proof steps of this syntax proof, type "show proof weq /all" in the Metamath program.) (Contributed by NM, 24-Jan-2006.) |
wff 𝑥 = 𝑦 | ||
Syntax | wcel 1463 |
Extend wff definition to include the membership connective between
classes.
(The purpose of introducing wff 𝐴 ∈ 𝐵 here is to allow us to express i.e. "prove" the wel 1464 of predicate calculus in terms of the wceq 1314 of set theory, so that we don't "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.) |
wff 𝐴 ∈ 𝐵 | ||
Theorem | wel 1464 |
Extend wff definition to include atomic formulas with the epsilon
(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 syntactical construction introduces a binary non-logical predicate symbol ∈ (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 1464 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 1463. This lets us avoid overloading the ∈ connective, thus preventing ambiguity that would complicate certain Metamath parsers. However, logically wel 1464 is considered to be a primitive syntax, even though here it is artificially "derived" from wcel 1463. Note: To see the proof steps of this syntax proof, type "show proof wel /all" in the Metamath program.) (Contributed by NM, 24-Jan-2006.) |
wff 𝑥 ∈ 𝑦 | ||
Axiom | ax-8 1465 |
Axiom of Equality. One of the equality and substitution axioms of
predicate calculus with equality. This is similar to, but not quite, a
transitive law for equality (proved later as equtr 1668). Axiom scheme C8'
in [Megill] p. 448 (p. 16 of the preprint).
Also appears as Axiom C7 of
[Monk2] p. 105.
Axioms ax-8 1465 through ax-16 1768 are the axioms having to do with equality, substitution, and logical properties of our binary predicate ∈ (which later in set theory will mean "is a member of"). Note that all axioms except ax-16 1768 and ax-17 1489 are still valid even when 𝑥, 𝑦, and 𝑧 are replaced with the same variable because they do not have any distinct variable (Metamath's $d) restrictions. Distinct variable restrictions are required for ax-16 1768 and ax-17 1489 only. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝑥 = 𝑧 → 𝑦 = 𝑧)) | ||
Axiom | ax-10 1466 |
Axiom of Quantifier Substitution. One of the equality and substitution
axioms of predicate calculus with equality. Appears as Lemma L12 in
[Megill] p. 445 (p. 12 of the preprint).
The original version of this axiom was ax-10o 1677 ("o" for "old") and was replaced with this shorter ax-10 1466 in May 2008. The old axiom is proved from this one as theorem ax10o 1676. Conversely, this axiom is proved from ax-10o 1677 as theorem ax10 1678. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥 𝑥 = 𝑦 → ∀𝑦 𝑦 = 𝑥) | ||
Axiom | ax-11 1467 |
Axiom of Variable Substitution. One of the 5 equality axioms of predicate
calculus. The final consequent ∀𝑥(𝑥 = 𝑦 → 𝜑) is a way of
expressing "𝑦 substituted for 𝑥 in wff
𝜑
" (cf. sb6 1840). 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.
Variants of this axiom which are equivalent in classical logic but which have not been shown to be equivalent for intuitionistic logic are ax11v 1781, ax11v2 1774 and ax-11o 1777. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (∀𝑦𝜑 → ∀𝑥(𝑥 = 𝑦 → 𝜑))) | ||
Axiom | ax-i12 1468 |
Axiom of Quantifier Introduction. One of the equality and substitution
axioms of predicate calculus with equality. Informally, it says that
whenever 𝑧 is distinct from 𝑥 and
𝑦,
and 𝑥 =
𝑦 is true,
then 𝑥 = 𝑦 quantified with 𝑧 is also
true. In other words, 𝑧
is irrelevant to the truth of 𝑥 = 𝑦. Axiom scheme C9' in [Megill]
p. 448 (p. 16 of the preprint). It apparently does not otherwise appear
in the literature but is easily proved from textbook predicate calculus by
cases.
This axiom has been modified from the original ax-12 1472 for compatibility with intuitionistic logic. (Contributed by Mario Carneiro, 31-Jan-2015.) |
⊢ (∀𝑧 𝑧 = 𝑥 ∨ (∀𝑧 𝑧 = 𝑦 ∨ ∀𝑧(𝑥 = 𝑦 → ∀𝑧 𝑥 = 𝑦))) | ||
Axiom | ax-bndl 1469 |
Axiom of bundling. The general idea of this axiom is that two variables
are either distinct or non-distinct. That idea could be expressed as
∀𝑧𝑧 = 𝑥 ∨ ¬ ∀𝑧𝑧 = 𝑥. However, we instead choose an axiom
which has many of the same consequences, but which is different with
respect to a universe which contains only one object. ∀𝑧𝑧 = 𝑥 holds
if 𝑧 and 𝑥 are the same variable,
likewise for 𝑧 and 𝑦,
and ∀𝑥∀𝑧(𝑥 = 𝑦 → ∀𝑧𝑥 = 𝑦) holds if 𝑧 is distinct from
the others (and the universe has at least two objects).
As with other statements of the form "x is decidable (either true or false)", this does not entail the full Law of the Excluded Middle (which is the proposition that all statements are decidable), but instead merely the assertion that particular kinds of statements are decidable (or in this case, an assertion similar to decidability). This axiom implies ax-i12 1468 as can be seen at axi12 1477. Whether ax-bndl 1469 can be proved from the remaining axioms including ax-i12 1468 is not known. The reason we call this "bundling" is that a statement without a distinct variable constraint "bundles" together two statements, one in which the two variables are the same and one in which they are different. (Contributed by Mario Carneiro and Jim Kingdon, 14-Mar-2018.) |
⊢ (∀𝑧 𝑧 = 𝑥 ∨ (∀𝑧 𝑧 = 𝑦 ∨ ∀𝑥∀𝑧(𝑥 = 𝑦 → ∀𝑧 𝑥 = 𝑦))) | ||
Axiom | ax-4 1470 |
Axiom of Specialization. A quantified wff implies the wff without a
quantifier (i.e. an instance, or special case, of the generalized wff).
In other words if something is true for all 𝑥, it is true for any
specific 𝑥 (that would typically occur as a free
variable in the wff
substituted for 𝜑). (A free variable is one that does
not occur in
the scope of a quantifier: 𝑥 and 𝑦 are both free in 𝑥 = 𝑦,
but only 𝑥 is free in ∀𝑦𝑥 = 𝑦.) Axiom scheme C5' in [Megill]
p. 448 (p. 16 of the preprint). Also appears as Axiom B5 of [Tarski]
p. 67 (under his system S2, defined in the last paragraph on p. 77).
Note that the converse of this axiom does not hold in general, but a weaker inference form of the converse holds and is expressed as rule ax-gen 1408. Conditional forms of the converse are given by ax-12 1472, ax-16 1768, and ax-17 1489. Unlike the more general textbook Axiom of Specialization, we cannot choose a variable different from 𝑥 for the special case. For use, that requires the assistance of equality axioms, and we deal with it later after we introduce the definition of proper substitution - see stdpc4 1731. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥𝜑 → 𝜑) | ||
Theorem | sp 1471 | Specialization. Another name for ax-4 1470. (Contributed by NM, 21-May-2008.) |
⊢ (∀𝑥𝜑 → 𝜑) | ||
Theorem | ax-12 1472 | Rederive the original version of the axiom from ax-i12 1468. (Contributed by Mario Carneiro, 3-Feb-2015.) |
⊢ (¬ ∀𝑧 𝑧 = 𝑥 → (¬ ∀𝑧 𝑧 = 𝑦 → (𝑥 = 𝑦 → ∀𝑧 𝑥 = 𝑦))) | ||
Theorem | ax12or 1473 | Another name for ax-i12 1468. (Contributed by NM, 3-Feb-2015.) |
⊢ (∀𝑧 𝑧 = 𝑥 ∨ (∀𝑧 𝑧 = 𝑦 ∨ ∀𝑧(𝑥 = 𝑦 → ∀𝑧 𝑥 = 𝑦))) | ||
Axiom | ax-13 1474 | Axiom of Equality. 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 the ∈ binary predicate. Axiom scheme C12' in [Megill] p. 448 (p. 16 of the preprint). It is a special case of Axiom B8 (p. 75) of system S2 of [Tarski] p. 77. "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. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝑥 ∈ 𝑧 → 𝑦 ∈ 𝑧)) | ||
Axiom | ax-14 1475 | Axiom of Equality. 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 the ∈ binary predicate. Axiom scheme C13' in [Megill] p. 448 (p. 16 of the preprint). It is a special case of Axiom B8 (p. 75) of system S2 of [Tarski] p. 77. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝑥 = 𝑦 → (𝑧 ∈ 𝑥 → 𝑧 ∈ 𝑦)) | ||
Theorem | hbequid 1476 | Bound-variable hypothesis builder for 𝑥 = 𝑥. This theorem tells us that any variable, including 𝑥, is effectively not free in 𝑥 = 𝑥, even though 𝑥 is technically free according to the traditional definition of free variable. (The proof uses only ax-5 1406, ax-8 1465, ax-12 1472, and ax-gen 1408. This shows that this can be proved without ax-9 1494, even though the theorem equid 1660 cannot be. A shorter proof using ax-9 1494 is obtainable from equid 1660 and hbth 1422.) (Contributed by NM, 13-Jan-2011.) (Proof shortened by Wolf Lammen, 23-Mar-2014.) |
⊢ (𝑥 = 𝑥 → ∀𝑦 𝑥 = 𝑥) | ||
Theorem | axi12 1477 | Proof that ax-i12 1468 follows from ax-bndl 1469. So that we can track which theorems rely on ax-bndl 1469, proofs should reference ax-i12 1468 rather than this theorem. (Contributed by Jim Kingdon, 17-Aug-2018.) (New usage is discouraged). (Proof modification is discouraged.) |
⊢ (∀𝑧 𝑧 = 𝑥 ∨ (∀𝑧 𝑧 = 𝑦 ∨ ∀𝑧(𝑥 = 𝑦 → ∀𝑧 𝑥 = 𝑦))) | ||
Theorem | alequcom 1478 | Commutation law for identical variable specifiers. The antecedent and consequent are true when 𝑥 and 𝑦 are substituted with the same variable. Lemma L12 in [Megill] p. 445 (p. 12 of the preprint). (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥 𝑥 = 𝑦 → ∀𝑦 𝑦 = 𝑥) | ||
Theorem | alequcoms 1479 | A commutation rule for identical variable specifiers. (Contributed by NM, 5-Aug-1993.) |
⊢ (∀𝑥 𝑥 = 𝑦 → 𝜑) ⇒ ⊢ (∀𝑦 𝑦 = 𝑥 → 𝜑) | ||
Theorem | nalequcoms 1480 | A commutation rule for distinct variable specifiers. (Contributed by NM, 2-Jan-2002.) (Revised by Mario Carneiro, 2-Feb-2015.) |
⊢ (¬ ∀𝑥 𝑥 = 𝑦 → 𝜑) ⇒ ⊢ (¬ ∀𝑦 𝑦 = 𝑥 → 𝜑) | ||
Theorem | nfr 1481 | Consequence of the definition of not-free. (Contributed by Mario Carneiro, 26-Sep-2016.) |
⊢ (Ⅎ𝑥𝜑 → (𝜑 → ∀𝑥𝜑)) | ||
Theorem | nfri 1482 | Consequence of the definition of not-free. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (𝜑 → ∀𝑥𝜑) | ||
Theorem | nfrd 1483 | Consequence of the definition of not-free in a context. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) | ||
Theorem | alimd 1484 | Deduction from Theorem 19.20 of [Margaris] p. 90. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 → ∀𝑥𝜒)) | ||
Theorem | alrimi 1485 | Inference from Theorem 19.21 of [Margaris] p. 90. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → 𝜓) ⇒ ⊢ (𝜑 → ∀𝑥𝜓) | ||
Theorem | nfd 1486 | Deduce that 𝑥 is not free in 𝜓 in a context. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
Theorem | nfdh 1487 | Deduce that 𝑥 is not free in 𝜓 in a context. (Contributed by Mario Carneiro, 24-Sep-2016.) |
⊢ (𝜑 → ∀𝑥𝜑) & ⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) ⇒ ⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
Theorem | nfrimi 1488 | Moving an antecedent outside Ⅎ. (Contributed by Jim Kingdon, 23-Mar-2018.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑥(𝜑 → 𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
Axiom | ax-17 1489* |
Axiom to quantify a variable over a formula in which it does not occur.
Axiom C5 in [Megill] p. 444 (p. 11 of the
preprint). Also appears as
Axiom B6 (p. 75) of system S2 of [Tarski]
p. 77 and Axiom C5-1 of
[Monk2] p. 113.
(Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → ∀𝑥𝜑) | ||
Theorem | a17d 1490* | ax-17 1489 with antecedent. (Contributed by NM, 1-Mar-2013.) |
⊢ (𝜑 → (𝜓 → ∀𝑥𝜓)) | ||
Theorem | nfv 1491* | If 𝑥 is not present in 𝜑, then 𝑥 is not free in 𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ Ⅎ𝑥𝜑 | ||
Theorem | nfvd 1492* | nfv 1491 with antecedent. Useful in proofs of deduction versions of bound-variable hypothesis builders such as nfimd 1547. (Contributed by Mario Carneiro, 6-Oct-2016.) |
⊢ (𝜑 → Ⅎ𝑥𝜓) | ||
Axiom | ax-i9 1493 | Axiom of Existence. One of the equality and substitution axioms of predicate calculus with equality. One thing this axiom tells us is that at least one thing exists (although ax-4 1470 and possibly others also tell us that, i.e. they are not valid in the empty domain of a "free logic"). In this form (not requiring that 𝑥 and 𝑦 be distinct) it was used in an axiom system of Tarski (see Axiom B7' in footnote 1 of [KalishMontague] p. 81.) Another name for this theorem is a9e 1657, which has additional remarks. (Contributed by Mario Carneiro, 31-Jan-2015.) |
⊢ ∃𝑥 𝑥 = 𝑦 | ||
Theorem | ax-9 1494 | Derive ax-9 1494 from ax-i9 1493, the modified version for intuitionistic logic. Although ax-9 1494 does hold intuistionistically, in intuitionistic logic it is weaker than ax-i9 1493. (Contributed by NM, 3-Feb-2015.) |
⊢ ¬ ∀𝑥 ¬ 𝑥 = 𝑦 | ||
Theorem | equidqe 1495 | equid 1660 with some quantification and negation without using ax-4 1470 or ax-17 1489. (Contributed by NM, 13-Jan-2011.) (Proof shortened by Wolf Lammen, 27-Feb-2014.) |
⊢ ¬ ∀𝑦 ¬ 𝑥 = 𝑥 | ||
Theorem | ax4sp1 1496 | A special case of ax-4 1470 without using ax-4 1470 or ax-17 1489. (Contributed by NM, 13-Jan-2011.) |
⊢ (∀𝑦 ¬ 𝑥 = 𝑥 → ¬ 𝑥 = 𝑥) | ||
Axiom | ax-ial 1497 | 𝑥 is not free in ∀𝑥𝜑. One of the axioms of predicate logic. (Contributed by Mario Carneiro, 31-Jan-2015.) |
⊢ (∀𝑥𝜑 → ∀𝑥∀𝑥𝜑) | ||
Axiom | ax-i5r 1498 | Axiom of quantifier collection. (Contributed by Mario Carneiro, 31-Jan-2015.) |
⊢ ((∀𝑥𝜑 → ∀𝑥𝜓) → ∀𝑥(∀𝑥𝜑 → 𝜓)) | ||
Theorem | spi 1499 | Inference reversing generalization (specialization). (Contributed by NM, 5-Aug-1993.) |
⊢ ∀𝑥𝜑 ⇒ ⊢ 𝜑 | ||
Theorem | sps 1500 | Generalization of antecedent. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (∀𝑥𝜑 → 𝜓) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |