Home | New
Foundations Explorer Theorem List (p. 14 of 64) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > NFE Home Page > Theorem List Contents This page: Page List |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | nanbi1d 1301 | Introduce a right anti-conjunct to both sides of a logical equivalence. (Contributed by SF, 2-Jan-2018.) |
⊢ (φ → (ψ ↔ χ)) ⇒ ⊢ (φ → ((ψ ⊼ θ) ↔ (χ ⊼ θ))) | ||
Theorem | nanbi2d 1302 | Introduce a left anti-conjunct to both sides of a logical equivalence. (Contributed by SF, 2-Jan-2018.) |
⊢ (φ → (ψ ↔ χ)) ⇒ ⊢ (φ → ((θ ⊼ ψ) ↔ (θ ⊼ χ))) | ||
Theorem | nanbi12d 1303 | Join two logical equivalences with anti-conjunction. (Contributed by Scott Fenton, 2-Jan-2018.) |
⊢ (φ → (ψ ↔ χ)) & ⊢ (φ → (θ ↔ τ)) ⇒ ⊢ (φ → ((ψ ⊼ θ) ↔ (χ ⊼ τ))) | ||
Syntax | wxo 1304 | Extend wff definition to include exclusive disjunction ('xor'). |
wff (φ ⊻ ψ) | ||
Definition | df-xor 1305 | Define exclusive disjunction (logical 'xor'). Return true if either the left or right, but not both, are true. After we define the constant true ⊤ (df-tru 1319) and the constant false ⊥ (df-fal 1320), we will be able to prove these truth table values: (( ⊤ ⊻ ⊤ ) ↔ ⊥ ) (truxortru 1358), (( ⊤ ⊻ ⊥ ) ↔ ⊤ ) (truxorfal 1359), (( ⊥ ⊻ ⊤ ) ↔ ⊤ ) (falxortru 1360), and (( ⊥ ⊻ ⊥ ) ↔ ⊥ ) (falxorfal 1361). Contrast with ∧ (df-an 360), ∨ (df-or 359), → (wi 4), and ⊼ (df-nan 1288) . (Contributed by FL, 22-Nov-2010.) |
⊢ ((φ ⊻ ψ) ↔ ¬ (φ ↔ ψ)) | ||
Theorem | xnor 1306 | Two ways to write XNOR. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ ((φ ↔ ψ) ↔ ¬ (φ ⊻ ψ)) | ||
Theorem | xorcom 1307 | ⊻ is commutative. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ ((φ ⊻ ψ) ↔ (ψ ⊻ φ)) | ||
Theorem | xorass 1308 | ⊻ is associative. (Contributed by FL, 22-Nov-2010.) (Proof shortened by Andrew Salmon, 8-Jun-2011.) |
⊢ (((φ ⊻ ψ) ⊻ χ) ↔ (φ ⊻ (ψ ⊻ χ))) | ||
Theorem | excxor 1309 | This tautology shows that xor is really exclusive. (Contributed by FL, 22-Nov-2010.) |
⊢ ((φ ⊻ ψ) ↔ ((φ ∧ ¬ ψ) ∨ (¬ φ ∧ ψ))) | ||
Theorem | xor2 1310 | Two ways to express "exclusive or." (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ ((φ ⊻ ψ) ↔ ((φ ∨ ψ) ∧ ¬ (φ ∧ ψ))) | ||
Theorem | xorneg1 1311 | ⊻ is negated under negation of one argument. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ ((¬ φ ⊻ ψ) ↔ ¬ (φ ⊻ ψ)) | ||
Theorem | xorneg2 1312 | ⊻ is negated under negation of one argument. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ ((φ ⊻ ¬ ψ) ↔ ¬ (φ ⊻ ψ)) | ||
Theorem | xorneg 1313 | ⊻ is unchanged under negation of both arguments. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ ((¬ φ ⊻ ¬ ψ) ↔ (φ ⊻ ψ)) | ||
Theorem | xorbi12i 1314 | Equality property for XOR. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (φ ↔ ψ) & ⊢ (χ ↔ θ) ⇒ ⊢ ((φ ⊻ χ) ↔ (ψ ⊻ θ)) | ||
Theorem | xorbi12d 1315 | Equality property for XOR. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (φ → (ψ ↔ χ)) & ⊢ (φ → (θ ↔ τ)) ⇒ ⊢ (φ → ((ψ ⊻ θ) ↔ (χ ⊻ τ))) | ||
Syntax | wtru 1316 | ⊤ is a wff. |
wff ⊤ | ||
Syntax | wfal 1317 | ⊥ is a wff. |
wff ⊥ | ||
Theorem | trujust 1318 | Soundness justification theorem for df-tru 1319. (Contributed by Mario Carneiro, 17-Nov-2013.) |
⊢ ((φ ↔ φ) ↔ (ψ ↔ ψ)) | ||
Definition | df-tru 1319 | Definition of ⊤, a tautology. ⊤ is a constant true. In this definition biid 227 is used as an antecedent, however, any true wff, such as an axiom, can be used in its place. (Contributed by Anthony Hart, 13-Oct-2010.) |
⊢ ( ⊤ ↔ (φ ↔ φ)) | ||
Definition | df-fal 1320 | Definition of ⊥, a contradiction. ⊥ is a constant false. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ ( ⊥ ↔ ¬ ⊤ ) | ||
Theorem | tru 1321 | ⊤ is provable. (Contributed by Anthony Hart, 13-Oct-2010.) |
⊢ ⊤ | ||
Theorem | fal 1322 | ⊥ is refutable. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Mel L. O'Cat, 11-Mar-2012.) |
⊢ ¬ ⊥ | ||
Theorem | trud 1323 | Eliminate ⊤ as an antecedent. (Contributed by Mario Carneiro, 13-Mar-2014.) |
⊢ ( ⊤ → φ) ⇒ ⊢ φ | ||
Theorem | tbtru 1324 | If something is true, it outputs ⊤. (Contributed by Anthony Hart, 14-Aug-2011.) |
⊢ (φ ↔ (φ ↔ ⊤ )) | ||
Theorem | nbfal 1325 | If something is not true, it outputs ⊥. (Contributed by Anthony Hart, 14-Aug-2011.) |
⊢ (¬ φ ↔ (φ ↔ ⊥ )) | ||
Theorem | bitru 1326 | A theorem is equivalent to truth. (Contributed by Mario Carneiro, 9-May-2015.) |
⊢ φ ⇒ ⊢ (φ ↔ ⊤ ) | ||
Theorem | bifal 1327 | A contradiction is equivalent to falsehood. (Contributed by Mario Carneiro, 9-May-2015.) |
⊢ ¬ φ ⇒ ⊢ (φ ↔ ⊥ ) | ||
Theorem | falim 1328 | ⊥ implies anything. (Contributed by FL, 20-Mar-2011.) (Proof shortened by Anthony Hart, 1-Aug-2011.) |
⊢ ( ⊥ → φ) | ||
Theorem | falimd 1329 | ⊥ implies anything. (Contributed by Mario Carneiro, 9-Feb-2017.) |
⊢ ((φ ∧ ⊥ ) → ψ) | ||
Theorem | a1tru 1330 | Anything implies ⊤. (Contributed by FL, 20-Mar-2011.) (Proof shortened by Anthony Hart, 1-Aug-2011.) |
⊢ (φ → ⊤ ) | ||
Theorem | truan 1331 | True can be removed from a conjunction. (Contributed by FL, 20-Mar-2011.) |
⊢ (( ⊤ ∧ φ) ↔ φ) | ||
Theorem | dfnot 1332 | Given falsum, we can define the negation of a wff φ as the statement that a contradiction follows from assuming φ. (Contributed by Mario Carneiro, 9-Feb-2017.) |
⊢ (¬ φ ↔ (φ → ⊥ )) | ||
Theorem | inegd 1333 | Negation introduction rule from natural deduction. (Contributed by Mario Carneiro, 9-Feb-2017.) |
⊢ ((φ ∧ ψ) → ⊥ ) ⇒ ⊢ (φ → ¬ ψ) | ||
Theorem | efald 1334 | Deduction based on reductio ad absurdum. (Contributed by Mario Carneiro, 9-Feb-2017.) |
⊢ ((φ ∧ ¬ ψ) → ⊥ ) ⇒ ⊢ (φ → ψ) | ||
Theorem | pm2.21fal 1335 | If a wff and its negation are provable, then falsum is provable. (Contributed by Mario Carneiro, 9-Feb-2017.) |
⊢ (φ → ψ) & ⊢ (φ → ¬ ψ) ⇒ ⊢ (φ → ⊥ ) | ||
Some sources define operations on true/false values using truth tables. These tables show the results of their operations for all possible combinations of true (⊤) and false (⊥). Here we show that our definitions and axioms produce equivalent results for ∧ (conjunction aka logical 'and') df-an 360, ∨ (disjunction aka logical inclusive 'or') df-or 359, → (implies) wi 4, ¬ (not) wn 3, ↔ (logical equivalence) df-bi 177, ⊼ (nand aka Sheffer stroke) df-nan 1288, and ⊻ (exclusive or) df-xor 1305. | ||
Theorem | truantru 1336 | A ∧ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (( ⊤ ∧ ⊤ ) ↔ ⊤ ) | ||
Theorem | truanfal 1337 | A ∧ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (( ⊤ ∧ ⊥ ) ↔ ⊥ ) | ||
Theorem | falantru 1338 | A ∧ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (( ⊥ ∧ ⊤ ) ↔ ⊥ ) | ||
Theorem | falanfal 1339 | A ∧ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (( ⊥ ∧ ⊥ ) ↔ ⊥ ) | ||
Theorem | truortru 1340 | A ∨ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊤ ∨ ⊤ ) ↔ ⊤ ) | ||
Theorem | truorfal 1341 | A ∨ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (( ⊤ ∨ ⊥ ) ↔ ⊤ ) | ||
Theorem | falortru 1342 | A ∨ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (( ⊥ ∨ ⊤ ) ↔ ⊤ ) | ||
Theorem | falorfal 1343 | A ∨ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊥ ∨ ⊥ ) ↔ ⊥ ) | ||
Theorem | truimtru 1344 | A → identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (( ⊤ → ⊤ ) ↔ ⊤ ) | ||
Theorem | truimfal 1345 | A → identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊤ → ⊥ ) ↔ ⊥ ) | ||
Theorem | falimtru 1346 | A → identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (( ⊥ → ⊤ ) ↔ ⊤ ) | ||
Theorem | falimfal 1347 | A → identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (( ⊥ → ⊥ ) ↔ ⊤ ) | ||
Theorem | nottru 1348 | A ¬ identity. (Contributed by Anthony Hart, 22-Oct-2010.) |
⊢ (¬ ⊤ ↔ ⊥ ) | ||
Theorem | notfal 1349 | A ¬ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (¬ ⊥ ↔ ⊤ ) | ||
Theorem | trubitru 1350 | A ↔ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊤ ↔ ⊤ ) ↔ ⊤ ) | ||
Theorem | trubifal 1351 | A ↔ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊤ ↔ ⊥ ) ↔ ⊥ ) | ||
Theorem | falbitru 1352 | A ↔ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊥ ↔ ⊤ ) ↔ ⊥ ) | ||
Theorem | falbifal 1353 | A ↔ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊥ ↔ ⊥ ) ↔ ⊤ ) | ||
Theorem | trunantru 1354 | A ⊼ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊤ ⊼ ⊤ ) ↔ ⊥ ) | ||
Theorem | trunanfal 1355 | A ⊼ identity. (Contributed by Anthony Hart, 23-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊤ ⊼ ⊥ ) ↔ ⊤ ) | ||
Theorem | falnantru 1356 | A ⊼ identity. (Contributed by Anthony Hart, 23-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊥ ⊼ ⊤ ) ↔ ⊤ ) | ||
Theorem | falnanfal 1357 | A ⊼ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.) |
⊢ (( ⊥ ⊼ ⊥ ) ↔ ⊤ ) | ||
Theorem | truxortru 1358 | A ⊻ identity. (Contributed by David A. Wheeler, 8-May-2015.) |
⊢ (( ⊤ ⊻ ⊤ ) ↔ ⊥ ) | ||
Theorem | truxorfal 1359 | A ⊻ identity. (Contributed by David A. Wheeler, 8-May-2015.) |
⊢ (( ⊤ ⊻ ⊥ ) ↔ ⊤ ) | ||
Theorem | falxortru 1360 | A ⊻ identity. (Contributed by David A. Wheeler, 9-May-2015.) |
⊢ (( ⊥ ⊻ ⊤ ) ↔ ⊤ ) | ||
Theorem | falxorfal 1361 | A ⊻ identity. (Contributed by David A. Wheeler, 9-May-2015.) |
⊢ (( ⊥ ⊻ ⊥ ) ↔ ⊥ ) | ||
Theorem | ee22 1362 | Virtual deduction rule e22 in set.mm without virtual deduction connectives. Special theorem needed for Alan Sare's virtual deduction translation tool. (Contributed by Alan Sare, 2-May-2011.) (New usage is discouraged.) TODO: decide if this is worth keeping. |
⊢ (φ → (ψ → χ)) & ⊢ (φ → (ψ → θ)) & ⊢ (χ → (θ → τ)) ⇒ ⊢ (φ → (ψ → τ)) | ||
Theorem | ee12an 1363 | e12an in set.mm without virtual deduction connectives. Special theorem needed for Alan Sare's virtual deduction translation tool. (Contributed by Alan Sare, 28-Oct-2011.) TODO: this is frequently used; come up with better label. |
⊢ (φ → ψ) & ⊢ (φ → (χ → θ)) & ⊢ ((ψ ∧ θ) → τ) ⇒ ⊢ (φ → (χ → τ)) | ||
Theorem | ee23 1364 | e23 in set.mm without virtual deductions. (Contributed by Alan Sare, 17-Jul-2011.) (New usage is discouraged.) TODO: decide if this is worth keeping. |
⊢ (φ → (ψ → χ)) & ⊢ (φ → (ψ → (θ → τ))) & ⊢ (χ → (τ → η)) ⇒ ⊢ (φ → (ψ → (θ → η))) | ||
Theorem | exbir 1365 | Exportation implication also converting head from biconditional to conditional. This proof is exbirVD in set.mm automatically translated and minimized. (Contributed by Alan Sare, 31-Dec-2011.) (New usage is discouraged.) TODO: decide if this is worth keeping. |
⊢ (((φ ∧ ψ) → (χ ↔ θ)) → (φ → (ψ → (θ → χ)))) | ||
Theorem | 3impexp 1366 | impexp 433 with a 3-conjunct antecedent. (Contributed by Alan Sare, 31-Dec-2011.) |
⊢ (((φ ∧ ψ ∧ χ) → θ) ↔ (φ → (ψ → (χ → θ)))) | ||
Theorem | 3impexpbicom 1367 | 3impexp 1366 with biconditional consequent of antecedent that is commuted in consequent. Derived automatically from 3impexpVD in set.mm. (Contributed by Alan Sare, 31-Dec-2011.) (New usage is discouraged.) TODO: decide if this is worth keeping. |
⊢ (((φ ∧ ψ ∧ χ) → (θ ↔ τ)) ↔ (φ → (ψ → (χ → (τ ↔ θ))))) | ||
Theorem | 3impexpbicomi 1368 | Deduction form of 3impexpbicom 1367. Derived automatically from 3impexpbicomiVD in set.mm. (Contributed by Alan Sare, 31-Dec-2011.) (New usage is discouraged.) TODO: decide if this is worth keeping. |
⊢ ((φ ∧ ψ ∧ χ) → (θ ↔ τ)) ⇒ ⊢ (φ → (ψ → (χ → (τ ↔ θ)))) | ||
Theorem | ancomsimp 1369 | Closed form of ancoms 439. Derived automatically from ancomsimpVD in set.mm. (Contributed by Alan Sare, 31-Dec-2011.) |
⊢ (((φ ∧ ψ) → χ) ↔ ((ψ ∧ φ) → χ)) | ||
Theorem | exp3acom3r 1370 | Export and commute antecedents. (Contributed by Alan Sare, 18-Mar-2012.) |
⊢ (φ → ((ψ ∧ χ) → θ)) ⇒ ⊢ (ψ → (χ → (φ → θ))) | ||
Theorem | exp3acom23g 1371 | Implication form of exp3acom23 1372. (Contributed by Alan Sare, 22-Jul-2012.) (New usage is discouraged.) TODO: decide if this is worth keeping. |
⊢ ((φ → ((ψ ∧ χ) → θ)) ↔ (φ → (χ → (ψ → θ)))) | ||
Theorem | exp3acom23 1372 | The exportation deduction exp3a 425 with commutation of the conjoined wwfs. (Contributed by Alan Sare, 22-Jul-2012.) |
⊢ (φ → ((ψ ∧ χ) → θ)) ⇒ ⊢ (φ → (χ → (ψ → θ))) | ||
Theorem | simplbi2comg 1373 | Implication form of simplbi2com 1374. (Contributed by Alan Sare, 22-Jul-2012.) (New usage is discouraged.) TODO: decide if this is worth keeping. |
⊢ ((φ ↔ (ψ ∧ χ)) → (χ → (ψ → φ))) | ||
Theorem | simplbi2com 1374 | A deduction eliminating a conjunct, similar to simplbi2 608. (Contributed by Alan Sare, 22-Jul-2012.) (Proof shortened by Wolf Lammen, 10-Nov-2012.) |
⊢ (φ ↔ (ψ ∧ χ)) ⇒ ⊢ (χ → (ψ → φ)) | ||
Theorem | ee21 1375 | e21 in set.mm without virtual deductions. (Contributed by Alan Sare, 18-Mar-2012.) (New usage is discouraged.) TODO: decide if this is worth keeping. |
⊢ (φ → (ψ → χ)) & ⊢ (φ → θ) & ⊢ (χ → (θ → τ)) ⇒ ⊢ (φ → (ψ → τ)) | ||
Theorem | ee10 1376 | e10 in set.mm without virtual deductions. (Contributed by Alan Sare, 25-Jul-2011.) TODO: this is frequently used; come up with better label. |
⊢ (φ → ψ) & ⊢ χ & ⊢ (ψ → (χ → θ)) ⇒ ⊢ (φ → θ) | ||
Theorem | ee02 1377 | e02 in set.mm without virtual deductions. (Contributed by Alan Sare, 22-Jul-2012.) (New usage is discouraged.) TODO: decide if this is worth keeping. |
⊢ φ & ⊢ (ψ → (χ → θ)) & ⊢ (φ → (θ → τ)) ⇒ ⊢ (ψ → (χ → τ)) | ||
Propositional calculus deals with truth values, which can be interpreted as bits. Using this, we can define the half-adder in pure propositional calculus, and show its basic properties. | ||
Syntax | whad 1378 | Define the half adder (triple XOR). (Contributed by Mario Carneiro, 4-Sep-2016.) |
wff hadd(φ, ψ, χ) | ||
Syntax | wcad 1379 | Define the half adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.) |
wff cadd(φ, ψ, χ) | ||
Definition | df-had 1380 | Define the half adder (triple XOR). (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (hadd(φ, ψ, χ) ↔ ((φ ⊻ ψ) ⊻ χ)) | ||
Definition | df-cad 1381 | Define the half adder carry, which is true when at least two arguments are true. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (cadd(φ, ψ, χ) ↔ ((φ ∧ ψ) ∨ (χ ∧ (φ ⊻ ψ)))) | ||
Theorem | hadbi123d 1382 | Equality theorem for half adder. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (φ → (ψ ↔ χ)) & ⊢ (φ → (θ ↔ τ)) & ⊢ (φ → (η ↔ ζ)) ⇒ ⊢ (φ → (hadd(ψ, θ, η) ↔ hadd(χ, τ, ζ))) | ||
Theorem | cadbi123d 1383 | Equality theorem for adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (φ → (ψ ↔ χ)) & ⊢ (φ → (θ ↔ τ)) & ⊢ (φ → (η ↔ ζ)) ⇒ ⊢ (φ → (cadd(ψ, θ, η) ↔ cadd(χ, τ, ζ))) | ||
Theorem | hadbi123i 1384 | Equality theorem for half adder. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (φ ↔ ψ) & ⊢ (χ ↔ θ) & ⊢ (τ ↔ η) ⇒ ⊢ (hadd(φ, χ, τ) ↔ hadd(ψ, θ, η)) | ||
Theorem | cadbi123i 1385 | Equality theorem for adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (φ ↔ ψ) & ⊢ (χ ↔ θ) & ⊢ (τ ↔ η) ⇒ ⊢ (cadd(φ, χ, τ) ↔ cadd(ψ, θ, η)) | ||
Theorem | hadass 1386 | Associative law for triple XOR. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (hadd(φ, ψ, χ) ↔ (φ ⊻ (ψ ⊻ χ))) | ||
Theorem | hadbi 1387 | The half adder is the same as the triple biconditional. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (hadd(φ, ψ, χ) ↔ ((φ ↔ ψ) ↔ χ)) | ||
Theorem | hadcoma 1388 | Commutative law for triple XOR. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (hadd(φ, ψ, χ) ↔ hadd(ψ, φ, χ)) | ||
Theorem | hadcomb 1389 | Commutative law for triple XOR. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (hadd(φ, ψ, χ) ↔ hadd(φ, χ, ψ)) | ||
Theorem | hadrot 1390 | Rotation law for triple XOR. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (hadd(φ, ψ, χ) ↔ hadd(ψ, χ, φ)) | ||
Theorem | cador 1391 | Write the adder carry in disjunctive normal form. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (cadd(φ, ψ, χ) ↔ ((φ ∧ ψ) ∨ (φ ∧ χ) ∨ (ψ ∧ χ))) | ||
Theorem | cadan 1392 | Write the adder carry in conjunctive normal form. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (cadd(φ, ψ, χ) ↔ ((φ ∨ ψ) ∧ (φ ∨ χ) ∧ (ψ ∨ χ))) | ||
Theorem | hadnot 1393 | The half adder distributes over negation. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (¬ hadd(φ, ψ, χ) ↔ hadd(¬ φ, ¬ ψ, ¬ χ)) | ||
Theorem | cadnot 1394 | The adder carry distributes over negation. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (¬ cadd(φ, ψ, χ) ↔ cadd(¬ φ, ¬ ψ, ¬ χ)) | ||
Theorem | cadcoma 1395 | Commutative law for adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (cadd(φ, ψ, χ) ↔ cadd(ψ, φ, χ)) | ||
Theorem | cadcomb 1396 | Commutative law for adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (cadd(φ, ψ, χ) ↔ cadd(φ, χ, ψ)) | ||
Theorem | cadrot 1397 | Rotation law for adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ (cadd(φ, ψ, χ) ↔ cadd(ψ, χ, φ)) | ||
Theorem | cad1 1398 | If one parameter is true, the adder carry is true exactly when at least one of the other parameters is true. (Contributed by Mario Carneiro, 8-Sep-2016.) |
⊢ (χ → (cadd(φ, ψ, χ) ↔ (φ ∨ ψ))) | ||
Theorem | cad11 1399 | If two parameters are true, the adder carry is true. (Contributed by Mario Carneiro, 4-Sep-2016.) |
⊢ ((φ ∧ ψ) → cadd(φ, ψ, χ)) | ||
Theorem | cad0 1400 | If one parameter is false, the adder carry is true exactly when both of the other two parameters are true. (Contributed by Mario Carneiro, 8-Sep-2016.) |
⊢ (¬ χ → (cadd(φ, ψ, χ) ↔ (φ ∧ ψ))) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |