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

Theorem dfopab2 7222
Description: A way to define an ordered-pair class abstraction without using existential quantifiers. (Contributed by NM, 18-Aug-2006.) (Revised by Mario Carneiro, 31-Aug-2015.)
Assertion
Ref Expression
dfopab2 {⟨𝑥, 𝑦⟩ ∣ 𝜑} = {𝑧 ∈ (V × V) ∣ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑}
Distinct variable groups:   𝜑,𝑧   𝑥,𝑦,𝑧
Allowed substitution hints:   𝜑(𝑥,𝑦)

Proof of Theorem dfopab2
StepHypRef Expression
1 nfsbc1v 3455 . . . . 5 𝑥[(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑
2119.41 2103 . . . 4 (∃𝑥(∃𝑦 𝑧 = ⟨𝑥, 𝑦⟩ ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑) ↔ (∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩ ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑))
3 sbcopeq1a 7220 . . . . . . . 8 (𝑧 = ⟨𝑥, 𝑦⟩ → ([(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑𝜑))
43pm5.32i 669 . . . . . . 7 ((𝑧 = ⟨𝑥, 𝑦⟩ ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑) ↔ (𝑧 = ⟨𝑥, 𝑦⟩ ∧ 𝜑))
54exbii 1774 . . . . . 6 (∃𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑) ↔ ∃𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ 𝜑))
6 nfcv 2764 . . . . . . . 8 𝑦(1st𝑧)
7 nfsbc1v 3455 . . . . . . . 8 𝑦[(2nd𝑧) / 𝑦]𝜑
86, 7nfsbc 3457 . . . . . . 7 𝑦[(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑
9819.41 2103 . . . . . 6 (∃𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑) ↔ (∃𝑦 𝑧 = ⟨𝑥, 𝑦⟩ ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑))
105, 9bitr3i 266 . . . . 5 (∃𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ 𝜑) ↔ (∃𝑦 𝑧 = ⟨𝑥, 𝑦⟩ ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑))
1110exbii 1774 . . . 4 (∃𝑥𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ 𝜑) ↔ ∃𝑥(∃𝑦 𝑧 = ⟨𝑥, 𝑦⟩ ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑))
12 elvv 5177 . . . . 5 (𝑧 ∈ (V × V) ↔ ∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩)
1312anbi1i 731 . . . 4 ((𝑧 ∈ (V × V) ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑) ↔ (∃𝑥𝑦 𝑧 = ⟨𝑥, 𝑦⟩ ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑))
142, 11, 133bitr4i 292 . . 3 (∃𝑥𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ 𝜑) ↔ (𝑧 ∈ (V × V) ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑))
1514abbii 2739 . 2 {𝑧 ∣ ∃𝑥𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ 𝜑)} = {𝑧 ∣ (𝑧 ∈ (V × V) ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑)}
16 df-opab 4713 . 2 {⟨𝑥, 𝑦⟩ ∣ 𝜑} = {𝑧 ∣ ∃𝑥𝑦(𝑧 = ⟨𝑥, 𝑦⟩ ∧ 𝜑)}
17 df-rab 2921 . 2 {𝑧 ∈ (V × V) ∣ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑} = {𝑧 ∣ (𝑧 ∈ (V × V) ∧ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑)}
1815, 16, 173eqtr4i 2654 1 {⟨𝑥, 𝑦⟩ ∣ 𝜑} = {𝑧 ∈ (V × V) ∣ [(1st𝑧) / 𝑥][(2nd𝑧) / 𝑦]𝜑}
Colors of variables: wff setvar class
Syntax hints:  wa 384   = wceq 1483  wex 1704  wcel 1990  {cab 2608  {crab 2916  Vcvv 3200  [wsbc 3435  cop 4183  {copab 4712   × cxp 5112  cfv 5888  1st c1st 7166  2nd c2nd 7167
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-8 1992  ax-9 1999  ax-10 2019  ax-11 2034  ax-12 2047  ax-13 2246  ax-ext 2602  ax-sep 4781  ax-nul 4789  ax-pow 4843  ax-pr 4906  ax-un 6949
This theorem depends on definitions:  df-bi 197  df-or 385  df-an 386  df-3an 1039  df-tru 1486  df-ex 1705  df-nf 1710  df-sb 1881  df-eu 2474  df-mo 2475  df-clab 2609  df-cleq 2615  df-clel 2618  df-nfc 2753  df-ral 2917  df-rex 2918  df-rab 2921  df-v 3202  df-sbc 3436  df-dif 3577  df-un 3579  df-in 3581  df-ss 3588  df-nul 3916  df-if 4087  df-sn 4178  df-pr 4180  df-op 4184  df-uni 4437  df-br 4654  df-opab 4713  df-mpt 4730  df-id 5024  df-xp 5120  df-rel 5121  df-cnv 5122  df-co 5123  df-dm 5124  df-rn 5125  df-iota 5851  df-fun 5890  df-fv 5896  df-1st 7168  df-2nd 7169
This theorem is referenced by: (None)
  Copyright terms: Public domain W3C validator