Home | Intuitionistic Logic Explorer Theorem List (p. 2 of 140) | < 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 | loolin 101 | The Linearity Axiom of the infinite-valued sentential logic (L-infinity) of Lukasiewicz. (Contributed by O'Cat, 12-Aug-2004.) |
⊢ (((𝜑 → 𝜓) → (𝜓 → 𝜑)) → (𝜓 → 𝜑)) | ||
Theorem | loowoz 102 | An alternate for the Linearity Axiom of the infinite-valued sentential logic (L-infinity) of Lukasiewicz, due to Barbara Wozniakowska, Reports on Mathematical Logic 10, 129-137 (1978). (Contributed by O'Cat, 8-Aug-2004.) |
⊢ (((𝜑 → 𝜓) → (𝜑 → 𝜒)) → ((𝜓 → 𝜑) → (𝜓 → 𝜒))) | ||
Syntax | wa 103 | Extend wff definition to include conjunction ('and'). |
wff (𝜑 ∧ 𝜓) | ||
Syntax | wb 104 | Extend our wff definition to include the biconditional connective. |
wff (𝜑 ↔ 𝜓) | ||
Axiom | ax-ia1 105 | Left 'and' elimination. One of the axioms of propositional logic. Use its alias simpl 108 instead for naming consistency with set.mm. (New usage is discouraged.) (Contributed by Mario Carneiro, 31-Jan-2015.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜑) | ||
Axiom | ax-ia2 106 | Right 'and' elimination. One of the axioms of propositional logic. (Contributed by Mario Carneiro, 31-Jan-2015.) Use its alias simpr 109 instead for naming consistency with set.mm. (New usage is discouraged.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜓) | ||
Axiom | ax-ia3 107 | 'And' introduction. One of the axioms of propositional logic. (Contributed by Mario Carneiro, 31-Jan-2015.) |
⊢ (𝜑 → (𝜓 → (𝜑 ∧ 𝜓))) | ||
Theorem | simpl 108 | 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 109 | 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 110 | Inference eliminating a conjunct. (Contributed by NM, 15-Jun-1994.) |
⊢ (𝜑 ∧ 𝜓) ⇒ ⊢ 𝜑 | ||
Theorem | simpld 111 | Deduction eliminating a conjunct. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ∧ 𝜒)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | simpri 112 | Inference eliminating a conjunct. (Contributed by NM, 15-Jun-1994.) |
⊢ (𝜑 ∧ 𝜓) ⇒ ⊢ 𝜓 | ||
Theorem | simprd 113 | Deduction eliminating a conjunct. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 3-Oct-2013.) |
⊢ (𝜑 → (𝜓 ∧ 𝜒)) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | ex 114 | 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 115 | Exportation inference with commuted antecedents. (Contributed by NM, 25-May-2005.) |
⊢ ((𝜑 ∧ 𝜓) → 𝜒) ⇒ ⊢ (𝜓 → (𝜑 → 𝜒)) | ||
Definition | df-bi 116 |
This is our first definition, which introduces and defines the
biconditional connective ↔, also called
biimplication. 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 969. 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 970) 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 631. It is impossible to use df-bi 116 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 116 in the proof with the corresponding biijust 631 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 116 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 386 which uses the biconditional instead. Other definitions of the biconditional, such as dfbi3dc 1387, only hold for decidable propositions, not all propositions. (Contributed by NM, 5-Aug-1993.) (Revised by Jim Kingdon, 24-Nov-2017.) |
⊢ (((𝜑 ↔ 𝜓) → ((𝜑 → 𝜓) ∧ (𝜓 → 𝜑))) ∧ (((𝜑 → 𝜓) ∧ (𝜓 → 𝜑)) → (𝜑 ↔ 𝜓))) | ||
Theorem | biimp 117 | Property of the biconditional connective. (Contributed by NM, 11-May-1999.) (Revised by NM, 31-Jan-2015.) |
⊢ ((𝜑 ↔ 𝜓) → (𝜑 → 𝜓)) | ||
Theorem | bi3 118 | Property of the biconditional connective. (Contributed by NM, 11-May-1999.) |
⊢ ((𝜑 → 𝜓) → ((𝜓 → 𝜑) → (𝜑 ↔ 𝜓))) | ||
Theorem | biimpi 119 | Infer an implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | sylbi 120 | A mixed syllogism inference from a biconditional and an implication. Useful for substituting an antecedent with a definition. (Contributed by NM, 3-Jan-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 → 𝜒) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | sylib 121 | A mixed syllogism inference from an implication and a biconditional. (Contributed by NM, 3-Jan-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | sylbb 122 | A mixed syllogism inference from two biconditionals. (Contributed by BJ, 30-Mar-2019.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | imp 123 | Importation inference. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Eric Schmidt, 22-Dec-2006.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ ((𝜑 ∧ 𝜓) → 𝜒) | ||
Theorem | impcom 124 | Importation inference with commuted antecedents. (Contributed by NM, 25-May-2005.) |
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ ((𝜓 ∧ 𝜑) → 𝜒) | ||
Theorem | impbii 125 | Infer an equivalence from an implication and its converse. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜓 → 𝜑) ⇒ ⊢ (𝜑 ↔ 𝜓) | ||
Theorem | impbidd 126 | Deduce an equivalence from two implications. (Contributed by Rodolfo Medina, 12-Oct-2010.) |
⊢ (𝜑 → (𝜓 → (𝜒 → 𝜃))) & ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜒))) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 ↔ 𝜃))) | ||
Theorem | impbid21d 127 | Deduce an equivalence from two implications. (Contributed by Wolf Lammen, 12-May-2013.) |
⊢ (𝜓 → (𝜒 → 𝜃)) & ⊢ (𝜑 → (𝜃 → 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 ↔ 𝜃))) | ||
Theorem | impbid 128 | Deduce an equivalence from two implications. (Contributed by NM, 5-Aug-1993.) (Revised by Wolf Lammen, 3-Nov-2012.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜒 → 𝜓)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜒)) | ||
Theorem | biimpr 129 | Property of the biconditional connective. (Contributed by NM, 11-May-1999.) (Proof shortened by Wolf Lammen, 11-Nov-2012.) |
⊢ ((𝜑 ↔ 𝜓) → (𝜓 → 𝜑)) | ||
Theorem | bicom1 130 | Commutative law for equivalence. (Contributed by Wolf Lammen, 10-Nov-2012.) |
⊢ ((𝜑 ↔ 𝜓) → (𝜓 ↔ 𝜑)) | ||
Theorem | bicomi 131 | Inference from commutative law for logical equivalence. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 16-Sep-2013.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (𝜓 ↔ 𝜑) | ||
Theorem | biimpri 132 | Infer a converse implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 16-Sep-2013.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (𝜓 → 𝜑) | ||
Theorem | sylibr 133 | 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 | sylbir 134 | A mixed syllogism inference from a biconditional and an implication. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜓 ↔ 𝜑) & ⊢ (𝜓 → 𝜒) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | sylbbr 135 |
A mixed syllogism inference from two biconditionals.
Note on the various syllogism-like statements in set.mm. The hypothetical syllogism syl 14 infers an implication from two implications (and there are 3syl 17 and 4syl 18 for chaining more inferences). There are four inferences inferring an implication from one implication and one biconditional: sylbi 120, sylib 121, sylbir 134, sylibr 133; four inferences inferring an implication from two biconditionals: sylbb 122, sylbbr 135, sylbb1 136, sylbb2 137; four inferences inferring a biconditional from two biconditionals: bitri 183, bitr2i 184, bitr3i 185, bitr4i 186 (and more for chaining more biconditionals). There are also closed forms and deduction versions of these, like, among many others, syld 45, syl5 32, syl6 33, mpbid 146, bitrd 187, syl5bb 191, bitrdi 195 and variants. (Contributed by BJ, 21-Apr-2019.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (𝜒 → 𝜑) | ||
Theorem | sylbb1 136 | A mixed syllogism inference from two biconditionals. (Contributed by BJ, 21-Apr-2019.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜑 ↔ 𝜒) ⇒ ⊢ (𝜓 → 𝜒) | ||
Theorem | sylbb2 137 | A mixed syllogism inference from two biconditionals. (Contributed by BJ, 21-Apr-2019.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 ↔ 𝜓) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | pm3.2 138 | 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 | bicom 139 | 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 140 | Commute two sides of a biconditional in a deduction. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜒 ↔ 𝜓)) | ||
Theorem | impbid1 141 | Infer an equivalence from two implications. (Contributed by NM, 6-Mar-2007.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜒 → 𝜓) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜒)) | ||
Theorem | impbid2 142 | Infer an equivalence from two implications. (Contributed by NM, 6-Mar-2007.) (Proof shortened by Wolf Lammen, 27-Sep-2013.) |
⊢ (𝜓 → 𝜒) & ⊢ (𝜑 → (𝜒 → 𝜓)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜒)) | ||
Theorem | biimpd 143 | Deduce an implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜒)) | ||
Theorem | mpbi 144 | An inference from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.) |
⊢ 𝜑 & ⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ 𝜓 | ||
Theorem | mpbir 145 | An inference from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.) |
⊢ 𝜓 & ⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ 𝜑 | ||
Theorem | mpbid 146 | A deduction from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → 𝜒) | ||
Theorem | mpbii 147 | 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 | sylibd 148 | A syllogism deduction. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜒 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | sylbid 149 | A syllogism deduction. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → (𝜒 → 𝜃)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | mpbidi 150 | A deduction from a biconditional, related to modus ponens. (Contributed by NM, 9-Aug-1994.) |
⊢ (𝜃 → (𝜑 → 𝜓)) & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜃 → (𝜑 → 𝜒)) | ||
Theorem | syl5bi 151 | 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 152 | A mixed syllogism inference from a nested implication and a biconditional. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜓 ↔ 𝜑) & ⊢ (𝜒 → (𝜓 → 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 → 𝜃)) | ||
Theorem | syl5ib 153 | A mixed syllogism inference. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 → 𝜃)) | ||
Theorem | syl5ibcom 154 | A mixed syllogism inference. (Contributed by NM, 19-Jun-2007.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜒 → 𝜃)) | ||
Theorem | syl5ibr 155 | A mixed syllogism inference. (Contributed by NM, 3-Apr-1994.) (Revised by NM, 22-Sep-2013.) |
⊢ (𝜑 → 𝜃) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 → 𝜓)) | ||
Theorem | syl5ibrcom 156 | A mixed syllogism inference. (Contributed by NM, 20-Jun-2007.) |
⊢ (𝜑 → 𝜃) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜒 → 𝜓)) | ||
Theorem | biimprd 157 | Deduce a converse implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 22-Sep-2013.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜒 → 𝜓)) | ||
Theorem | biimpcd 158 | Deduce a commuted implication from a logical equivalence. (Contributed by NM, 3-May-1994.) (Proof shortened by Wolf Lammen, 22-Sep-2013.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜓 → (𝜑 → 𝜒)) | ||
Theorem | biimprcd 159 | 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 160 | A mixed syllogism inference from a nested implication and a biconditional. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜒 ↔ 𝜃) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | syl6ibr 161 | 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 162 | A mixed syllogism inference. (Contributed by NM, 2-Jan-1994.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜒 → 𝜃) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | syl6bir 163 | A mixed syllogism inference. (Contributed by NM, 18-May-1994.) |
⊢ (𝜑 → (𝜒 ↔ 𝜓)) & ⊢ (𝜒 → 𝜃) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | syl7bi 164 | A mixed syllogism inference from a doubly nested implication and a biconditional. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 → (𝜃 → (𝜓 → 𝜏))) ⇒ ⊢ (𝜒 → (𝜃 → (𝜑 → 𝜏))) | ||
Theorem | syl8ib 165 | 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 166 | A deduction from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜒) & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | mpbiri 167 | 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 168 | A syllogism deduction. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜃 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | sylbird 169 | A syllogism deduction. (Contributed by NM, 3-Aug-1994.) |
⊢ (𝜑 → (𝜒 ↔ 𝜓)) & ⊢ (𝜑 → (𝜒 → 𝜃)) ⇒ ⊢ (𝜑 → (𝜓 → 𝜃)) | ||
Theorem | biid 170 | Principle of identity for logical equivalence. Theorem *4.2 of [WhiteheadRussell] p. 117. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜑) | ||
Theorem | biidd 171 | Principle of identity with antecedent. (Contributed by NM, 25-Nov-1995.) |
⊢ (𝜑 → (𝜓 ↔ 𝜓)) | ||
Theorem | pm5.1im 172 | Two propositions are equivalent if they are both true. Closed form of 2th 173. Equivalent to a biimp 117-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 173 | Two truths are equivalent. (Contributed by NM, 18-Aug-1993.) |
⊢ 𝜑 & ⊢ 𝜓 ⇒ ⊢ (𝜑 ↔ 𝜓) | ||
Theorem | 2thd 174 | Two truths are equivalent (deduction form). (Contributed by NM, 3-Jun-2012.) (Revised by NM, 29-Jan-2013.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜑 → 𝜒) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜒)) | ||
Theorem | ibi 175 | Inference that converts a biconditional implied by one of its arguments, into an implication. (Contributed by NM, 17-Oct-2003.) |
⊢ (𝜑 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | ibir 176 | Inference that converts a biconditional implied by one of its arguments, into an implication. (Contributed by NM, 22-Jul-2004.) |
⊢ (𝜑 → (𝜓 ↔ 𝜑)) ⇒ ⊢ (𝜑 → 𝜓) | ||
Theorem | ibd 177 | Deduction that converts a biconditional implied by one of its arguments, into an implication. (Contributed by NM, 26-Jun-2004.) |
⊢ (𝜑 → (𝜓 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → (𝜓 → 𝜒)) | ||
Theorem | pm5.74 178 | 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 179 | Distribution of implication over biconditional (inference form). (Contributed by NM, 1-Aug-1994.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ ((𝜑 → 𝜓) ↔ (𝜑 → 𝜒)) | ||
Theorem | pm5.74ri 180 | Distribution of implication over biconditional (reverse inference form). (Contributed by NM, 1-Aug-1994.) |
⊢ ((𝜑 → 𝜓) ↔ (𝜑 → 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜒)) | ||
Theorem | pm5.74d 181 | Distribution of implication over biconditional (deduction form). (Contributed by NM, 21-Mar-1996.) |
⊢ (𝜑 → (𝜓 → (𝜒 ↔ 𝜃))) ⇒ ⊢ (𝜑 → ((𝜓 → 𝜒) ↔ (𝜓 → 𝜃))) | ||
Theorem | pm5.74rd 182 | Distribution of implication over biconditional (deduction form). (Contributed by NM, 19-Mar-1997.) |
⊢ (𝜑 → ((𝜓 → 𝜒) ↔ (𝜓 → 𝜃))) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 ↔ 𝜃))) | ||
Theorem | bitri 183 | An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 13-Oct-2012.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (𝜑 ↔ 𝜒) | ||
Theorem | bitr2i 184 | An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (𝜒 ↔ 𝜑) | ||
Theorem | bitr3i 185 | An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜓 ↔ 𝜑) & ⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (𝜑 ↔ 𝜒) | ||
Theorem | bitr4i 186 | An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 ↔ 𝜓) ⇒ ⊢ (𝜑 ↔ 𝜒) | ||
Theorem | bitrd 187 | Deduction form of bitri 183. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 14-Apr-2013.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → (𝜒 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜃)) | ||
Theorem | bitr2d 188 | Deduction form of bitr2i 184. (Contributed by NM, 9-Jun-2004.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → (𝜒 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜃 ↔ 𝜓)) | ||
Theorem | bitr3d 189 | Deduction form of bitr3i 185. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜑 → (𝜒 ↔ 𝜃)) | ||
Theorem | bitr4d 190 | Deduction form of bitr4i 186. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → (𝜃 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜃)) | ||
Theorem | syl5bb 191 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 ↔ 𝜃)) | ||
Theorem | bitr2id 192 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 ↔ 𝜓) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜃 ↔ 𝜑)) | ||
Theorem | bitr3id 193 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜓 ↔ 𝜑) & ⊢ (𝜒 → (𝜓 ↔ 𝜃)) ⇒ ⊢ (𝜒 → (𝜑 ↔ 𝜃)) | ||
Theorem | bitr3di 194 | A syllogism inference from two biconditionals. (Contributed by NM, 25-Nov-1994.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜓 ↔ 𝜃) ⇒ ⊢ (𝜑 → (𝜒 ↔ 𝜃)) | ||
Theorem | bitrdi 195 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜒 ↔ 𝜃) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜃)) | ||
Theorem | bitr2di 196 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜒 ↔ 𝜃) ⇒ ⊢ (𝜑 → (𝜃 ↔ 𝜓)) | ||
Theorem | bitr4di 197 | A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) & ⊢ (𝜃 ↔ 𝜒) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜃)) | ||
Theorem | bitr4id 198 | A syllogism inference from two biconditionals. (Contributed by NM, 25-Nov-1994.) |
⊢ (𝜓 ↔ 𝜒) & ⊢ (𝜑 → (𝜃 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (𝜓 ↔ 𝜃)) | ||
Theorem | 3imtr3i 199 | A mixed syllogism inference, useful for removing a definition from both sides of an implication. (Contributed by NM, 10-Aug-1994.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜑 ↔ 𝜒) & ⊢ (𝜓 ↔ 𝜃) ⇒ ⊢ (𝜒 → 𝜃) | ||
Theorem | 3imtr4i 200 | A mixed syllogism inference, useful for applying a definition to both sides of an implication. (Contributed by NM, 5-Aug-1993.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜒 ↔ 𝜑) & ⊢ (𝜃 ↔ 𝜓) ⇒ ⊢ (𝜒 → 𝜃) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |