Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
Mirrors > Home > MPE Home > Th. List > cadcoma | Structured version Visualization version GIF version |
Description: Commutative law for the adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.) |
Ref | Expression |
---|---|
cadcoma | ⊢ (cadd(𝜑, 𝜓, 𝜒) ↔ cadd(𝜓, 𝜑, 𝜒)) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | ancom 460 | . . 3 ⊢ ((𝜑 ∧ 𝜓) ↔ (𝜓 ∧ 𝜑)) | |
2 | xorcom 1506 | . . . 4 ⊢ ((𝜑 ⊻ 𝜓) ↔ (𝜓 ⊻ 𝜑)) | |
3 | 2 | anbi2i 622 | . . 3 ⊢ ((𝜒 ∧ (𝜑 ⊻ 𝜓)) ↔ (𝜒 ∧ (𝜓 ⊻ 𝜑))) |
4 | 1, 3 | orbi12i 911 | . 2 ⊢ (((𝜑 ∧ 𝜓) ∨ (𝜒 ∧ (𝜑 ⊻ 𝜓))) ↔ ((𝜓 ∧ 𝜑) ∨ (𝜒 ∧ (𝜓 ⊻ 𝜑)))) |
5 | df-cad 1610 | . 2 ⊢ (cadd(𝜑, 𝜓, 𝜒) ↔ ((𝜑 ∧ 𝜓) ∨ (𝜒 ∧ (𝜑 ⊻ 𝜓)))) | |
6 | df-cad 1610 | . 2 ⊢ (cadd(𝜓, 𝜑, 𝜒) ↔ ((𝜓 ∧ 𝜑) ∨ (𝜒 ∧ (𝜓 ⊻ 𝜑)))) | |
7 | 4, 5, 6 | 3bitr4i 302 | 1 ⊢ (cadd(𝜑, 𝜓, 𝜒) ↔ cadd(𝜓, 𝜑, 𝜒)) |
Colors of variables: wff setvar class |
Syntax hints: ↔ wb 205 ∧ wa 395 ∨ wo 843 ⊻ wxo 1503 caddwcad 1609 |
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-an 396 df-or 844 df-xor 1504 df-cad 1610 |
This theorem is referenced by: cadrot 1617 sadcom 16098 |
Copyright terms: Public domain | W3C validator |