MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  difxp Structured version   Visualization version   GIF version

Theorem difxp 6067
Description: Difference of Cartesian products, expressed in terms of a union of Cartesian products of differences. (Contributed by Jeff Madsen, 2-Sep-2009.) (Revised by Mario Carneiro, 26-Jun-2014.)
Assertion
Ref Expression
difxp ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) = (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵)))

Proof of Theorem difxp
Dummy variables 𝑥 𝑦 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 difss 4066 . . 3 ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) ⊆ (𝐶 × 𝐷)
2 relxp 5607 . . 3 Rel (𝐶 × 𝐷)
3 relss 5692 . . 3 (((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) ⊆ (𝐶 × 𝐷) → (Rel (𝐶 × 𝐷) → Rel ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵))))
41, 2, 3mp2 9 . 2 Rel ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵))
5 relxp 5607 . . 3 Rel ((𝐶𝐴) × 𝐷)
6 relxp 5607 . . 3 Rel (𝐶 × (𝐷𝐵))
7 relun 5721 . . 3 (Rel (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵))) ↔ (Rel ((𝐶𝐴) × 𝐷) ∧ Rel (𝐶 × (𝐷𝐵))))
85, 6, 7mpbir2an 708 . 2 Rel (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵)))
9 ianor 979 . . . . . 6 (¬ (𝑥𝐴𝑦𝐵) ↔ (¬ 𝑥𝐴 ∨ ¬ 𝑦𝐵))
109anbi2i 623 . . . . 5 (((𝑥𝐶𝑦𝐷) ∧ ¬ (𝑥𝐴𝑦𝐵)) ↔ ((𝑥𝐶𝑦𝐷) ∧ (¬ 𝑥𝐴 ∨ ¬ 𝑦𝐵)))
11 andi 1005 . . . . 5 (((𝑥𝐶𝑦𝐷) ∧ (¬ 𝑥𝐴 ∨ ¬ 𝑦𝐵)) ↔ (((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴) ∨ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑦𝐵)))
1210, 11bitri 274 . . . 4 (((𝑥𝐶𝑦𝐷) ∧ ¬ (𝑥𝐴𝑦𝐵)) ↔ (((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴) ∨ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑦𝐵)))
13 opelxp 5625 . . . . 5 (⟨𝑥, 𝑦⟩ ∈ (𝐶 × 𝐷) ↔ (𝑥𝐶𝑦𝐷))
14 opelxp 5625 . . . . . 6 (⟨𝑥, 𝑦⟩ ∈ (𝐴 × 𝐵) ↔ (𝑥𝐴𝑦𝐵))
1514notbii 320 . . . . 5 (¬ ⟨𝑥, 𝑦⟩ ∈ (𝐴 × 𝐵) ↔ ¬ (𝑥𝐴𝑦𝐵))
1613, 15anbi12i 627 . . . 4 ((⟨𝑥, 𝑦⟩ ∈ (𝐶 × 𝐷) ∧ ¬ ⟨𝑥, 𝑦⟩ ∈ (𝐴 × 𝐵)) ↔ ((𝑥𝐶𝑦𝐷) ∧ ¬ (𝑥𝐴𝑦𝐵)))
17 opelxp 5625 . . . . . 6 (⟨𝑥, 𝑦⟩ ∈ ((𝐶𝐴) × 𝐷) ↔ (𝑥 ∈ (𝐶𝐴) ∧ 𝑦𝐷))
18 eldif 3897 . . . . . . . 8 (𝑥 ∈ (𝐶𝐴) ↔ (𝑥𝐶 ∧ ¬ 𝑥𝐴))
1918anbi1i 624 . . . . . . 7 ((𝑥 ∈ (𝐶𝐴) ∧ 𝑦𝐷) ↔ ((𝑥𝐶 ∧ ¬ 𝑥𝐴) ∧ 𝑦𝐷))
20 an32 643 . . . . . . 7 (((𝑥𝐶 ∧ ¬ 𝑥𝐴) ∧ 𝑦𝐷) ↔ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴))
2119, 20bitri 274 . . . . . 6 ((𝑥 ∈ (𝐶𝐴) ∧ 𝑦𝐷) ↔ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴))
2217, 21bitri 274 . . . . 5 (⟨𝑥, 𝑦⟩ ∈ ((𝐶𝐴) × 𝐷) ↔ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴))
23 eldif 3897 . . . . . . 7 (𝑦 ∈ (𝐷𝐵) ↔ (𝑦𝐷 ∧ ¬ 𝑦𝐵))
2423anbi2i 623 . . . . . 6 ((𝑥𝐶𝑦 ∈ (𝐷𝐵)) ↔ (𝑥𝐶 ∧ (𝑦𝐷 ∧ ¬ 𝑦𝐵)))
25 opelxp 5625 . . . . . 6 (⟨𝑥, 𝑦⟩ ∈ (𝐶 × (𝐷𝐵)) ↔ (𝑥𝐶𝑦 ∈ (𝐷𝐵)))
26 anass 469 . . . . . 6 (((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑦𝐵) ↔ (𝑥𝐶 ∧ (𝑦𝐷 ∧ ¬ 𝑦𝐵)))
2724, 25, 263bitr4i 303 . . . . 5 (⟨𝑥, 𝑦⟩ ∈ (𝐶 × (𝐷𝐵)) ↔ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑦𝐵))
2822, 27orbi12i 912 . . . 4 ((⟨𝑥, 𝑦⟩ ∈ ((𝐶𝐴) × 𝐷) ∨ ⟨𝑥, 𝑦⟩ ∈ (𝐶 × (𝐷𝐵))) ↔ (((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴) ∨ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑦𝐵)))
2912, 16, 283bitr4i 303 . . 3 ((⟨𝑥, 𝑦⟩ ∈ (𝐶 × 𝐷) ∧ ¬ ⟨𝑥, 𝑦⟩ ∈ (𝐴 × 𝐵)) ↔ (⟨𝑥, 𝑦⟩ ∈ ((𝐶𝐴) × 𝐷) ∨ ⟨𝑥, 𝑦⟩ ∈ (𝐶 × (𝐷𝐵))))
30 eldif 3897 . . 3 (⟨𝑥, 𝑦⟩ ∈ ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) ↔ (⟨𝑥, 𝑦⟩ ∈ (𝐶 × 𝐷) ∧ ¬ ⟨𝑥, 𝑦⟩ ∈ (𝐴 × 𝐵)))
31 elun 4083 . . 3 (⟨𝑥, 𝑦⟩ ∈ (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵))) ↔ (⟨𝑥, 𝑦⟩ ∈ ((𝐶𝐴) × 𝐷) ∨ ⟨𝑥, 𝑦⟩ ∈ (𝐶 × (𝐷𝐵))))
3229, 30, 313bitr4i 303 . 2 (⟨𝑥, 𝑦⟩ ∈ ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) ↔ ⟨𝑥, 𝑦⟩ ∈ (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵))))
334, 8, 32eqrelriiv 5700 1 ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) = (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵)))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wa 396  wo 844   = wceq 1539  wcel 2106  cdif 3884  cun 3885  wss 3887  cop 4567   × cxp 5587  Rel wrel 5594
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1798  ax-4 1812  ax-5 1913  ax-6 1971  ax-7 2011  ax-8 2108  ax-9 2116  ax-ext 2709  ax-sep 5223  ax-nul 5230  ax-pr 5352
This theorem depends on definitions:  df-bi 206  df-an 397  df-or 845  df-3an 1088  df-tru 1542  df-fal 1552  df-ex 1783  df-sb 2068  df-clab 2716  df-cleq 2730  df-clel 2816  df-ral 3069  df-rex 3070  df-rab 3073  df-v 3434  df-dif 3890  df-un 3892  df-in 3894  df-ss 3904  df-nul 4257  df-if 4460  df-sn 4562  df-pr 4564  df-op 4568  df-opab 5137  df-xp 5595  df-rel 5596
This theorem is referenced by:  difxp1  6068  difxp2  6069  evlslem4  21284  txcld  22754  suppovss  31017
  Copyright terms: Public domain W3C validator