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

Theorem difxp 6184
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.) (Proof shortened by Wolf Lammen, 16-May-2025.)
Assertion
Ref Expression
difxp ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) = (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵)))

Proof of Theorem difxp
Dummy variables 𝑥 𝑦 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 difss 4136 . . 3 ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) ⊆ (𝐶 × 𝐷)
2 relxp 5703 . . 3 Rel (𝐶 × 𝐷)
3 relss 5791 . . 3 (((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) ⊆ (𝐶 × 𝐷) → (Rel (𝐶 × 𝐷) → Rel ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵))))
41, 2, 3mp2 9 . 2 Rel ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵))
5 relxp 5703 . . 3 Rel ((𝐶𝐴) × 𝐷)
6 relxp 5703 . . 3 Rel (𝐶 × (𝐷𝐵))
7 relun 5821 . . 3 (Rel (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵))) ↔ (Rel ((𝐶𝐴) × 𝐷) ∧ Rel (𝐶 × (𝐷𝐵))))
85, 6, 7mpbir2an 711 . 2 Rel (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵)))
9 ianor 984 . . . . . 6 (¬ (𝑥𝐴𝑦𝐵) ↔ (¬ 𝑥𝐴 ∨ ¬ 𝑦𝐵))
109anbi2i 623 . . . . 5 (((𝑥𝐶𝑦𝐷) ∧ ¬ (𝑥𝐴𝑦𝐵)) ↔ ((𝑥𝐶𝑦𝐷) ∧ (¬ 𝑥𝐴 ∨ ¬ 𝑦𝐵)))
11 andi 1010 . . . . 5 (((𝑥𝐶𝑦𝐷) ∧ (¬ 𝑥𝐴 ∨ ¬ 𝑦𝐵)) ↔ (((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴) ∨ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑦𝐵)))
1210, 11bitri 275 . . . 4 (((𝑥𝐶𝑦𝐷) ∧ ¬ (𝑥𝐴𝑦𝐵)) ↔ (((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴) ∨ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑦𝐵)))
13 opelxp 5721 . . . . 5 (⟨𝑥, 𝑦⟩ ∈ (𝐶 × 𝐷) ↔ (𝑥𝐶𝑦𝐷))
14 opelxp 5721 . . . . . 6 (⟨𝑥, 𝑦⟩ ∈ (𝐴 × 𝐵) ↔ (𝑥𝐴𝑦𝐵))
1514notbii 320 . . . . 5 (¬ ⟨𝑥, 𝑦⟩ ∈ (𝐴 × 𝐵) ↔ ¬ (𝑥𝐴𝑦𝐵))
1613, 15anbi12i 628 . . . 4 ((⟨𝑥, 𝑦⟩ ∈ (𝐶 × 𝐷) ∧ ¬ ⟨𝑥, 𝑦⟩ ∈ (𝐴 × 𝐵)) ↔ ((𝑥𝐶𝑦𝐷) ∧ ¬ (𝑥𝐴𝑦𝐵)))
17 opelxp 5721 . . . . . 6 (⟨𝑥, 𝑦⟩ ∈ ((𝐶𝐴) × 𝐷) ↔ (𝑥 ∈ (𝐶𝐴) ∧ 𝑦𝐷))
18 eldif 3961 . . . . . . 7 (𝑥 ∈ (𝐶𝐴) ↔ (𝑥𝐶 ∧ ¬ 𝑥𝐴))
1918anbi1i 624 . . . . . 6 ((𝑥 ∈ (𝐶𝐴) ∧ 𝑦𝐷) ↔ ((𝑥𝐶 ∧ ¬ 𝑥𝐴) ∧ 𝑦𝐷))
20 an32 646 . . . . . 6 (((𝑥𝐶 ∧ ¬ 𝑥𝐴) ∧ 𝑦𝐷) ↔ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴))
2117, 19, 203bitri 297 . . . . 5 (⟨𝑥, 𝑦⟩ ∈ ((𝐶𝐴) × 𝐷) ↔ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴))
22 eldif 3961 . . . . . . 7 (𝑦 ∈ (𝐷𝐵) ↔ (𝑦𝐷 ∧ ¬ 𝑦𝐵))
2322anbi2i 623 . . . . . 6 ((𝑥𝐶𝑦 ∈ (𝐷𝐵)) ↔ (𝑥𝐶 ∧ (𝑦𝐷 ∧ ¬ 𝑦𝐵)))
24 opelxp 5721 . . . . . 6 (⟨𝑥, 𝑦⟩ ∈ (𝐶 × (𝐷𝐵)) ↔ (𝑥𝐶𝑦 ∈ (𝐷𝐵)))
25 anass 468 . . . . . 6 (((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑦𝐵) ↔ (𝑥𝐶 ∧ (𝑦𝐷 ∧ ¬ 𝑦𝐵)))
2623, 24, 253bitr4i 303 . . . . 5 (⟨𝑥, 𝑦⟩ ∈ (𝐶 × (𝐷𝐵)) ↔ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑦𝐵))
2721, 26orbi12i 915 . . . 4 ((⟨𝑥, 𝑦⟩ ∈ ((𝐶𝐴) × 𝐷) ∨ ⟨𝑥, 𝑦⟩ ∈ (𝐶 × (𝐷𝐵))) ↔ (((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑥𝐴) ∨ ((𝑥𝐶𝑦𝐷) ∧ ¬ 𝑦𝐵)))
2812, 16, 273bitr4i 303 . . 3 ((⟨𝑥, 𝑦⟩ ∈ (𝐶 × 𝐷) ∧ ¬ ⟨𝑥, 𝑦⟩ ∈ (𝐴 × 𝐵)) ↔ (⟨𝑥, 𝑦⟩ ∈ ((𝐶𝐴) × 𝐷) ∨ ⟨𝑥, 𝑦⟩ ∈ (𝐶 × (𝐷𝐵))))
29 eldif 3961 . . 3 (⟨𝑥, 𝑦⟩ ∈ ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) ↔ (⟨𝑥, 𝑦⟩ ∈ (𝐶 × 𝐷) ∧ ¬ ⟨𝑥, 𝑦⟩ ∈ (𝐴 × 𝐵)))
30 elun 4153 . . 3 (⟨𝑥, 𝑦⟩ ∈ (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵))) ↔ (⟨𝑥, 𝑦⟩ ∈ ((𝐶𝐴) × 𝐷) ∨ ⟨𝑥, 𝑦⟩ ∈ (𝐶 × (𝐷𝐵))))
3128, 29, 303bitr4i 303 . 2 (⟨𝑥, 𝑦⟩ ∈ ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) ↔ ⟨𝑥, 𝑦⟩ ∈ (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵))))
324, 8, 31eqrelriiv 5800 1 ((𝐶 × 𝐷) ∖ (𝐴 × 𝐵)) = (((𝐶𝐴) × 𝐷) ∪ (𝐶 × (𝐷𝐵)))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wa 395  wo 848   = wceq 1540  wcel 2108  cdif 3948  cun 3949  wss 3951  cop 4632   × cxp 5683  Rel wrel 5690
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1795  ax-4 1809  ax-5 1910  ax-6 1967  ax-7 2007  ax-8 2110  ax-9 2118  ax-ext 2708  ax-sep 5296  ax-nul 5306  ax-pr 5432
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 849  df-3an 1089  df-tru 1543  df-fal 1553  df-ex 1780  df-sb 2065  df-clab 2715  df-cleq 2729  df-clel 2816  df-ral 3062  df-rex 3071  df-rab 3437  df-v 3482  df-dif 3954  df-un 3956  df-ss 3968  df-nul 4334  df-if 4526  df-sn 4627  df-pr 4629  df-op 4633  df-opab 5206  df-xp 5691  df-rel 5692
This theorem is referenced by:  difxp1  6185  difxp2  6186  evlslem4  22100  txcld  23611  suppovss  32690  elrgspnlem2  33247  elrgspnsubrunlem2  33252
  Copyright terms: Public domain W3C validator