![]() |
Metamath
Proof Explorer Theorem List (p. 11 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 | ecase3 1001 | Inference for elimination by cases. (Contributed by NM, 23-Mar-1995.) (Proof shortened by Wolf Lammen, 26-Nov-2012.) |
⊢ (𝜑 → 𝜒) & ⊢ (𝜓 → 𝜒) & ⊢ (¬ (𝜑 ∨ 𝜓) → 𝜒) ⇒ ⊢ 𝜒 | ||
Theorem | ecase 1002 | Inference for elimination by cases. (Contributed by NM, 13-Jul-2005.) |
⊢ (¬ 𝜑 → 𝜒) & ⊢ (¬ 𝜓 → 𝜒) & ⊢ ((𝜑 ∧ 𝜓) → 𝜒) ⇒ ⊢ 𝜒 | ||
Theorem | ecase3d 1003 | Deduction for elimination by cases. (Contributed by NM, 2-May-1996.) (Proof shortened by Andrew Salmon, 7-May-2011.) |
⊢ (𝜑 → (𝜓 → 𝜃)) & ⊢ (𝜑 → (𝜒 → 𝜃)) & ⊢ (𝜑 → (¬ (𝜓 ∨ 𝜒) → 𝜃)) ⇒ ⊢ (𝜑 → 𝜃) | ||
Theorem | ecased 1004 | Deduction for elimination by cases. (Contributed by NM, 8-Oct-2012.) |
⊢ (𝜑 → (¬ 𝜓 → 𝜃)) & ⊢ (𝜑 → (¬ 𝜒 → 𝜃)) & ⊢ (𝜑 → ((𝜓 ∧ 𝜒) → 𝜃)) ⇒ ⊢ (𝜑 → 𝜃) | ||
Theorem | ecase3ad 1005 | Deduction for elimination by cases. (Contributed by NM, 24-May-2013.) |
⊢ (𝜑 → (𝜓 → 𝜃)) & ⊢ (𝜑 → (𝜒 → 𝜃)) & ⊢ (𝜑 → ((¬ 𝜓 ∧ ¬ 𝜒) → 𝜃)) ⇒ ⊢ (𝜑 → 𝜃) | ||
Theorem | ccase 1006 | Inference for combining cases. (Contributed by NM, 29-Jul-1999.) (Proof shortened by Wolf Lammen, 6-Jan-2013.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜏) & ⊢ ((𝜒 ∧ 𝜓) → 𝜏) & ⊢ ((𝜑 ∧ 𝜃) → 𝜏) & ⊢ ((𝜒 ∧ 𝜃) → 𝜏) ⇒ ⊢ (((𝜑 ∨ 𝜒) ∧ (𝜓 ∨ 𝜃)) → 𝜏) | ||
Theorem | ccased 1007 | Deduction for combining cases. (Contributed by NM, 9-May-2004.) |
⊢ (𝜑 → ((𝜓 ∧ 𝜒) → 𝜂)) & ⊢ (𝜑 → ((𝜃 ∧ 𝜒) → 𝜂)) & ⊢ (𝜑 → ((𝜓 ∧ 𝜏) → 𝜂)) & ⊢ (𝜑 → ((𝜃 ∧ 𝜏) → 𝜂)) ⇒ ⊢ (𝜑 → (((𝜓 ∨ 𝜃) ∧ (𝜒 ∨ 𝜏)) → 𝜂)) | ||
Theorem | ccase2 1008 | Inference for combining cases. (Contributed by NM, 29-Jul-1999.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜏) & ⊢ (𝜒 → 𝜏) & ⊢ (𝜃 → 𝜏) ⇒ ⊢ (((𝜑 ∨ 𝜒) ∧ (𝜓 ∨ 𝜃)) → 𝜏) | ||
Theorem | 4cases 1009 | Inference eliminating two antecedents from the four possible cases that result from their true/false combinations. (Contributed by NM, 25-Oct-2003.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) & ⊢ ((𝜑 ∧ ¬ 𝜓) → 𝜒) & ⊢ ((¬ 𝜑 ∧ 𝜓) → 𝜒) & ⊢ ((¬ 𝜑 ∧ ¬ 𝜓) → 𝜒) ⇒ ⊢ 𝜒 | ||
Theorem | 4casesdan 1010 | Deduction eliminating two antecedents from the four possible cases that result from their true/false combinations. (Contributed by NM, 19-Mar-2013.) |
⊢ ((𝜑 ∧ (𝜓 ∧ 𝜒)) → 𝜃) & ⊢ ((𝜑 ∧ (𝜓 ∧ ¬ 𝜒)) → 𝜃) & ⊢ ((𝜑 ∧ (¬ 𝜓 ∧ 𝜒)) → 𝜃) & ⊢ ((𝜑 ∧ (¬ 𝜓 ∧ ¬ 𝜒)) → 𝜃) ⇒ ⊢ (𝜑 → 𝜃) | ||
Theorem | cases 1011 | Case disjunction according to the value of 𝜑. (Contributed by NM, 25-Apr-2019.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (¬ 𝜑 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜓 ↔ ((𝜑 ∧ 𝜒) ∨ (¬ 𝜑 ∧ 𝜃))) | ||
Theorem | dedlem0a 1012 | Lemma for an alternate version of weak deduction theorem. (Contributed by NM, 2-Apr-1994.) (Proof shortened by Andrew Salmon, 7-May-2011.) (Proof shortened by Wolf Lammen, 4-Dec-2012.) |
⊢ (𝜑 → (𝜓 ↔ ((𝜒 → 𝜑) → (𝜓 ∧ 𝜑)))) | ||
Theorem | dedlem0b 1013 | Lemma for an alternate version of weak deduction theorem. (Contributed by NM, 2-Apr-1994.) |
⊢ (¬ 𝜑 → (𝜓 ↔ ((𝜓 → 𝜑) → (𝜒 ∧ 𝜑)))) | ||
Theorem | dedlema 1014 | Lemma for weak deduction theorem. See also ifptru 1043. (Contributed by NM, 26-Jun-2002.) (Proof shortened by Andrew Salmon, 7-May-2011.) |
⊢ (𝜑 → (𝜓 ↔ ((𝜓 ∧ 𝜑) ∨ (𝜒 ∧ ¬ 𝜑)))) | ||
Theorem | dedlemb 1015 | Lemma for weak deduction theorem. See also ifpfal 1044. (Contributed by NM, 15-May-1999.) (Proof shortened by Andrew Salmon, 7-May-2011.) |
⊢ (¬ 𝜑 → (𝜒 ↔ ((𝜓 ∧ 𝜑) ∨ (𝜒 ∧ ¬ 𝜑)))) | ||
Theorem | cases2 1016 | Case disjunction according to the value of 𝜑. (Contributed by BJ, 6-Apr-2019.) (Proof shortened by Wolf Lammen, 28-Feb-2022.) |
⊢ (((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ 𝜒)) ↔ ((𝜑 → 𝜓) ∧ (¬ 𝜑 → 𝜒))) | ||
Theorem | cases2ALT 1017 | Alternate proof of cases2 1016, not using dedlema 1014 or dedlemb 1015. (Contributed by BJ, 6-Apr-2019.) (Proof shortened by Wolf Lammen, 2-Jan-2020.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ 𝜒)) ↔ ((𝜑 → 𝜓) ∧ (¬ 𝜑 → 𝜒))) | ||
Theorem | dfbi3 1018 | An alternate definition of the biconditional. Theorem *5.23 of [WhiteheadRussell] p. 124. (Contributed by NM, 27-Jun-2002.) (Proof shortened by Wolf Lammen, 3-Nov-2013.) (Proof shortened by NM, 29-Oct-2021.) |
⊢ ((𝜑 ↔ 𝜓) ↔ ((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ ¬ 𝜓))) | ||
Theorem | dfbi3OLD 1019 | Obsolete proof of dfbi3 1018 as of 29-Oct-2021. (Contributed by NM, 27-Jun-2002.) (Proof shortened by Wolf Lammen, 3-Nov-2013.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ ((𝜑 ↔ 𝜓) ↔ ((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ ¬ 𝜓))) | ||
Theorem | pm5.24 1020 | Theorem *5.24 of [WhiteheadRussell] p. 124. (Contributed by NM, 3-Jan-2005.) |
⊢ (¬ ((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ ¬ 𝜓)) ↔ ((𝜑 ∧ ¬ 𝜓) ∨ (𝜓 ∧ ¬ 𝜑))) | ||
Theorem | 4exmid 1021 | The disjunction of the four possible combinations of two wffs and their negations is always true. A four-way excluded middle (see exmid 430). (Contributed by David Abernethy, 28-Jan-2014.) (Proof shortened by NM, 29-Oct-2021.) |
⊢ (((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ ¬ 𝜓)) ∨ ((𝜑 ∧ ¬ 𝜓) ∨ (𝜓 ∧ ¬ 𝜑))) | ||
Theorem | 4exmidOLD 1022 | Obsolete proof of 4exmid 1021 as of 29-Oct-2021. (Contributed by David Abernethy, 28-Jan-2014.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ ¬ 𝜓)) ∨ ((𝜑 ∧ ¬ 𝜓) ∨ (𝜓 ∧ ¬ 𝜑))) | ||
Theorem | consensus 1023 | The consensus theorem. This theorem and its dual (with ∨ and ∧ interchanged) are commonly used in computer logic design to eliminate redundant terms from Boolean expressions. Specifically, we prove that the term (𝜓 ∧ 𝜒) on the left-hand side is redundant. (Contributed by NM, 16-May-2003.) (Proof shortened by Andrew Salmon, 13-May-2011.) (Proof shortened by Wolf Lammen, 20-Jan-2013.) |
⊢ ((((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ 𝜒)) ∨ (𝜓 ∧ 𝜒)) ↔ ((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ 𝜒))) | ||
Theorem | pm4.42 1024 | Theorem *4.42 of [WhiteheadRussell] p. 119. See also ifpid 1045. (Contributed by Roy F. Longton, 21-Jun-2005.) |
⊢ (𝜑 ↔ ((𝜑 ∧ 𝜓) ∨ (𝜑 ∧ ¬ 𝜓))) | ||
Theorem | prlem1 1025 | A specialized lemma for set theory (to derive the Axiom of Pairing). (Contributed by NM, 18-Oct-1995.) (Proof shortened by Andrew Salmon, 13-May-2011.) (Proof shortened by Wolf Lammen, 5-Jan-2013.) |
⊢ (𝜑 → (𝜂 ↔ 𝜒)) & ⊢ (𝜓 → ¬ 𝜃) ⇒ ⊢ (𝜑 → (𝜓 → (((𝜓 ∧ 𝜒) ∨ (𝜃 ∧ 𝜏)) → 𝜂))) | ||
Theorem | prlem2 1026 | A specialized lemma for set theory (to derive the Axiom of Pairing). (Contributed by NM, 21-Jun-1993.) (Proof shortened by Andrew Salmon, 13-May-2011.) (Proof shortened by Wolf Lammen, 9-Dec-2012.) |
⊢ (((𝜑 ∧ 𝜓) ∨ (𝜒 ∧ 𝜃)) ↔ ((𝜑 ∨ 𝜒) ∧ ((𝜑 ∧ 𝜓) ∨ (𝜒 ∧ 𝜃)))) | ||
Theorem | oplem1 1027 | A specialized lemma for set theory (ordered pair theorem). (Contributed by NM, 18-Oct-1995.) (Proof shortened by Wolf Lammen, 8-Dec-2012.) |
⊢ (𝜑 → (𝜓 ∨ 𝜒)) & ⊢ (𝜑 → (𝜃 ∨ 𝜏)) & ⊢ (𝜓 ↔ 𝜃) & ⊢ (𝜒 → (𝜃 ↔ 𝜏)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | dn1 1028 | A single axiom for Boolean algebra known as DN1. See McCune, Veroff, Fitelson, Harris, Feist, Wos, Short single axioms for Boolean algebra, Journal of Automated Reasoning, 29(1):1--16, 2002. (https://www.cs.unm.edu/~mccune/papers/basax/v12.pdf). (Contributed by Jeff Hankins, 3-Jul-2009.) (Proof shortened by Andrew Salmon, 13-May-2011.) (Proof shortened by Wolf Lammen, 6-Jan-2013.) |
⊢ (¬ (¬ (¬ (𝜑 ∨ 𝜓) ∨ 𝜒) ∨ ¬ (𝜑 ∨ ¬ (¬ 𝜒 ∨ ¬ (𝜒 ∨ 𝜃)))) ↔ 𝜒) | ||
Theorem | bianir 1029 | A closed form of mpbir 221, analogous to pm2.27 42 (assertion). (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) (Proof shortened by Roger Witte, 17-Aug-2020.) |
⊢ ((𝜑 ∧ (𝜓 ↔ 𝜑)) → 𝜓) | ||
Theorem | jaoi2 1030 | Inference removing a negated conjunct in a disjunction of an antecedent if this conjunct is part of the disjunction. (Contributed by Alexander van der Vekens, 3-Nov-2017.) (Proof shortened by Wolf Lammen, 21-Sep-2018.) |
⊢ ((𝜑 ∨ (¬ 𝜑 ∧ 𝜒)) → 𝜓) ⇒ ⊢ ((𝜑 ∨ 𝜒) → 𝜓) | ||
Theorem | jaoi3 1031 | Inference separating a disjunct of an antecedent. (Contributed by Alexander van der Vekens, 25-May-2018.) |
⊢ (𝜑 → 𝜓) & ⊢ ((¬ 𝜑 ∧ 𝜒) → 𝜓) ⇒ ⊢ ((𝜑 ∨ 𝜒) → 𝜓) | ||
This subsection introduces the conditional operator for propositions, denoted by if- (see df-ifp 1033). It is the analogue for propositions of the conditional operator for classes if (see df-if 4120). | ||
Syntax | wif 1032 | Extend class notation to include the conditional operator for propositions. |
wff if-(𝜑, 𝜓, 𝜒) | ||
Definition | df-ifp 1033 |
Definition of the conditional operator for propositions. The value of
if-(𝜑,
𝜓, 𝜒) is 𝜓 if 𝜑 is true and 𝜒 if 𝜑
false. See dfifp2 1034, dfifp3 1035, dfifp4 1036, dfifp5 1037, dfifp6 1038 and
dfifp7 1039 for alternate definitions.
This definition (in the form of dfifp2 1034) appears in Section II.24 of [Church] p. 129 (Definition D12 page 132), where it is called "conditioned disjunction". Church's [𝜓, 𝜑, 𝜒] corresponds to our if-(𝜑, 𝜓, 𝜒) (note the permutation of the first two variables). Church uses the conditional operator as an intermediate step to prove completeness of some systems of connectives. The first result is that the system {if-, ⊤, ⊥} is complete: for the induction step, consider a wff with n+1 variables; single out one variable, say 𝜑; when one sets 𝜑 to True (resp. False), then what remains is a wff of n variables, so by the induction hypothesis it corresponds to a formula using only {if-, ⊤, ⊥}, say 𝜓 (resp. 𝜒); therefore, the formula if-(𝜑, 𝜓, 𝜒) represents the initial wff. Now, since { → , ¬ } and similar systems suffice to express if-, ⊤, ⊥, they are also complete. (Contributed by BJ, 22-Jun-2019.) |
⊢ (if-(𝜑, 𝜓, 𝜒) ↔ ((𝜑 ∧ 𝜓) ∨ (¬ 𝜑 ∧ 𝜒))) | ||
Theorem | dfifp2 1034 | Alternate definition of the conditional operator for propositions. The value of if-(𝜑, 𝜓, 𝜒) is "if 𝜑 then 𝜓, and if not 𝜑 then 𝜒." This is the definition used in Section II.24 of [Church] p. 129 (Definition D12 page 132) (see comment of df-ifp 1033). (Contributed by BJ, 22-Jun-2019.) |
⊢ (if-(𝜑, 𝜓, 𝜒) ↔ ((𝜑 → 𝜓) ∧ (¬ 𝜑 → 𝜒))) | ||
Theorem | dfifp3 1035 | Alternate definition of the conditional operator for propositions. (Contributed by BJ, 30-Sep-2019.) |
⊢ (if-(𝜑, 𝜓, 𝜒) ↔ ((𝜑 → 𝜓) ∧ (𝜑 ∨ 𝜒))) | ||
Theorem | dfifp4 1036 | Alternate definition of the conditional operator for propositions. (Contributed by BJ, 30-Sep-2019.) |
⊢ (if-(𝜑, 𝜓, 𝜒) ↔ ((¬ 𝜑 ∨ 𝜓) ∧ (𝜑 ∨ 𝜒))) | ||
Theorem | dfifp5 1037 | Alternate definition of the conditional operator for propositions. (Contributed by BJ, 2-Oct-2019.) |
⊢ (if-(𝜑, 𝜓, 𝜒) ↔ ((¬ 𝜑 ∨ 𝜓) ∧ (¬ 𝜑 → 𝜒))) | ||
Theorem | dfifp6 1038 | Alternate definition of the conditional operator for propositions. (Contributed by BJ, 2-Oct-2019.) |
⊢ (if-(𝜑, 𝜓, 𝜒) ↔ ((𝜑 ∧ 𝜓) ∨ ¬ (𝜒 → 𝜑))) | ||
Theorem | dfifp7 1039 | Alternate definition of the conditional operator for propositions. (Contributed by BJ, 2-Oct-2019.) |
⊢ (if-(𝜑, 𝜓, 𝜒) ↔ ((𝜒 → 𝜑) → (𝜑 ∧ 𝜓))) | ||
Theorem | anifp 1040 | The conditional operator is implied by the conjunction of its possible outputs. Dual statement of ifpor 1041. (Contributed by BJ, 30-Sep-2019.) |
⊢ ((𝜓 ∧ 𝜒) → if-(𝜑, 𝜓, 𝜒)) | ||
Theorem | ifpor 1041 | The conditional operator implies the disjunction of its possible outputs. Dual statement of anifp 1040. (Contributed by BJ, 1-Oct-2019.) |
⊢ (if-(𝜑, 𝜓, 𝜒) → (𝜓 ∨ 𝜒)) | ||
Theorem | ifpn 1042 | Conditional operator for the negation of a proposition. (Contributed by BJ, 30-Sep-2019.) |
⊢ (if-(𝜑, 𝜓, 𝜒) ↔ if-(¬ 𝜑, 𝜒, 𝜓)) | ||
Theorem | ifptru 1043 | Value of the conditional operator for propositions when its first argument is true. Analogue for propositions of iftrue 4125. This is essentially dedlema 1014. (Contributed by BJ, 20-Sep-2019.) (Proof shortened by Wolf Lammen, 10-Jul-2020.) |
⊢ (𝜑 → (if-(𝜑, 𝜓, 𝜒) ↔ 𝜓)) | ||
Theorem | ifpfal 1044 | Value of the conditional operator for propositions when its first argument is false. Analogue for propositions of iffalse 4128. This is essentially dedlemb 1015. (Contributed by BJ, 20-Sep-2019.) (Proof shortened by Wolf Lammen, 25-Jun-2020.) |
⊢ (¬ 𝜑 → (if-(𝜑, 𝜓, 𝜒) ↔ 𝜒)) | ||
Theorem | ifpid 1045 | Value of the conditional operator for propositions when the same proposition is returned in either case. Analogue for propositions of ifid 4158. This is essentially pm4.42 1024. (Contributed by BJ, 20-Sep-2019.) |
⊢ (if-(𝜑, 𝜓, 𝜓) ↔ 𝜓) | ||
Theorem | casesifp 1046 | Version of cases 1011 expressed using if-. Case disjunction according to the value of 𝜑. One can see this as a proof that the two hypotheses characterize the conditional operator for propositions. For the converses, see ifptru 1043 and ifpfal 1044. (Contributed by BJ, 20-Sep-2019.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (¬ 𝜑 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜓 ↔ if-(𝜑, 𝜒, 𝜃)) | ||
Theorem | ifpbi123d 1047 | Equality deduction for conditional operator for propositions. (Contributed by AV, 30-Dec-2020.) |
⊢ (𝜑 → (𝜓 ↔ 𝜏)) & ⊢ (𝜑 → (𝜒 ↔ 𝜂)) & ⊢ (𝜑 → (𝜃 ↔ 𝜁)) ⇒ ⊢ (𝜑 → (if-(𝜓, 𝜒, 𝜃) ↔ if-(𝜏, 𝜂, 𝜁))) | ||
Theorem | ifpimpda 1048 | Separation of the values of the conditional operator for propositions. (Contributed by AV, 30-Dec-2020.) (Proof shortened by Wolf Lammen, 27-Feb-2021.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) & ⊢ ((𝜑 ∧ ¬ 𝜓) → 𝜃) ⇒ ⊢ (𝜑 → if-(𝜓, 𝜒, 𝜃)) | ||
Theorem | 1fpid3 1049 | The value of the conditional operator for propositions is its third argument if the first and second argument imply the third argument. (Contributed by AV, 4-Apr-2021.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) ⇒ ⊢ (if-(𝜑, 𝜓, 𝜒) → 𝜒) | ||
This subsection contains a few results related to the weak deduction theorem. For more information, see the Weak Deduction Theorem page mmdeduction.html. | ||
Theorem | elimh 1050 | Hypothesis builder for the weak deduction theorem. For more information, see the Weak Deduction Theorem page mmdeduction.html. (Contributed by NM, 26-Jun-2002.) Revised to use the conditional operator. (Revised by BJ, 30-Sep-2019.) |
⊢ ((if-(𝜒, 𝜑, 𝜓) ↔ 𝜑) → (𝜒 ↔ 𝜏)) & ⊢ ((if-(𝜒, 𝜑, 𝜓) ↔ 𝜓) → (𝜃 ↔ 𝜏)) & ⊢ 𝜃 ⇒ ⊢ 𝜏 | ||
Theorem | dedt 1051 | The weak deduction theorem. For more information, see the Weak Deduction Theorem page mmdeduction.html. (Contributed by NM, 26-Jun-2002.) Revised to use the conditional operator. (Revised by BJ, 30-Sep-2019.) |
⊢ ((if-(𝜒, 𝜑, 𝜓) ↔ 𝜑) → (𝜃 ↔ 𝜏)) & ⊢ 𝜏 ⇒ ⊢ (𝜒 → 𝜃) | ||
Theorem | con3ALT 1052 | Proof of con3 149 from its associated inference con3i 150 that illustrates the use of the weak deduction theorem dedt 1051. (Contributed by NM, 27-Jun-2002.) Revised to use the conditional operator. (Revised by BJ, 30-Sep-2019.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ ((𝜑 → 𝜓) → (¬ 𝜓 → ¬ 𝜑)) | ||
Syntax | w3o 1053 | Extend wff definition to include three-way disjunction ('or'). |
wff (𝜑 ∨ 𝜓 ∨ 𝜒) | ||
Syntax | w3a 1054 | Extend wff definition to include three-way conjunction ('and'). |
wff (𝜑 ∧ 𝜓 ∧ 𝜒) | ||
Definition | df-3or 1055 | Define disjunction ('or') of three wff's. Definition *2.33 of [WhiteheadRussell] p. 105. This abbreviation reduces the number of parentheses and emphasizes that the order of bracketing is not important by virtue of the associative law orass 545. (Contributed by NM, 8-Apr-1994.) |
⊢ ((𝜑 ∨ 𝜓 ∨ 𝜒) ↔ ((𝜑 ∨ 𝜓) ∨ 𝜒)) | ||
Definition | df-3an 1056 | Define conjunction ('and') of three wff's. Definition *4.34 of [WhiteheadRussell] p. 118. This abbreviation reduces the number of parentheses and emphasizes that the order of bracketing is not important by virtue of the associative law anass 682. (Contributed by NM, 8-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ ((𝜑 ∧ 𝜓) ∧ 𝜒)) | ||
Theorem | 3orass 1057 | Associative law for triple disjunction. (Contributed by NM, 8-Apr-1994.) |
⊢ ((𝜑 ∨ 𝜓 ∨ 𝜒) ↔ (𝜑 ∨ (𝜓 ∨ 𝜒))) | ||
Theorem | 3orel1 1058 | Partial elimination of a triple disjunction by denial of a disjunct. (Contributed by Scott Fenton, 26-Mar-2011.) |
⊢ (¬ 𝜑 → ((𝜑 ∨ 𝜓 ∨ 𝜒) → (𝜓 ∨ 𝜒))) | ||
Theorem | 3anass 1059 | Associative law for triple conjunction. (Contributed by NM, 8-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ (𝜑 ∧ (𝜓 ∧ 𝜒))) | ||
Theorem | 3anrot 1060 | Rotation law for triple conjunction. (Contributed by NM, 8-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ (𝜓 ∧ 𝜒 ∧ 𝜑)) | ||
Theorem | 3orrot 1061 | Rotation law for triple disjunction. (Contributed by NM, 4-Apr-1995.) |
⊢ ((𝜑 ∨ 𝜓 ∨ 𝜒) ↔ (𝜓 ∨ 𝜒 ∨ 𝜑)) | ||
Theorem | 3ancoma 1062 | Commutation law for triple conjunction. (Contributed by NM, 21-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ (𝜓 ∧ 𝜑 ∧ 𝜒)) | ||
Theorem | 3orcoma 1063 | Commutation law for triple disjunction. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ ((𝜑 ∨ 𝜓 ∨ 𝜒) ↔ (𝜓 ∨ 𝜑 ∨ 𝜒)) | ||
Theorem | 3ancomb 1064 | Commutation law for triple conjunction. (Contributed by NM, 21-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ (𝜑 ∧ 𝜒 ∧ 𝜓)) | ||
Theorem | 3orcomb 1065 | Commutation law for triple disjunction. (Contributed by Scott Fenton, 20-Apr-2011.) (Proof shortened by Wolf Lammen, 8-Apr-2022.) |
⊢ ((𝜑 ∨ 𝜓 ∨ 𝜒) ↔ (𝜑 ∨ 𝜒 ∨ 𝜓)) | ||
Theorem | 3orcombOLD 1066 | Obsolete version of 3orcomb 1065 as of 8-Apr-2022. (Contributed by Scott Fenton, 20-Apr-2011.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ ((𝜑 ∨ 𝜓 ∨ 𝜒) ↔ (𝜑 ∨ 𝜒 ∨ 𝜓)) | ||
Theorem | 3anrev 1067 | Reversal law for triple conjunction. (Contributed by NM, 21-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ (𝜒 ∧ 𝜓 ∧ 𝜑)) | ||
Theorem | 3anan32 1068 | Convert triple conjunction to conjunction, then commute. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ ((𝜑 ∧ 𝜒) ∧ 𝜓)) | ||
Theorem | 3anan12 1069 | Convert triple conjunction to conjunction, then commute. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) (Proof shortened by Andrew Salmon, 14-Jun-2011.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ (𝜓 ∧ (𝜑 ∧ 𝜒))) | ||
Theorem | anandi3 1070 | Distribution of triple conjunction over conjunction. (Contributed by David A. Wheeler, 4-Nov-2018.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ ((𝜑 ∧ 𝜓) ∧ (𝜑 ∧ 𝜒))) | ||
Theorem | anandi3r 1071 | Distribution of triple conjunction over conjunction. (Contributed by David A. Wheeler, 4-Nov-2018.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ ((𝜑 ∧ 𝜓) ∧ (𝜒 ∧ 𝜓))) | ||
Theorem | 3anorOLD 1072 | Obsolete version of 3anor 1075 as of 8-Apr-2022. (Contributed by Jeff Hankins, 15-Aug-2009.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ ¬ (¬ 𝜑 ∨ ¬ 𝜓 ∨ ¬ 𝜒)) | ||
Theorem | 3ioran 1073 | Negated triple disjunction as triple conjunction. (Contributed by Scott Fenton, 19-Apr-2011.) |
⊢ (¬ (𝜑 ∨ 𝜓 ∨ 𝜒) ↔ (¬ 𝜑 ∧ ¬ 𝜓 ∧ ¬ 𝜒)) | ||
Theorem | 3ianor 1074 | Negated triple conjunction expressed in terms of triple disjunction. (Contributed by Jeff Hankins, 15-Aug-2009.) (Proof shortened by Andrew Salmon, 13-May-2011.) (Revised by Wolf Lammen, 8-Apr-2022.) |
⊢ (¬ (𝜑 ∧ 𝜓 ∧ 𝜒) ↔ (¬ 𝜑 ∨ ¬ 𝜓 ∨ ¬ 𝜒)) | ||
Theorem | 3anor 1075 | Triple conjunction expressed in terms of triple disjunction. (Contributed by Jeff Hankins, 15-Aug-2009.) (Proof shortened by Wolf Lammen, 8-Apr-2022.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) ↔ ¬ (¬ 𝜑 ∨ ¬ 𝜓 ∨ ¬ 𝜒)) | ||
Theorem | 3ianorOLD 1076 | Obsolete version of 3ianor 1074 as of 8-Apr-2022. (Contributed by Jeff Hankins, 15-Aug-2009.) (Proof shortened by Andrew Salmon, 13-May-2011.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (¬ (𝜑 ∧ 𝜓 ∧ 𝜒) ↔ (¬ 𝜑 ∨ ¬ 𝜓 ∨ ¬ 𝜒)) | ||
Theorem | 3oran 1077 | Triple disjunction in terms of triple conjunction. (Contributed by NM, 8-Oct-2012.) |
⊢ ((𝜑 ∨ 𝜓 ∨ 𝜒) ↔ ¬ (¬ 𝜑 ∧ ¬ 𝜓 ∧ ¬ 𝜒)) | ||
Theorem | 3simpa 1078 | Simplification of triple conjunction. (Contributed by NM, 21-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → (𝜑 ∧ 𝜓)) | ||
Theorem | 3simpb 1079 | Simplification of triple conjunction. (Contributed by NM, 21-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → (𝜑 ∧ 𝜒)) | ||
Theorem | 3simpc 1080 | Simplification of triple conjunction. (Contributed by NM, 21-Apr-1994.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → (𝜓 ∧ 𝜒)) | ||
Theorem | simp1 1081 | Simplification of triple conjunction. (Contributed by NM, 21-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜑) | ||
Theorem | simp2 1082 | Simplification of triple conjunction. (Contributed by NM, 21-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜓) | ||
Theorem | simp3 1083 | Simplification of triple conjunction. (Contributed by NM, 21-Apr-1994.) |
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜒) | ||
Theorem | simpl1 1084 | Simplification rule. (Contributed by Jeff Hankins, 17-Nov-2009.) |
⊢ (((𝜑 ∧ 𝜓 ∧ 𝜒) ∧ 𝜃) → 𝜑) | ||
Theorem | simpl2 1085 | Simplification rule. (Contributed by Jeff Hankins, 17-Nov-2009.) |
⊢ (((𝜑 ∧ 𝜓 ∧ 𝜒) ∧ 𝜃) → 𝜓) | ||
Theorem | simpl3 1086 | Simplification rule. (Contributed by Jeff Hankins, 17-Nov-2009.) |
⊢ (((𝜑 ∧ 𝜓 ∧ 𝜒) ∧ 𝜃) → 𝜒) | ||
Theorem | simpr1 1087 | Simplification rule. (Contributed by Jeff Hankins, 17-Nov-2009.) |
⊢ ((𝜑 ∧ (𝜓 ∧ 𝜒 ∧ 𝜃)) → 𝜓) | ||
Theorem | simpr2 1088 | Simplification rule. (Contributed by Jeff Hankins, 17-Nov-2009.) |
⊢ ((𝜑 ∧ (𝜓 ∧ 𝜒 ∧ 𝜃)) → 𝜒) | ||
Theorem | simpr3 1089 | Simplification rule. (Contributed by Jeff Hankins, 17-Nov-2009.) |
⊢ ((𝜑 ∧ (𝜓 ∧ 𝜒 ∧ 𝜃)) → 𝜃) | ||
Theorem | simp1i 1090 | Infer a conjunct from a triple conjunction. (Contributed by NM, 19-Apr-2005.) |
⊢ (𝜑 ∧ 𝜓 ∧ 𝜒) ⇒ ⊢ 𝜑 | ||
Theorem | simp2i 1091 | Infer a conjunct from a triple conjunction. (Contributed by NM, 19-Apr-2005.) |
⊢ (𝜑 ∧ 𝜓 ∧ 𝜒) ⇒ ⊢ 𝜓 | ||
Theorem | simp3i 1092 | Infer a conjunct from a triple conjunction. (Contributed by NM, 19-Apr-2005.) |
⊢ (𝜑 ∧ 𝜓 ∧ 𝜒) ⇒ ⊢ 𝜒 | ||
Theorem | simp1d 1093 | Deduce a conjunct from a triple conjunction. (Contributed by NM, 4-Sep-2005.) |
⊢ (𝜑 → (𝜓 ∧ 𝜒 ∧ 𝜃)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | simp2d 1094 | Deduce a conjunct from a triple conjunction. (Contributed by NM, 4-Sep-2005.) |
⊢ (𝜑 → (𝜓 ∧ 𝜒 ∧ 𝜃)) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | simp3d 1095 | Deduce a conjunct from a triple conjunction. (Contributed by NM, 4-Sep-2005.) |
⊢ (𝜑 → (𝜓 ∧ 𝜒 ∧ 𝜃)) ⇒ ⊢ (𝜑 → 𝜃) | ||
Theorem | simp1bi 1096 | Deduce a conjunct from a triple conjunction. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) |
⊢ (𝜑 ↔ (𝜓 ∧ 𝜒 ∧ 𝜃)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | simp2bi 1097 | Deduce a conjunct from a triple conjunction. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) |
⊢ (𝜑 ↔ (𝜓 ∧ 𝜒 ∧ 𝜃)) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | simp3bi 1098 | Deduce a conjunct from a triple conjunction. (Contributed by Jonathan Ben-Naim, 3-Jun-2011.) |
⊢ (𝜑 ↔ (𝜓 ∧ 𝜒 ∧ 𝜃)) ⇒ ⊢ (𝜑 → 𝜃) | ||
Theorem | 3adant1 1099 | Deduction adding a conjunct to antecedent. (Contributed by NM, 16-Jul-1995.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) ⇒ ⊢ ((𝜃 ∧ 𝜑 ∧ 𝜓) → 𝜒) | ||
Theorem | 3adant2 1100 | Deduction adding a conjunct to antecedent. (Contributed by NM, 16-Jul-1995.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) ⇒ ⊢ ((𝜑 ∧ 𝜃 ∧ 𝜓) → 𝜒) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |