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

Theorem ssrel 5207
Description: A subclass relationship depends only on a relation's ordered pairs. Theorem 3.2(i) of [Monk1] p. 33. (Contributed by NM, 2-Aug-1994.) (Proof shortened by Andrew Salmon, 27-Aug-2011.) Remove dependency on ax-sep 4781, ax-nul 4789, ax-pr 4906. (Revised by KP, 25-Oct-2021.)
Assertion
Ref Expression
ssrel (Rel 𝐴 → (𝐴𝐵 ↔ ∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵)))
Distinct variable groups:   𝑥,𝑦,𝐴   𝑥,𝐵,𝑦

Proof of Theorem ssrel
Dummy variable 𝑧 is distinct from all other variables.
StepHypRef Expression
1 ssel 3597 . . 3 (𝐴𝐵 → (⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵))
21alrimivv 1856 . 2 (𝐴𝐵 → ∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵))
3 df-rel 5121 . . . . . . 7 (Rel 𝐴𝐴 ⊆ (V × V))
4 dfss2 3591 . . . . . . 7 (𝐴 ⊆ (V × V) ↔ ∀𝑧(𝑧𝐴𝑧 ∈ (V × V)))
53, 4sylbb 209 . . . . . 6 (Rel 𝐴 → ∀𝑧(𝑧𝐴𝑧 ∈ (V × V)))
6 df-xp 5120 . . . . . . . . . 10 (V × V) = {⟨𝑥, 𝑦⟩ ∣ (𝑥 ∈ V ∧ 𝑦 ∈ V)}
7 df-opab 4713 . . . . . . . . . 10 {⟨𝑥, 𝑦⟩ ∣ (𝑥 ∈ V ∧ 𝑦 ∈ V)} = {𝑧 ∣ ∃𝑥𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ (𝑥 ∈ V ∧ 𝑦 ∈ V))}
86, 7eqtri 2644 . . . . . . . . 9 (V × V) = {𝑧 ∣ ∃𝑥𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ (𝑥 ∈ V ∧ 𝑦 ∈ V))}
98abeq2i 2735 . . . . . . . 8 (𝑧 ∈ (V × V) ↔ ∃𝑥𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ (𝑥 ∈ V ∧ 𝑦 ∈ V)))
10 simpl 473 . . . . . . . . 9 ((𝑧 = ⟨𝑥, 𝑦⟩ ∧ (𝑥 ∈ V ∧ 𝑦 ∈ V)) → 𝑧 = ⟨𝑥, 𝑦⟩)
11102eximi 1763 . . . . . . . 8 (∃𝑥𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ (𝑥 ∈ V ∧ 𝑦 ∈ V)) → ∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩)
129, 11sylbi 207 . . . . . . 7 (𝑧 ∈ (V × V) → ∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩)
1312imim2i 16 . . . . . 6 ((𝑧𝐴𝑧 ∈ (V × V)) → (𝑧𝐴 → ∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩))
145, 13sylg 1750 . . . . 5 (Rel 𝐴 → ∀𝑧(𝑧𝐴 → ∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩))
15 eleq1 2689 . . . . . . . . . . . 12 (𝑧 = ⟨𝑥, 𝑦⟩ → (𝑧𝐴 ↔ ⟨𝑥, 𝑦⟩ ∈ 𝐴))
16 eleq1 2689 . . . . . . . . . . . 12 (𝑧 = ⟨𝑥, 𝑦⟩ → (𝑧𝐵 ↔ ⟨𝑥, 𝑦⟩ ∈ 𝐵))
1715, 16imbi12d 334 . . . . . . . . . . 11 (𝑧 = ⟨𝑥, 𝑦⟩ → ((𝑧𝐴𝑧𝐵) ↔ (⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵)))
1817biimprcd 240 . . . . . . . . . 10 ((⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵) → (𝑧 = ⟨𝑥, 𝑦⟩ → (𝑧𝐴𝑧𝐵)))
19182alimi 1740 . . . . . . . . 9 (∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵) → ∀𝑥𝑦(𝑧 = ⟨𝑥, 𝑦⟩ → (𝑧𝐴𝑧𝐵)))
20 19.23vv 1903 . . . . . . . . 9 (∀𝑥𝑦(𝑧 = ⟨𝑥, 𝑦⟩ → (𝑧𝐴𝑧𝐵)) ↔ (∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩ → (𝑧𝐴𝑧𝐵)))
2119, 20sylib 208 . . . . . . . 8 (∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵) → (∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩ → (𝑧𝐴𝑧𝐵)))
2221com23 86 . . . . . . 7 (∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵) → (𝑧𝐴 → (∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩ → 𝑧𝐵)))
2322a2d 29 . . . . . 6 (∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵) → ((𝑧𝐴 → ∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩) → (𝑧𝐴𝑧𝐵)))
2423alimdv 1845 . . . . 5 (∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵) → (∀𝑧(𝑧𝐴 → ∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩) → ∀𝑧(𝑧𝐴𝑧𝐵)))
2514, 24syl5 34 . . . 4 (∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵) → (Rel 𝐴 → ∀𝑧(𝑧𝐴𝑧𝐵)))
26 dfss2 3591 . . . 4 (𝐴𝐵 ↔ ∀𝑧(𝑧𝐴𝑧𝐵))
2725, 26syl6ibr 242 . . 3 (∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵) → (Rel 𝐴𝐴𝐵))
2827com12 32 . 2 (Rel 𝐴 → (∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵) → 𝐴𝐵))
292, 28impbid2 216 1 (Rel 𝐴 → (𝐴𝐵 ↔ ∀𝑥𝑦(⟨𝑥, 𝑦⟩ ∈ 𝐴 → ⟨𝑥, 𝑦⟩ ∈ 𝐵)))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 196  wa 384  wal 1481   = wceq 1483  wex 1704  wcel 1990  {cab 2608  Vcvv 3200  wss 3574  cop 4183  {copab 4712   × cxp 5112  Rel wrel 5119
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1722  ax-4 1737  ax-5 1839  ax-6 1888  ax-7 1935  ax-9 1999  ax-10 2019  ax-11 2034  ax-12 2047  ax-13 2246  ax-ext 2602
This theorem depends on definitions:  df-bi 197  df-or 385  df-an 386  df-tru 1486  df-ex 1705  df-nf 1710  df-sb 1881  df-clab 2609  df-cleq 2615  df-clel 2618  df-in 3581  df-ss 3588  df-opab 4713  df-xp 5120  df-rel 5121
This theorem is referenced by:  eqrel  5209  relssi  5211  relssdv  5212  cotrg  5507  cnvsym  5510  intasym  5511  intirr  5514  codir  5516  qfto  5517  dfso2  31644  dfpo2  31645  dffun10  32021  imagesset  32060  ssrel3  34067  undmrnresiss  37910  cnvssco  37912
  Copyright terms: Public domain W3C validator