Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
Mirrors > Home > MPE Home > Th. List > had0 | Structured version Visualization version GIF version |
Description: If the first input is false, then the adder sum is equivalent to the exclusive disjunction of the other two inputs. (Contributed by Mario Carneiro, 4-Sep-2016.) (Proof shortened by Wolf Lammen, 12-Jul-2020.) |
Ref | Expression |
---|---|
had0 | ⊢ (¬ 𝜑 → (hadd(𝜑, 𝜓, 𝜒) ↔ (𝜓 ⊻ 𝜒))) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | had1 1606 | . . 3 ⊢ (¬ 𝜑 → (hadd(¬ 𝜑, ¬ 𝜓, ¬ 𝜒) ↔ (¬ 𝜓 ↔ ¬ 𝜒))) | |
2 | hadnot 1605 | . . 3 ⊢ (¬ hadd(𝜑, 𝜓, 𝜒) ↔ hadd(¬ 𝜑, ¬ 𝜓, ¬ 𝜒)) | |
3 | xnor 1505 | . . . 4 ⊢ ((𝜓 ↔ 𝜒) ↔ ¬ (𝜓 ⊻ 𝜒)) | |
4 | notbi 318 | . . . 4 ⊢ ((𝜓 ↔ 𝜒) ↔ (¬ 𝜓 ↔ ¬ 𝜒)) | |
5 | 3, 4 | bitr3i 276 | . . 3 ⊢ (¬ (𝜓 ⊻ 𝜒) ↔ (¬ 𝜓 ↔ ¬ 𝜒)) |
6 | 1, 2, 5 | 3bitr4g 313 | . 2 ⊢ (¬ 𝜑 → (¬ hadd(𝜑, 𝜓, 𝜒) ↔ ¬ (𝜓 ⊻ 𝜒))) |
7 | 6 | con4bid 316 | 1 ⊢ (¬ 𝜑 → (hadd(𝜑, 𝜓, 𝜒) ↔ (𝜓 ⊻ 𝜒))) |
Colors of variables: wff setvar class |
Syntax hints: ¬ wn 3 → wi 4 ↔ wb 205 ⊻ wxo 1503 haddwhad 1595 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 |
This theorem depends on definitions: df-bi 206 df-xor 1504 df-had 1596 |
This theorem is referenced by: hadifp 1608 sadadd2lem2 16085 saddisjlem 16099 |
Copyright terms: Public domain | W3C validator |