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

Theorem wfrlem13OLD 8323
Description: Lemma for well-ordered recursion. From here through wfrlem16OLD 8326, we aim to prove that dom 𝐹 = 𝐴. We do this by supposing that there is an element 𝑧 of 𝐴 that is not in dom 𝐹. We then define 𝐶 by extending dom 𝐹 with the appropriate value at 𝑧. We then show that 𝑧 cannot be an 𝑅 minimal element of (𝐴 ∖ dom 𝐹), meaning that (𝐴 ∖ dom 𝐹) must be empty, so dom 𝐹 = 𝐴. Here, we show that 𝐶 is a function extending the domain of 𝐹 by one. Obsolete as of 18-Nov-2024. (New usage is discouraged.) (Proof modification is discouraged.) (Contributed by Scott Fenton, 21-Apr-2011.) (Revised by Mario Carneiro, 26-Jun-2015.)
Hypotheses
Ref Expression
wfrlem13OLD.1 𝑅 We 𝐴
wfrlem13OLD.2 𝑅 Se 𝐴
wfrlem13OLD.3 𝐹 = wrecs(𝑅, 𝐴, 𝐺)
wfrlem13OLD.4 𝐶 = (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
Assertion
Ref Expression
wfrlem13OLD (𝑧 ∈ (𝐴 ∖ dom 𝐹) → 𝐶 Fn (dom 𝐹 ∪ {𝑧}))
Distinct variable groups:   𝑧,𝐴   𝑧,𝐹   𝑧,𝑅
Allowed substitution hints:   𝐶(𝑧)   𝐺(𝑧)

Proof of Theorem wfrlem13OLD
StepHypRef Expression
1 wfrlem13OLD.1 . . . . 5 𝑅 We 𝐴
2 wfrlem13OLD.2 . . . . 5 𝑅 Se 𝐴
3 wfrlem13OLD.3 . . . . 5 𝐹 = wrecs(𝑅, 𝐴, 𝐺)
41, 2, 3wfrfunOLD 8321 . . . 4 Fun 𝐹
5 vex 3476 . . . . 5 𝑧 ∈ V
6 fvex 6903 . . . . 5 (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧))) ∈ V
75, 6funsn 6600 . . . 4 Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}
84, 7pm3.2i 469 . . 3 (Fun 𝐹 ∧ Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
96dmsnop 6214 . . . . 5 dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩} = {𝑧}
109ineq2i 4208 . . . 4 (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∩ {𝑧})
11 eldifn 4126 . . . . 5 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → ¬ 𝑧 ∈ dom 𝐹)
12 disjsn 4714 . . . . 5 ((dom 𝐹 ∩ {𝑧}) = ∅ ↔ ¬ 𝑧 ∈ dom 𝐹)
1311, 12sylibr 233 . . . 4 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → (dom 𝐹 ∩ {𝑧}) = ∅)
1410, 13eqtrid 2782 . . 3 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = ∅)
15 funun 6593 . . 3 (((Fun 𝐹 ∧ Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) ∧ (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = ∅) → Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}))
168, 14, 15sylancr 585 . 2 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}))
17 dmun 5909 . . 3 dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
189uneq2i 4159 . . 3 (dom 𝐹 ∪ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})
1917, 18eqtri 2758 . 2 dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})
20 wfrlem13OLD.4 . . . 4 𝐶 = (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
2120fneq1i 6645 . . 3 (𝐶 Fn (dom 𝐹 ∪ {𝑧}) ↔ (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) Fn (dom 𝐹 ∪ {𝑧}))
22 df-fn 6545 . . 3 ((𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) Fn (dom 𝐹 ∪ {𝑧}) ↔ (Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) ∧ dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})))
2321, 22bitri 274 . 2 (𝐶 Fn (dom 𝐹 ∪ {𝑧}) ↔ (Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) ∧ dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})))
2416, 19, 23sylanblrc 588 1 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → 𝐶 Fn (dom 𝐹 ∪ {𝑧}))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 394   = wceq 1539  wcel 2104  cdif 3944  cun 3945  cin 3946  c0 4321  {csn 4627  cop 4633   Se wse 5628   We wwe 5629  dom cdm 5675  cres 5677  Predcpred 6298  Fun wfun 6536   Fn wfn 6537  cfv 6542  wrecscwrecs 8298
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 1911  ax-6 1969  ax-7 2009  ax-8 2106  ax-9 2114  ax-10 2135  ax-11 2152  ax-12 2169  ax-ext 2701  ax-sep 5298  ax-nul 5305  ax-pr 5426  ax-un 7727
This theorem depends on definitions:  df-bi 206  df-an 395  df-or 844  df-3an 1087  df-tru 1542  df-fal 1552  df-ex 1780  df-nf 1784  df-sb 2066  df-mo 2532  df-eu 2561  df-clab 2708  df-cleq 2722  df-clel 2808  df-nfc 2883  df-ne 2939  df-ral 3060  df-rex 3069  df-rab 3431  df-v 3474  df-sbc 3777  df-csb 3893  df-dif 3950  df-un 3952  df-in 3954  df-ss 3964  df-nul 4322  df-if 4528  df-pw 4603  df-sn 4628  df-pr 4630  df-op 4634  df-uni 4908  df-iun 4998  df-br 5148  df-opab 5210  df-mpt 5231  df-id 5573  df-po 5587  df-so 5588  df-fr 5630  df-se 5631  df-we 5632  df-xp 5681  df-rel 5682  df-cnv 5683  df-co 5684  df-dm 5685  df-rn 5686  df-res 5687  df-ima 5688  df-pred 6299  df-iota 6494  df-fun 6544  df-fn 6545  df-f 6546  df-fo 6548  df-fv 6550  df-ov 7414  df-2nd 7978  df-frecs 8268  df-wrecs 8299
This theorem is referenced by:  wfrlem14OLD  8324  wfrlem15OLD  8325
  Copyright terms: Public domain W3C validator