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

Theorem cbvopab1g 5218
Description: Change first bound variable in an ordered-pair class abstraction, using explicit substitution. Usage of this theorem is discouraged because it depends on ax-13 2377. See cbvopab1 5217 for a version with more disjoint variable conditions, but not requiring ax-13 2377. (Contributed by NM, 6-Oct-2004.) (Revised by Mario Carneiro, 14-Oct-2016.) (New usage is discouraged.)
Hypotheses
Ref Expression
cbvopab1g.1 𝑧𝜑
cbvopab1g.2 𝑥𝜓
cbvopab1g.3 (𝑥 = 𝑧 → (𝜑𝜓))
Assertion
Ref Expression
cbvopab1g {⟨𝑥, 𝑦⟩ ∣ 𝜑} = {⟨𝑧, 𝑦⟩ ∣ 𝜓}
Distinct variable groups:   𝑥,𝑦   𝑦,𝑧
Allowed substitution hints:   𝜑(𝑥,𝑦,𝑧)   𝜓(𝑥,𝑦,𝑧)

Proof of Theorem cbvopab1g
Dummy variables 𝑤 𝑣 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 nfv 1914 . . . . 5 𝑣𝑦(𝑤 = ⟨𝑥, 𝑦⟩ ∧ 𝜑)
2 nfv 1914 . . . . . . 7 𝑥 𝑤 = ⟨𝑣, 𝑦
3 nfs1v 2156 . . . . . . 7 𝑥[𝑣 / 𝑥]𝜑
42, 3nfan 1899 . . . . . 6 𝑥(𝑤 = ⟨𝑣, 𝑦⟩ ∧ [𝑣 / 𝑥]𝜑)
54nfex 2324 . . . . 5 𝑥𝑦(𝑤 = ⟨𝑣, 𝑦⟩ ∧ [𝑣 / 𝑥]𝜑)
6 opeq1 4873 . . . . . . . 8 (𝑥 = 𝑣 → ⟨𝑥, 𝑦⟩ = ⟨𝑣, 𝑦⟩)
76eqeq2d 2748 . . . . . . 7 (𝑥 = 𝑣 → (𝑤 = ⟨𝑥, 𝑦⟩ ↔ 𝑤 = ⟨𝑣, 𝑦⟩))
8 sbequ12 2251 . . . . . . 7 (𝑥 = 𝑣 → (𝜑 ↔ [𝑣 / 𝑥]𝜑))
97, 8anbi12d 632 . . . . . 6 (𝑥 = 𝑣 → ((𝑤 = ⟨𝑥, 𝑦⟩ ∧ 𝜑) ↔ (𝑤 = ⟨𝑣, 𝑦⟩ ∧ [𝑣 / 𝑥]𝜑)))
109exbidv 1921 . . . . 5 (𝑥 = 𝑣 → (∃𝑦(𝑤 = ⟨𝑥, 𝑦⟩ ∧ 𝜑) ↔ ∃𝑦(𝑤 = ⟨𝑣, 𝑦⟩ ∧ [𝑣 / 𝑥]𝜑)))
111, 5, 10cbvexv1 2344 . . . 4 (∃𝑥𝑦(𝑤 = ⟨𝑥, 𝑦⟩ ∧ 𝜑) ↔ ∃𝑣𝑦(𝑤 = ⟨𝑣, 𝑦⟩ ∧ [𝑣 / 𝑥]𝜑))
12 nfv 1914 . . . . . . 7 𝑧 𝑤 = ⟨𝑣, 𝑦
13 cbvopab1g.1 . . . . . . . 8 𝑧𝜑
1413nfsb 2528 . . . . . . 7 𝑧[𝑣 / 𝑥]𝜑
1512, 14nfan 1899 . . . . . 6 𝑧(𝑤 = ⟨𝑣, 𝑦⟩ ∧ [𝑣 / 𝑥]𝜑)
1615nfex 2324 . . . . 5 𝑧𝑦(𝑤 = ⟨𝑣, 𝑦⟩ ∧ [𝑣 / 𝑥]𝜑)
17 nfv 1914 . . . . 5 𝑣𝑦(𝑤 = ⟨𝑧, 𝑦⟩ ∧ 𝜓)
18 opeq1 4873 . . . . . . . 8 (𝑣 = 𝑧 → ⟨𝑣, 𝑦⟩ = ⟨𝑧, 𝑦⟩)
1918eqeq2d 2748 . . . . . . 7 (𝑣 = 𝑧 → (𝑤 = ⟨𝑣, 𝑦⟩ ↔ 𝑤 = ⟨𝑧, 𝑦⟩))
20 sbequ 2083 . . . . . . . 8 (𝑣 = 𝑧 → ([𝑣 / 𝑥]𝜑 ↔ [𝑧 / 𝑥]𝜑))
21 cbvopab1g.2 . . . . . . . . 9 𝑥𝜓
22 cbvopab1g.3 . . . . . . . . 9 (𝑥 = 𝑧 → (𝜑𝜓))
2321, 22sbie 2507 . . . . . . . 8 ([𝑧 / 𝑥]𝜑𝜓)
2420, 23bitrdi 287 . . . . . . 7 (𝑣 = 𝑧 → ([𝑣 / 𝑥]𝜑𝜓))
2519, 24anbi12d 632 . . . . . 6 (𝑣 = 𝑧 → ((𝑤 = ⟨𝑣, 𝑦⟩ ∧ [𝑣 / 𝑥]𝜑) ↔ (𝑤 = ⟨𝑧, 𝑦⟩ ∧ 𝜓)))
2625exbidv 1921 . . . . 5 (𝑣 = 𝑧 → (∃𝑦(𝑤 = ⟨𝑣, 𝑦⟩ ∧ [𝑣 / 𝑥]𝜑) ↔ ∃𝑦(𝑤 = ⟨𝑧, 𝑦⟩ ∧ 𝜓)))
2716, 17, 26cbvexv1 2344 . . . 4 (∃𝑣𝑦(𝑤 = ⟨𝑣, 𝑦⟩ ∧ [𝑣 / 𝑥]𝜑) ↔ ∃𝑧𝑦(𝑤 = ⟨𝑧, 𝑦⟩ ∧ 𝜓))
2811, 27bitri 275 . . 3 (∃𝑥𝑦(𝑤 = ⟨𝑥, 𝑦⟩ ∧ 𝜑) ↔ ∃𝑧𝑦(𝑤 = ⟨𝑧, 𝑦⟩ ∧ 𝜓))
2928abbii 2809 . 2 {𝑤 ∣ ∃𝑥𝑦(𝑤 = ⟨𝑥, 𝑦⟩ ∧ 𝜑)} = {𝑤 ∣ ∃𝑧𝑦(𝑤 = ⟨𝑧, 𝑦⟩ ∧ 𝜓)}
30 df-opab 5206 . 2 {⟨𝑥, 𝑦⟩ ∣ 𝜑} = {𝑤 ∣ ∃𝑥𝑦(𝑤 = ⟨𝑥, 𝑦⟩ ∧ 𝜑)}
31 df-opab 5206 . 2 {⟨𝑧, 𝑦⟩ ∣ 𝜓} = {𝑤 ∣ ∃𝑧𝑦(𝑤 = ⟨𝑧, 𝑦⟩ ∧ 𝜓)}
3229, 30, 313eqtr4i 2775 1 {⟨𝑥, 𝑦⟩ ∣ 𝜑} = {⟨𝑧, 𝑦⟩ ∣ 𝜓}
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 206  wa 395   = wceq 1540  wex 1779  wnf 1783  [wsb 2064  {cab 2714  cop 4632  {copab 5205
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-10 2141  ax-11 2157  ax-12 2177  ax-13 2377  ax-ext 2708
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-nf 1784  df-sb 2065  df-clab 2715  df-cleq 2729  df-clel 2816  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
This theorem is referenced by:  cbvmptfg  5252
  Copyright terms: Public domain W3C validator