Home | Intuitionistic Logic Explorer Theorem List (p. 2 of 102) | < 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 | ||
Axiom | ax-ia3 101 | 'And' introduction. One of the axioms of propositional logic. (Contributed by Mario Carneiro, 31-Jan-2015.) |
⊢ (𝜑 → (𝜓 → (𝜑 ∧ 𝜓))) | ||
Theorem | simpl 102 | Elimination of a conjunct. Theorem *3.26 (Simp) of [WhiteheadRussell] p. 112. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 13-Nov-2012.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜑) | ||
Theorem | simpr 103 | Elimination of a conjunct. Theorem *3.27 (Simp) of [WhiteheadRussell] p. 112. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 13-Nov-2012.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜓) | ||
Theorem | simpli 104 | Inference eliminating a conjunct. (Contributed by NM, 15-Jun-1994.) |
⊢ (𝜑 ∧ 𝜓) ⇒ ⊢ 𝜑 | ||
Theorem | simpld 105 | Deduction eliminating a conjunct. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ∧ 𝜒)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | simpri 106 | Inference eliminating a conjunct. (Contributed by NM, 15-Jun-1994.) |
⊢ (𝜑 ∧ 𝜓) ⇒ ⊢ 𝜓 | ||
Theorem | simprd 107 | Deduction eliminating a conjunct. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 3-Oct-2013.) |
⊢ (𝜑 → (𝜓 ∧ 𝜒)) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | ex 108 | Exportation inference. (This theorem used to be labeled "exp" but was changed to "ex" so as not to conflict with the math token "exp", per the June 2006 Metamath spec change.) (Contributed by NM, 5-Aug-1993.) (Proof shortened by Eric Schmidt, 22-Dec-2006.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) ⇒ ⊢ (𝜑 → (𝜓 → 𝜒)) | ||
Theorem | expcom 109 | Exportation inference with commuted antecedents. (Contributed by NM, 25-May-2005.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) ⇒ ⊢ (𝜓 → (𝜑 → 𝜒)) | ||
Definition | df-bi 110 |
This is our first definition, which introduces and defines the
biconditional connective ↔. We define a wff
of the form
(𝜑
↔ 𝜓) as an
abbreviation for
((𝜑
→ 𝜓) ∧ (𝜓 → 𝜑)).
Unlike most traditional developments, we have chosen not to have a separate symbol such as "Df." to mean "is defined as." Instead, we will later use the biconditional connective for this purpose, as it allows us to use logic to manipulate definitions directly. For an example of such a definition, see df-3or 886. This greatly simplifies many proofs since it eliminates the need for a separate mechanism for introducing and eliminating definitions. Of course, we cannot use this mechanism to define the biconditional itself, since it hasn't been introduced yet. Instead, we use a more general form of definition, described as follows. In its most general form, a definition is simply an assertion that introduces a new symbol (or a new combination of existing symbols, as in df-3an 887) that is eliminable and does not strengthen the existing language. The latter requirement means that the set of provable statements not containing the new symbol (or new combination) should remain exactly the same after the definition is introduced. Our definition of the biconditional may look unusual compared to most definitions, but it strictly satisfies these requirements. The justification for our definition is that if we mechanically replace (𝜑 ↔ 𝜓) (the definiendum i.e. the thing being defined) with ((𝜑 → 𝜓) ∧ (𝜓 → 𝜑)) (the definiens i.e. the defining expression) in the definition, the definition becomes the previously proved theorem biijust 570. It is impossible to use df-bi 110 to prove any statement expressed in the original language that can't be proved from the original axioms, because if we simply replace each instance of df-bi 110 in the proof with the corresponding biijust 570 instance, we will end up with a proof from the original axioms. Note that from Metamath's point of view, a definition is just another axiom - i.e. an assertion we claim to be true - but from our high level point of view, we are are not strengthening the language. To indicate this fact, we prefix definition labels with "df-" instead of "ax-". (This prefixing is an informal convention that means nothing to the Metamath proof verifier; it is just for human readability.) df-bi 110 itself is a conjunction of two implications (to avoid using the biconditional in its own definition), but once we have the biconditional, we can prove dfbi2 368 which uses the biconditional instead. Other definitions of the biconditional, such as dfbi3dc 1288, only hold for decidable propositions, not all propositions. (Contributed by NM, 5-Aug-1993.) (Revised by Jim Kingdon, 24-Nov-2017.) |
⊢ (((𝜑 ↔ 𝜓) → ((𝜑 → 𝜓) ∧ (𝜓 → 𝜑))) ∧ (((𝜑 → 𝜓) ∧ (𝜓 → 𝜑)) → (𝜑 ↔ 𝜓))) | ||
Theorem | bi1 111 | Property of the biconditional connective. (Contributed by NM, 11-May-1999.) (Revised by NM, 31-Jan-2015.) |
⊢ ((𝜑 ↔ 𝜓) → (𝜑 → 𝜓)) | ||
Theorem | bi3 112 | Property of the biconditional connective. (Contributed by NM, 11-May-1999.) |
⊢ ((𝜑 → 𝜓) → ((𝜓 → 𝜑) → (𝜑 ↔ 𝜓))) | ||
Theorem | biimpi 113 | Infer an implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | sylbi 114 | A mixed syllogism inference from a biconditional and an implication. Useful for substituting an antecedent with a definition. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 → 𝜒) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | imp 115 | Importation inference. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Eric Schmidt, 22-Dec-2006.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ ((𝜑 ∧ 𝜓) → 𝜒) | ||
Theorem | impcom 116 | Importation inference with commuted antecedents. (Contributed by NM, 25-May-2005.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ ((𝜓 ∧ 𝜑) → 𝜒) | ||
Theorem | impbii 117 | Infer an equivalence from an implication and its converse. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜓 → 𝜑) ⇒ ⊢ (𝜑 ↔ 𝜓) | ||
Theorem | impbidd 118 | Deduce an equivalence from two implications. (Contributed by Rodolfo Medina, 12-Oct-2010.) |
⊢ (𝜑 → (𝜓 → (𝜒 → 𝜃))) & ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜒))) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 ↔ 𝜃))) | ||
Theorem | impbid21d 119 | Deduce an equivalence from two implications. (Contributed by Wolf Lammen, 12-May-2013.) |
⊢ (𝜓 → (𝜒 → 𝜃)) & ⊢ (𝜑 → (𝜃 → 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 ↔ 𝜃))) | ||
Theorem | impbid 120 | Deduce an equivalence from two implications. (Contributed by NM, 5-Aug-1993.) (Revised by Wolf Lammen, 3-Nov-2012.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜒 → 𝜓)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜒)) | ||
Theorem | bi2 121 | Property of the biconditional connective. (Contributed by NM, 11-May-1999.) (Proof shortened by Wolf Lammen, 11-Nov-2012.) |
⊢ ((𝜑 ↔ 𝜓) → (𝜓 → 𝜑)) | ||
Theorem | bicom1 122 | Commutative law for equivalence. (Contributed by Wolf Lammen, 10-Nov-2012.) |
⊢ ((𝜑 ↔ 𝜓) → (𝜓 ↔ 𝜑)) | ||
Theorem | bicomi 123 | Inference from commutative law for logical equivalence. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 16-Sep-2013.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (𝜓 ↔ 𝜑) | ||
Theorem | biimpri 124 | Infer a converse implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 16-Sep-2013.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (𝜓 → 𝜑) | ||
Theorem | sylbir 125 | A mixed syllogism inference from a biconditional and an implication. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜓 ↔ 𝜑) & ⊢ (𝜓 → 𝜒) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | pm3.2 126 | Join antecedents with conjunction. Theorem *3.2 of [WhiteheadRussell] p. 111. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 12-Nov-2012.) (Proof shortened by Jia Ming, 17-Nov-2020.) |
⊢ (𝜑 → (𝜓 → (𝜑 ∧ 𝜓))) | ||
Theorem | sylib 127 | A mixed syllogism inference from an implication and a biconditional. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | bicom 128 | Commutative law for equivalence. Theorem *4.21 of [WhiteheadRussell] p. 117. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 11-Nov-2012.) |
⊢ ((𝜑 ↔ 𝜓) ↔ (𝜓 ↔ 𝜑)) | ||
Theorem | bicomd 129 | Commute two sides of a biconditional in a deduction. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜒 ↔ 𝜓)) | ||
Theorem | impbid1 130 | Infer an equivalence from two implications. (Contributed by NM, 6-Mar-2007.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜒 → 𝜓) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜒)) | ||
Theorem | impbid2 131 | Infer an equivalence from two implications. (Contributed by NM, 6-Mar-2007.) (Proof shortened by Wolf Lammen, 27-Sep-2013.) |
⊢ (𝜓 → 𝜒) & ⊢ (𝜑 → (𝜒 → 𝜓)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜒)) | ||
Theorem | biimpd 132 | Deduce an implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜒)) | ||
Theorem | mpbi 133 | An inference from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.) |
⊢ 𝜑 & ⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ 𝜓 | ||
Theorem | mpbir 134 | An inference from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.) |
⊢ 𝜓 & ⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ 𝜑 | ||
Theorem | mpbid 135 | A deduction from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | mpbii 136 | An inference from a nested biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 25-Oct-2012.) |
⊢ 𝜓 & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | sylibr 137 | A mixed syllogism inference from an implication and a biconditional. Useful for substituting a consequent with a definition. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜒 ↔ 𝜓) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | sylibd 138 | A syllogism deduction. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜒 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | sylbid 139 | A syllogism deduction. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → (𝜒 → 𝜃)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | mpbidi 140 | A deduction from a biconditional, related to modus ponens. (Contributed by NM, 9-Aug-1994.) |
⊢ (𝜃 → (𝜑 → 𝜓)) & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜃 → (𝜑 → 𝜒)) | ||
Theorem | syl5bi 141 | A mixed syllogism inference from a nested implication and a biconditional. Useful for substituting an embedded antecedent with a definition. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 → (𝜓 → 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 → 𝜃)) | ||
Theorem | syl5bir 142 | A mixed syllogism inference from a nested implication and a biconditional. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜓 ↔ 𝜑) & ⊢ (𝜒 → (𝜓 → 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 → 𝜃)) | ||
Theorem | syl5ib 143 | A mixed syllogism inference. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 → 𝜃)) | ||
Theorem | syl5ibcom 144 | A mixed syllogism inference. (Contributed by NM, 19-Jun-2007.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜒 → 𝜃)) | ||
Theorem | syl5ibr 145 | A mixed syllogism inference. (Contributed by NM, 3-Apr-1994.) (Revised by NM, 22-Sep-2013.) |
⊢ (𝜑 → 𝜃) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 → 𝜓)) | ||
Theorem | syl5ibrcom 146 | A mixed syllogism inference. (Contributed by NM, 20-Jun-2007.) |
⊢ (𝜑 → 𝜃) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜒 → 𝜓)) | ||
Theorem | biimprd 147 | Deduce a converse implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 22-Sep-2013.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜒 → 𝜓)) | ||
Theorem | biimpcd 148 | Deduce a commuted implication from a logical equivalence. (Contributed by NM, 3-May-1994.) (Proof shortened by Wolf Lammen, 22-Sep-2013.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜓 → (𝜑 → 𝜒)) | ||
Theorem | biimprcd 149 | Deduce a converse commuted implication from a logical equivalence. (Contributed by NM, 3-May-1994.) (Proof shortened by Wolf Lammen, 20-Dec-2013.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜒 → (𝜑 → 𝜓)) | ||
Theorem | syl6ib 150 | A mixed syllogism inference from a nested implication and a biconditional. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜒 ↔ 𝜃) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | syl6ibr 151 | A mixed syllogism inference from a nested implication and a biconditional. Useful for substituting an embedded consequent with a definition. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜃 ↔ 𝜒) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | syl6bi 152 | A mixed syllogism inference. (Contributed by NM, 2-Jan-1994.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜒 → 𝜃) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | syl6bir 153 | A mixed syllogism inference. (Contributed by NM, 18-May-1994.) |
⊢ (𝜑 → (𝜒 ↔ 𝜓)) & ⊢ (𝜒 → 𝜃) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | syl7bi 154 | A mixed syllogism inference from a doubly nested implication and a biconditional. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 → (𝜃 → (𝜓 → 𝜏))) ⇒ ⊢ (𝜒 → (𝜃 → (𝜑 → 𝜏))) | ||
Theorem | syl8ib 155 | A syllogism rule of inference. The second premise is used to replace the consequent of the first premise. (Contributed by NM, 1-Aug-1994.) |
⊢ (𝜑 → (𝜓 → (𝜒 → 𝜃))) & ⊢ (𝜃 ↔ 𝜏) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜏))) | ||
Theorem | mpbird 156 | A deduction from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜒) & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | mpbiri 157 | An inference from a nested biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 25-Oct-2012.) |
⊢ 𝜒 & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | sylibrd 158 | A syllogism deduction. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜃 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | sylbird 159 | A syllogism deduction. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝜑 → (𝜒 ↔ 𝜓)) & ⊢ (𝜑 → (𝜒 → 𝜃)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | biid 160 | Principle of identity for logical equivalence. Theorem *4.2 of [WhiteheadRussell] p. 117. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜑) | ||
Theorem | biidd 161 | Principle of identity with antecedent. (Contributed by NM, 25-Nov-1995.) |
⊢ (𝜑 → (𝜓 ↔ 𝜓)) | ||
Theorem | pm5.1im 162 | Two propositions are equivalent if they are both true. Closed form of 2th 163. Equivalent to a bi1 111-like version of the xor-connective. This theorem stays true, no matter how you permute its operands. This is evident from its sharper version (𝜑 ↔ (𝜓 ↔ (𝜑 ↔ 𝜓))). (Contributed by Wolf Lammen, 12-May-2013.) |
⊢ (𝜑 → (𝜓 → (𝜑 ↔ 𝜓))) | ||
Theorem | 2th 163 | Two truths are equivalent. (Contributed by NM, 18-Aug-1993.) |
⊢ 𝜑 & ⊢ 𝜓 ⇒ ⊢ (𝜑 ↔ 𝜓) | ||
Theorem | 2thd 164 | Two truths are equivalent (deduction rule). (Contributed by NM, 3-Jun-2012.) (Revised by NM, 29-Jan-2013.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜑 → 𝜒) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜒)) | ||
Theorem | ibi 165 | Inference that converts a biconditional implied by one of its arguments, into an implication. (Contributed by NM, 17-Oct-2003.) |
⊢ (𝜑 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | ibir 166 | Inference that converts a biconditional implied by one of its arguments, into an implication. (Contributed by NM, 22-Jul-2004.) |
⊢ (𝜑 → (𝜓 ↔ 𝜑)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | ibd 167 | Deduction that converts a biconditional implied by one of its arguments, into an implication. (Contributed by NM, 26-Jun-2004.) |
⊢ (𝜑 → (𝜓 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → (𝜓 → 𝜒)) | ||
Theorem | pm5.74 168 | Distribution of implication over biconditional. Theorem *5.74 of [WhiteheadRussell] p. 126. (Contributed by NM, 1-Aug-1994.) (Proof shortened by Wolf Lammen, 11-Apr-2013.) |
⊢ ((𝜑 → (𝜓 ↔ 𝜒)) ↔ ((𝜑 → 𝜓) ↔ (𝜑 → 𝜒))) | ||
Theorem | pm5.74i 169 | Distribution of implication over biconditional (inference rule). (Contributed by NM, 1-Aug-1994.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ ((𝜑 → 𝜓) ↔ (𝜑 → 𝜒)) | ||
Theorem | pm5.74ri 170 | Distribution of implication over biconditional (reverse inference rule). (Contributed by NM, 1-Aug-1994.) |
⊢ ((𝜑 → 𝜓) ↔ (𝜑 → 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜒)) | ||
Theorem | pm5.74d 171 | Distribution of implication over biconditional (deduction rule). (Contributed by NM, 21-Mar-1996.) |
⊢ (𝜑 → (𝜓 → (𝜒 ↔ 𝜃))) ⇒ ⊢ (𝜑 → ((𝜓 → 𝜒) ↔ (𝜓 → 𝜃))) | ||
Theorem | pm5.74rd 172 | Distribution of implication over biconditional (deduction rule). (Contributed by NM, 19-Mar-1997.) |
⊢ (𝜑 → ((𝜓 → 𝜒) ↔ (𝜓 → 𝜃))) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 ↔ 𝜃))) | ||
Theorem | bitri 173 | An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 13-Oct-2012.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (𝜑 ↔ 𝜒) | ||
Theorem | bitr2i 174 | An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (𝜒 ↔ 𝜑) | ||
Theorem | bitr3i 175 | An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜓 ↔ 𝜑) & ⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (𝜑 ↔ 𝜒) | ||
Theorem | bitr4i 176 | An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 ↔ 𝜓) ⇒ ⊢ (𝜑 ↔ 𝜒) | ||
Theorem | bitrd 177 | Deduction form of bitri 173. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 14-Apr-2013.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → (𝜒 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜃)) | ||
Theorem | bitr2d 178 | Deduction form of bitr2i 174. (Contributed by NM, 9-Jun-2004.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → (𝜒 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜃 ↔ 𝜓)) | ||
Theorem | bitr3d 179 | Deduction form of bitr3i 175. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜒 ↔ 𝜃)) | ||
Theorem | bitr4d 180 | Deduction form of bitr4i 176. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → (𝜃 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜃)) | ||
Theorem | syl5bb 181 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 ↔ 𝜃)) | ||
Theorem | syl5rbb 182 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜃 ↔ 𝜑)) | ||
Theorem | syl5bbr 183 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜓 ↔ 𝜑) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 ↔ 𝜃)) | ||
Theorem | syl5rbbr 184 | A syllogism inference from two biconditionals. (Contributed by NM, 25-Nov-1994.) |
⊢ (𝜓 ↔ 𝜑) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜃 ↔ 𝜑)) | ||
Theorem | syl6bb 185 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜒 ↔ 𝜃) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜃)) | ||
Theorem | syl6rbb 186 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜒 ↔ 𝜃) ⇒ ⊢ (𝜑 → (𝜃 ↔ 𝜓)) | ||
Theorem | syl6bbr 187 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜃 ↔ 𝜒) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜃)) | ||
Theorem | syl6rbbr 188 | A syllogism inference from two biconditionals. (Contributed by NM, 25-Nov-1994.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜃 ↔ 𝜒) ⇒ ⊢ (𝜑 → (𝜃 ↔ 𝜓)) | ||
Theorem | 3imtr3i 189 | A mixed syllogism inference, useful for removing a definition from both sides of an implication. (Contributed by NM, 10-Aug-1994.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜑 ↔ 𝜒) & ⊢ (𝜓 ↔ 𝜃) ⇒ ⊢ (𝜒 → 𝜃) | ||
Theorem | 3imtr4i 190 | A mixed syllogism inference, useful for applying a definition to both sides of an implication. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜒 ↔ 𝜑) & ⊢ (𝜃 ↔ 𝜓) ⇒ ⊢ (𝜒 → 𝜃) | ||
Theorem | 3imtr3d 191 | More general version of 3imtr3i 189. Useful for converting conditional definitions in a formula. (Contributed by NM, 8-Apr-1996.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜓 ↔ 𝜃)) & ⊢ (𝜑 → (𝜒 ↔ 𝜏)) ⇒ ⊢ (𝜑 → (𝜃 → 𝜏)) | ||
Theorem | 3imtr4d 192 | More general version of 3imtr4i 190. Useful for converting conditional definitions in a formula. (Contributed by NM, 26-Oct-1995.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜃 ↔ 𝜓)) & ⊢ (𝜑 → (𝜏 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜃 → 𝜏)) | ||
Theorem | 3imtr3g 193 | More general version of 3imtr3i 189. Useful for converting definitions in a formula. (Contributed by NM, 20-May-1996.) (Proof shortened by Wolf Lammen, 20-Dec-2013.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜓 ↔ 𝜃) & ⊢ (𝜒 ↔ 𝜏) ⇒ ⊢ (𝜑 → (𝜃 → 𝜏)) | ||
Theorem | 3imtr4g 194 | More general version of 3imtr4i 190. Useful for converting definitions in a formula. (Contributed by NM, 20-May-1996.) (Proof shortened by Wolf Lammen, 20-Dec-2013.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜃 ↔ 𝜓) & ⊢ (𝜏 ↔ 𝜒) ⇒ ⊢ (𝜑 → (𝜃 → 𝜏)) | ||
Theorem | 3bitri 195 | A chained inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 ↔ 𝜒) & ⊢ (𝜒 ↔ 𝜃) ⇒ ⊢ (𝜑 ↔ 𝜃) | ||
Theorem | 3bitrri 196 | A chained inference from transitive law for logical equivalence. (Contributed by NM, 4-Aug-2006.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 ↔ 𝜒) & ⊢ (𝜒 ↔ 𝜃) ⇒ ⊢ (𝜃 ↔ 𝜑) | ||
Theorem | 3bitr2i 197 | A chained inference from transitive law for logical equivalence. (Contributed by NM, 4-Aug-2006.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 ↔ 𝜓) & ⊢ (𝜒 ↔ 𝜃) ⇒ ⊢ (𝜑 ↔ 𝜃) | ||
Theorem | 3bitr2ri 198 | A chained inference from transitive law for logical equivalence. (Contributed by NM, 4-Aug-2006.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 ↔ 𝜓) & ⊢ (𝜒 ↔ 𝜃) ⇒ ⊢ (𝜃 ↔ 𝜑) | ||
Theorem | 3bitr3i 199 | A chained inference from transitive law for logical equivalence. (Contributed by NM, 19-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜑 ↔ 𝜒) & ⊢ (𝜓 ↔ 𝜃) ⇒ ⊢ (𝜒 ↔ 𝜃) | ||
Theorem | 3bitr3ri 200 | A chained inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜑 ↔ 𝜒) & ⊢ (𝜓 ↔ 𝜃) ⇒ ⊢ (𝜃 ↔ 𝜒) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |