HomeHome Intuitionistic Logic Explorer
Theorem List (p. 2 of 132)
< Previous  Next >
Browser slow? Try the
Unicode version.

Mirrors  >  Metamath Home Page  >  ILE Home Page  >  Theorem List Contents  >  Recent Proofs       This page: Page List

Theorem List for Intuitionistic Logic Explorer - 101-200   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremloolin 101 The Linearity Axiom of the infinite-valued sentential logic (L-infinity) of Lukasiewicz. (Contributed by O'Cat, 12-Aug-2004.)
 |-  ( ( ( ph  ->  ps )  ->  ( ps  ->  ph ) )  ->  ( ps  ->  ph ) )
 
Theoremloowoz 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.)
 |-  ( ( ( ph  ->  ps )  ->  ( ph  ->  ch ) )  ->  ( ( ps  ->  ph )  ->  ( ps  ->  ch ) ) )
 
1.2.4  Logical conjunction and logical equivalence
 
Syntaxwa 103 Extend wff definition to include conjunction ('and').
 wff  ( ph  /\  ps )
 
Syntaxwb 104 Extend our wff definition to include the biconditional connective.
 wff  ( ph  <->  ps )
 
Axiomax-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.)
 |-  ( ( ph  /\  ps )  ->  ph )
 
Axiomax-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.)
 |-  ( ( ph  /\  ps )  ->  ps )
 
Axiomax-ia3 107 'And' introduction. One of the axioms of propositional logic. (Contributed by Mario Carneiro, 31-Jan-2015.)
 |-  ( ph  ->  ( ps  ->  ( ph  /\  ps ) ) )
 
Theoremsimpl 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.)
 |-  ( ( ph  /\  ps )  ->  ph )
 
Theoremsimpr 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.)
 |-  ( ( ph  /\  ps )  ->  ps )
 
Theoremsimpli 110 Inference eliminating a conjunct. (Contributed by NM, 15-Jun-1994.)
 |-  ( ph  /\  ps )   =>    |-  ph
 
Theoremsimpld 111 Deduction eliminating a conjunct. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ( ps  /\  ch ) )   =>    |-  ( ph  ->  ps )
 
Theoremsimpri 112 Inference eliminating a conjunct. (Contributed by NM, 15-Jun-1994.)
 |-  ( ph  /\  ps )   =>    |- 
 ps
 
Theoremsimprd 113 Deduction eliminating a conjunct. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 3-Oct-2013.)
 |-  ( ph  ->  ( ps  /\  ch ) )   =>    |-  ( ph  ->  ch )
 
Theoremex 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.)
 |-  ( ( ph  /\  ps )  ->  ch )   =>    |-  ( ph  ->  ( ps  ->  ch ) )
 
Theoremexpcom 115 Exportation inference with commuted antecedents. (Contributed by NM, 25-May-2005.)
 |-  ( ( ph  /\  ps )  ->  ch )   =>    |-  ( ps  ->  ( ph  ->  ch ) )
 
Definitiondf-bi 116 This is our first definition, which introduces and defines the biconditional connective  <->. We define a wff of the form  ( ph  <->  ps ) as an abbreviation for  ( ( ph  ->  ps )  /\  ( ps  ->  ph )
).

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 948. 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 949) 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  ( ph  <->  ps ) (the definiendum i.e. the thing being defined) with  ( ( ph  ->  ps )  /\  ( ps  ->  ph )
) (the definiens i.e. the defining expression) in the definition, the definition becomes the previously proved theorem biijust 615. 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 615 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 385 which uses the biconditional instead.

Other definitions of the biconditional, such as dfbi3dc 1360, only hold for decidable propositions, not all propositions. (Contributed by NM, 5-Aug-1993.) (Revised by Jim Kingdon, 24-Nov-2017.)

 |-  ( ( ( ph  <->  ps )  ->  ( ( ph  ->  ps )  /\  ( ps  ->  ph ) ) ) 
 /\  ( ( (
 ph  ->  ps )  /\  ( ps  ->  ph ) )  ->  ( ph  <->  ps ) ) )
 
Theorembi1 117 Property of the biconditional connective. (Contributed by NM, 11-May-1999.) (Revised by NM, 31-Jan-2015.)
 |-  ( ( ph  <->  ps )  ->  ( ph  ->  ps ) )
 
Theorembi3 118 Property of the biconditional connective. (Contributed by NM, 11-May-1999.)
 |-  ( ( ph  ->  ps )  ->  ( ( ps  ->  ph )  ->  ( ph 
 <->  ps ) ) )
 
Theorembiimpi 119 Infer an implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  <->  ps )   =>    |-  ( ph  ->  ps )
 
Theoremsylbi 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.)
 |-  ( ph  <->  ps )   &    |-  ( ps  ->  ch )   =>    |-  ( ph  ->  ch )
 
Theoremsylib 121 A mixed syllogism inference from an implication and a biconditional. (Contributed by NM, 3-Jan-1993.)
 |-  ( ph  ->  ps )   &    |-  ( ps 
 <->  ch )   =>    |-  ( ph  ->  ch )
 
Theoremsylbb 122 A mixed syllogism inference from two biconditionals. (Contributed by BJ, 30-Mar-2019.)
 |-  ( ph  <->  ps )   &    |-  ( ps  <->  ch )   =>    |-  ( ph  ->  ch )
 
Theoremimp 123 Importation inference. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Eric Schmidt, 22-Dec-2006.)
 |-  ( ph  ->  ( ps  ->  ch ) )   =>    |-  ( ( ph  /\ 
 ps )  ->  ch )
 
Theoremimpcom 124 Importation inference with commuted antecedents. (Contributed by NM, 25-May-2005.)
 |-  ( ph  ->  ( ps  ->  ch ) )   =>    |-  ( ( ps 
 /\  ph )  ->  ch )
 
Theoremimpbii 125 Infer an equivalence from an implication and its converse. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ps )   &    |-  ( ps  ->  ph )   =>    |-  ( ph  <->  ps )
 
Theoremimpbidd 126 Deduce an equivalence from two implications. (Contributed by Rodolfo Medina, 12-Oct-2010.)
 |-  ( ph  ->  ( ps  ->  ( ch  ->  th ) ) )   &    |-  ( ph  ->  ( ps  ->  ( th  ->  ch )
 ) )   =>    |-  ( ph  ->  ( ps  ->  ( ch  <->  th ) ) )
 
Theoremimpbid21d 127 Deduce an equivalence from two implications. (Contributed by Wolf Lammen, 12-May-2013.)
 |-  ( ps  ->  ( ch  ->  th ) )   &    |-  ( ph  ->  ( th  ->  ch ) )   =>    |-  ( ph  ->  ( ps  ->  ( ch  <->  th ) ) )
 
Theoremimpbid 128 Deduce an equivalence from two implications. (Contributed by NM, 5-Aug-1993.) (Revised by Wolf Lammen, 3-Nov-2012.)
 |-  ( ph  ->  ( ps  ->  ch ) )   &    |-  ( ph  ->  ( ch  ->  ps ) )   =>    |-  ( ph  ->  ( ps 
 <->  ch ) )
 
Theorembi2 129 Property of the biconditional connective. (Contributed by NM, 11-May-1999.) (Proof shortened by Wolf Lammen, 11-Nov-2012.)
 |-  ( ( ph  <->  ps )  ->  ( ps  ->  ph ) )
 
Theorembicom1 130 Commutative law for equivalence. (Contributed by Wolf Lammen, 10-Nov-2012.)
 |-  ( ( ph  <->  ps )  ->  ( ps 
 <-> 
 ph ) )
 
Theorembicomi 131 Inference from commutative law for logical equivalence. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 16-Sep-2013.)
 |-  ( ph  <->  ps )   =>    |-  ( ps  <->  ph )
 
Theorembiimpri 132 Infer a converse implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 16-Sep-2013.)
 |-  ( ph  <->  ps )   =>    |-  ( ps  ->  ph )
 
Theoremsylibr 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.)
 |-  ( ph  ->  ps )   &    |-  ( ch 
 <->  ps )   =>    |-  ( ph  ->  ch )
 
Theoremsylbir 134 A mixed syllogism inference from a biconditional and an implication. (Contributed by NM, 5-Aug-1993.)
 |-  ( ps  <->  ph )   &    |-  ( ps  ->  ch )   =>    |-  ( ph  ->  ch )
 
Theoremsylbbr 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, syl6bb 195 and variants. (Contributed by BJ, 21-Apr-2019.)

 |-  ( ph  <->  ps )   &    |-  ( ps  <->  ch )   =>    |-  ( ch  ->  ph )
 
Theoremsylbb1 136 A mixed syllogism inference from two biconditionals. (Contributed by BJ, 21-Apr-2019.)
 |-  ( ph  <->  ps )   &    |-  ( ph  <->  ch )   =>    |-  ( ps  ->  ch )
 
Theoremsylbb2 137 A mixed syllogism inference from two biconditionals. (Contributed by BJ, 21-Apr-2019.)
 |-  ( ph  <->  ps )   &    |-  ( ch  <->  ps )   =>    |-  ( ph  ->  ch )
 
Theorempm3.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.)
 |-  ( ph  ->  ( ps  ->  ( ph  /\  ps ) ) )
 
Theorembicom 139 Commutative law for equivalence. Theorem *4.21 of [WhiteheadRussell] p. 117. (Contributed by NM, 5-Aug-1993.) (Revised by NM, 11-Nov-2012.)
 |-  ( ( ph  <->  ps )  <->  ( ps  <->  ph ) )
 
Theorembicomd 140 Commute two sides of a biconditional in a deduction. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   =>    |-  ( ph  ->  ( ch  <->  ps ) )
 
Theoremimpbid1 141 Infer an equivalence from two implications. (Contributed by NM, 6-Mar-2007.)
 |-  ( ph  ->  ( ps  ->  ch ) )   &    |-  ( ch  ->  ps )   =>    |-  ( ph  ->  ( ps 
 <->  ch ) )
 
Theoremimpbid2 142 Infer an equivalence from two implications. (Contributed by NM, 6-Mar-2007.) (Proof shortened by Wolf Lammen, 27-Sep-2013.)
 |-  ( ps  ->  ch )   &    |-  ( ph  ->  ( ch  ->  ps ) )   =>    |-  ( ph  ->  ( ps 
 <->  ch ) )
 
Theorembiimpd 143 Deduce an implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   =>    |-  ( ph  ->  ( ps  ->  ch )
 )
 
Theoremmpbi 144 An inference from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.)
 |-  ph   &    |-  ( ph  <->  ps )   =>    |- 
 ps
 
Theoremmpbir 145 An inference from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.)
 |- 
 ps   &    |-  ( ph  <->  ps )   =>    |-  ph
 
Theoremmpbid 146 A deduction from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ps )   &    |-  ( ph  ->  ( ps  <->  ch ) )   =>    |-  ( ph  ->  ch )
 
Theoremmpbii 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.)
 |- 
 ps   &    |-  ( ph  ->  ( ps 
 <->  ch ) )   =>    |-  ( ph  ->  ch )
 
Theoremsylibd 148 A syllogism deduction. (Contributed by NM, 3-Aug-1994.)
 |-  ( ph  ->  ( ps  ->  ch ) )   &    |-  ( ph  ->  ( ch  <->  th ) )   =>    |-  ( ph  ->  ( ps  ->  th )
 )
 
Theoremsylbid 149 A syllogism deduction. (Contributed by NM, 3-Aug-1994.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   &    |-  ( ph  ->  ( ch  ->  th ) )   =>    |-  ( ph  ->  ( ps  ->  th ) )
 
Theoremmpbidi 150 A deduction from a biconditional, related to modus ponens. (Contributed by NM, 9-Aug-1994.)
 |-  ( th  ->  ( ph  ->  ps ) )   &    |-  ( ph  ->  ( ps  <->  ch ) )   =>    |-  ( th  ->  (
 ph  ->  ch ) )
 
Theoremsyl5bi 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.)
 |-  ( ph  <->  ps )   &    |-  ( ch  ->  ( ps  ->  th )
 )   =>    |-  ( ch  ->  ( ph  ->  th ) )
 
Theoremsyl5bir 152 A mixed syllogism inference from a nested implication and a biconditional. (Contributed by NM, 5-Aug-1993.)
 |-  ( ps  <->  ph )   &    |-  ( ch  ->  ( ps  ->  th )
 )   =>    |-  ( ch  ->  ( ph  ->  th ) )
 
Theoremsyl5ib 153 A mixed syllogism inference. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ps )   &    |-  ( ch  ->  ( ps  <->  th ) )   =>    |-  ( ch  ->  (
 ph  ->  th ) )
 
Theoremsyl5ibcom 154 A mixed syllogism inference. (Contributed by NM, 19-Jun-2007.)
 |-  ( ph  ->  ps )   &    |-  ( ch  ->  ( ps  <->  th ) )   =>    |-  ( ph  ->  ( ch  ->  th )
 )
 
Theoremsyl5ibr 155 A mixed syllogism inference. (Contributed by NM, 3-Apr-1994.) (Revised by NM, 22-Sep-2013.)
 |-  ( ph  ->  th )   &    |-  ( ch  ->  ( ps  <->  th ) )   =>    |-  ( ch  ->  (
 ph  ->  ps ) )
 
Theoremsyl5ibrcom 156 A mixed syllogism inference. (Contributed by NM, 20-Jun-2007.)
 |-  ( ph  ->  th )   &    |-  ( ch  ->  ( ps  <->  th ) )   =>    |-  ( ph  ->  ( ch  ->  ps )
 )
 
Theorembiimprd 157 Deduce a converse implication from a logical equivalence. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 22-Sep-2013.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   =>    |-  ( ph  ->  ( ch  ->  ps )
 )
 
Theorembiimpcd 158 Deduce a commuted implication from a logical equivalence. (Contributed by NM, 3-May-1994.) (Proof shortened by Wolf Lammen, 22-Sep-2013.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   =>    |-  ( ps  ->  (
 ph  ->  ch ) )
 
Theorembiimprcd 159 Deduce a converse commuted implication from a logical equivalence. (Contributed by NM, 3-May-1994.) (Proof shortened by Wolf Lammen, 20-Dec-2013.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   =>    |-  ( ch  ->  (
 ph  ->  ps ) )
 
Theoremsyl6ib 160 A mixed syllogism inference from a nested implication and a biconditional. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ( ps  ->  ch ) )   &    |-  ( ch 
 <-> 
 th )   =>    |-  ( ph  ->  ( ps  ->  th ) )
 
Theoremsyl6ibr 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.)
 |-  ( ph  ->  ( ps  ->  ch ) )   &    |-  ( th 
 <->  ch )   =>    |-  ( ph  ->  ( ps  ->  th ) )
 
Theoremsyl6bi 162 A mixed syllogism inference. (Contributed by NM, 2-Jan-1994.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   &    |-  ( ch  ->  th )   =>    |-  ( ph  ->  ( ps  ->  th ) )
 
Theoremsyl6bir 163 A mixed syllogism inference. (Contributed by NM, 18-May-1994.)
 |-  ( ph  ->  ( ch 
 <->  ps ) )   &    |-  ( ch  ->  th )   =>    |-  ( ph  ->  ( ps  ->  th ) )
 
Theoremsyl7bi 164 A mixed syllogism inference from a doubly nested implication and a biconditional. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  <->  ps )   &    |-  ( ch  ->  ( th  ->  ( ps  ->  ta ) ) )   =>    |-  ( ch  ->  ( th  ->  ( ph  ->  ta )
 ) )
 
Theoremsyl8ib 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.)
 |-  ( ph  ->  ( ps  ->  ( ch  ->  th ) ) )   &    |-  ( th 
 <->  ta )   =>    |-  ( ph  ->  ( ps  ->  ( ch  ->  ta ) ) )
 
Theoremmpbird 166 A deduction from a biconditional, related to modus ponens. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ch )   &    |-  ( ph  ->  ( ps  <->  ch ) )   =>    |-  ( ph  ->  ps )
 
Theoremmpbiri 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.)
 |- 
 ch   &    |-  ( ph  ->  ( ps 
 <->  ch ) )   =>    |-  ( ph  ->  ps )
 
Theoremsylibrd 168 A syllogism deduction. (Contributed by NM, 3-Aug-1994.)
 |-  ( ph  ->  ( ps  ->  ch ) )   &    |-  ( ph  ->  ( th  <->  ch ) )   =>    |-  ( ph  ->  ( ps  ->  th )
 )
 
Theoremsylbird 169 A syllogism deduction. (Contributed by NM, 3-Aug-1994.)
 |-  ( ph  ->  ( ch 
 <->  ps ) )   &    |-  ( ph  ->  ( ch  ->  th ) )   =>    |-  ( ph  ->  ( ps  ->  th ) )
 
Theorembiid 170 Principle of identity for logical equivalence. Theorem *4.2 of [WhiteheadRussell] p. 117. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  <->  ph )
 
Theorembiidd 171 Principle of identity with antecedent. (Contributed by NM, 25-Nov-1995.)
 |-  ( ph  ->  ( ps 
 <->  ps ) )
 
Theorempm5.1im 172 Two propositions are equivalent if they are both true. Closed form of 2th 173. Equivalent to a bi1 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 
( ph  <->  ( ps  <->  ( ph  <->  ps ) ) ). (Contributed by Wolf Lammen, 12-May-2013.)
 |-  ( ph  ->  ( ps  ->  ( ph  <->  ps ) ) )
 
Theorem2th 173 Two truths are equivalent. (Contributed by NM, 18-Aug-1993.)
 |-  ph   &    |- 
 ps   =>    |-  ( ph  <->  ps )
 
Theorem2thd 174 Two truths are equivalent (deduction form). (Contributed by NM, 3-Jun-2012.) (Revised by NM, 29-Jan-2013.)
 |-  ( ph  ->  ps )   &    |-  ( ph  ->  ch )   =>    |-  ( ph  ->  ( ps 
 <->  ch ) )
 
Theoremibi 175 Inference that converts a biconditional implied by one of its arguments, into an implication. (Contributed by NM, 17-Oct-2003.)
 |-  ( ph  ->  ( ph 
 <->  ps ) )   =>    |-  ( ph  ->  ps )
 
Theoremibir 176 Inference that converts a biconditional implied by one of its arguments, into an implication. (Contributed by NM, 22-Jul-2004.)
 |-  ( ph  ->  ( ps 
 <-> 
 ph ) )   =>    |-  ( ph  ->  ps )
 
Theoremibd 177 Deduction that converts a biconditional implied by one of its arguments, into an implication. (Contributed by NM, 26-Jun-2004.)
 |-  ( ph  ->  ( ps  ->  ( ps  <->  ch ) ) )   =>    |-  ( ph  ->  ( ps  ->  ch ) )
 
Theorempm5.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.)
 |-  ( ( ph  ->  ( ps  <->  ch ) )  <->  ( ( ph  ->  ps )  <->  ( ph  ->  ch ) ) )
 
Theorempm5.74i 179 Distribution of implication over biconditional (inference form). (Contributed by NM, 1-Aug-1994.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   =>    |-  ( ( ph  ->  ps )  <->  ( ph  ->  ch ) )
 
Theorempm5.74ri 180 Distribution of implication over biconditional (reverse inference form). (Contributed by NM, 1-Aug-1994.)
 |-  ( ( ph  ->  ps )  <->  ( ph  ->  ch ) )   =>    |-  ( ph  ->  ( ps 
 <->  ch ) )
 
Theorempm5.74d 181 Distribution of implication over biconditional (deduction form). (Contributed by NM, 21-Mar-1996.)
 |-  ( ph  ->  ( ps  ->  ( ch  <->  th ) ) )   =>    |-  ( ph  ->  ( ( ps  ->  ch )  <->  ( ps  ->  th ) ) )
 
Theorempm5.74rd 182 Distribution of implication over biconditional (deduction form). (Contributed by NM, 19-Mar-1997.)
 |-  ( ph  ->  (
 ( ps  ->  ch )  <->  ( ps  ->  th )
 ) )   =>    |-  ( ph  ->  ( ps  ->  ( ch  <->  th ) ) )
 
Theorembitri 183 An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 13-Oct-2012.)
 |-  ( ph  <->  ps )   &    |-  ( ps  <->  ch )   =>    |-  ( ph  <->  ch )
 
Theorembitr2i 184 An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  <->  ps )   &    |-  ( ps  <->  ch )   =>    |-  ( ch  <->  ph )
 
Theorembitr3i 185 An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.)
 |-  ( ps  <->  ph )   &    |-  ( ps  <->  ch )   =>    |-  ( ph  <->  ch )
 
Theorembitr4i 186 An inference from transitive law for logical equivalence. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  <->  ps )   &    |-  ( ch  <->  ps )   =>    |-  ( ph  <->  ch )
 
Theorembitrd 187 Deduction form of bitri 183. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Wolf Lammen, 14-Apr-2013.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   &    |-  ( ph  ->  ( ch  <->  th ) )   =>    |-  ( ph  ->  ( ps  <->  th ) )
 
Theorembitr2d 188 Deduction form of bitr2i 184. (Contributed by NM, 9-Jun-2004.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   &    |-  ( ph  ->  ( ch  <->  th ) )   =>    |-  ( ph  ->  ( th  <->  ps ) )
 
Theorembitr3d 189 Deduction form of bitr3i 185. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   &    |-  ( ph  ->  ( ps  <->  th ) )   =>    |-  ( ph  ->  ( ch  <->  th ) )
 
Theorembitr4d 190 Deduction form of bitr4i 186. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   &    |-  ( ph  ->  ( th  <->  ch ) )   =>    |-  ( ph  ->  ( ps  <->  th ) )
 
Theoremsyl5bb 191 A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  <->  ps )   &    |-  ( ch  ->  ( ps  <->  th ) )   =>    |-  ( ch  ->  (
 ph 
 <-> 
 th ) )
 
Theoremsyl5rbb 192 A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  <->  ps )   &    |-  ( ch  ->  ( ps  <->  th ) )   =>    |-  ( ch  ->  ( th  <->  ph ) )
 
Theoremsyl5bbr 193 A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.)
 |-  ( ps  <->  ph )   &    |-  ( ch  ->  ( ps  <->  th ) )   =>    |-  ( ch  ->  (
 ph 
 <-> 
 th ) )
 
Theoremsyl5rbbr 194 A syllogism inference from two biconditionals. (Contributed by NM, 25-Nov-1994.)
 |-  ( ps  <->  ph )   &    |-  ( ch  ->  ( ps  <->  th ) )   =>    |-  ( ch  ->  ( th  <->  ph ) )
 
Theoremsyl6bb 195 A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   &    |-  ( ch 
 <-> 
 th )   =>    |-  ( ph  ->  ( ps 
 <-> 
 th ) )
 
Theoremsyl6rbb 196 A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   &    |-  ( ch 
 <-> 
 th )   =>    |-  ( ph  ->  ( th 
 <->  ps ) )
 
Theoremsyl6bbr 197 A syllogism inference from two biconditionals. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   &    |-  ( th 
 <->  ch )   =>    |-  ( ph  ->  ( ps 
 <-> 
 th ) )
 
Theoremsyl6rbbr 198 A syllogism inference from two biconditionals. (Contributed by NM, 25-Nov-1994.)
 |-  ( ph  ->  ( ps 
 <->  ch ) )   &    |-  ( th 
 <->  ch )   =>    |-  ( ph  ->  ( th 
 <->  ps ) )
 
Theorem3imtr3i 199 A mixed syllogism inference, useful for removing a definition from both sides of an implication. (Contributed by NM, 10-Aug-1994.)
 |-  ( ph  ->  ps )   &    |-  ( ph 
 <->  ch )   &    |-  ( ps  <->  th )   =>    |-  ( ch  ->  th )
 
Theorem3imtr4i 200 A mixed syllogism inference, useful for applying a definition to both sides of an implication. (Contributed by NM, 5-Aug-1993.)
 |-  ( ph  ->  ps )   &    |-  ( ch 
 <-> 
 ph )   &    |-  ( th  <->  ps )   =>    |-  ( ch  ->  th )
    < Previous  Next >

Page List
Jump to page: Contents  1 1-100101-200 3 201-300 4 301-400 5 401-500 6 501-600 7 601-700 8 701-800 9 801-900 10 901-1000 11 1001-1100 12 1101-1200 13 1201-1300 14 1301-1400 15 1401-1500 16 1501-1600 17 1601-1700 18 1701-1800 19 1801-1900 20 1901-2000 21 2001-2100 22 2101-2200 23 2201-2300 24 2301-2400 25 2401-2500 26 2501-2600 27 2601-2700 28 2701-2800 29 2801-2900 30 2901-3000 31 3001-3100 32 3101-3200 33 3201-3300 34 3301-3400 35 3401-3500 36 3501-3600 37 3601-3700 38 3701-3800 39 3801-3900 40 3901-4000 41 4001-4100 42 4101-4200 43 4201-4300 44 4301-4400 45 4401-4500 46 4501-4600 47 4601-4700 48 4701-4800 49 4801-4900 50 4901-5000 51 5001-5100 52 5101-5200 53 5201-5300 54 5301-5400 55 5401-5500 56 5501-5600 57 5601-5700 58 5701-5800 59 5801-5900 60 5901-6000 61 6001-6100 62 6101-6200 63 6201-6300 64 6301-6400 65 6401-6500 66 6501-6600 67 6601-6700 68 6701-6800 69 6801-6900 70 6901-7000 71 7001-7100 72 7101-7200 73 7201-7300 74 7301-7400 75 7401-7500 76 7501-7600 77 7601-7700 78 7701-7800 79 7801-7900 80 7901-8000 81 8001-8100 82 8101-8200 83 8201-8300 84 8301-8400 85 8401-8500 86 8501-8600 87 8601-8700 88 8701-8800 89 8801-8900 90 8901-9000 91 9001-9100 92 9101-9200 93 9201-9300 94 9301-9400 95 9401-9500 96 9501-9600 97 9601-9700 98 9701-9800 99 9801-9900 100 9901-10000 101 10001-10100 102 10101-10200 103 10201-10300 104 10301-10400 105 10401-10500 106 10501-10600 107 10601-10700 108 10701-10800 109 10801-10900 110 10901-11000 111 11001-11100 112 11101-11200 113 11201-11300 114 11301-11400 115 11401-11500 116 11501-11600 117 11601-11700 118 11701-11800 119 11801-11900 120 11901-12000 121 12001-12100 122 12101-12200 123 12201-12300 124 12301-12400 125 12401-12500 126 12501-12600 127 12601-12700 128 12701-12800 129 12801-12900 130 12901-13000 131 13001-13100 132 13101-13177
  Copyright terms: Public domain < Previous  Next >