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

Theorem wfrlem13OLD 8272
Description: Lemma for well-ordered recursion. From here through wfrlem16OLD 8275, 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 8270 . . . 4 Fun 𝐹
5 vex 3450 . . . . 5 𝑧 ∈ V
6 fvex 6860 . . . . 5 (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧))) ∈ V
75, 6funsn 6559 . . . 4 Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}
84, 7pm3.2i 471 . . 3 (Fun 𝐹 ∧ Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
96dmsnop 6173 . . . . 5 dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩} = {𝑧}
109ineq2i 4174 . . . 4 (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∩ {𝑧})
11 eldifn 4092 . . . . 5 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → ¬ 𝑧 ∈ dom 𝐹)
12 disjsn 4677 . . . . 5 ((dom 𝐹 ∩ {𝑧}) = ∅ ↔ ¬ 𝑧 ∈ dom 𝐹)
1311, 12sylibr 233 . . . 4 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → (dom 𝐹 ∩ {𝑧}) = ∅)
1410, 13eqtrid 2783 . . 3 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = ∅)
15 funun 6552 . . 3 (((Fun 𝐹 ∧ Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) ∧ (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = ∅) → Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}))
168, 14, 15sylancr 587 . 2 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}))
17 dmun 5871 . . 3 dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
189uneq2i 4125 . . 3 (dom 𝐹 ∪ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})
1917, 18eqtri 2759 . 2 dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})
20 wfrlem13OLD.4 . . . 4 𝐶 = (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
2120fneq1i 6604 . . 3 (𝐶 Fn (dom 𝐹 ∪ {𝑧}) ↔ (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) Fn (dom 𝐹 ∪ {𝑧}))
22 df-fn 6504 . . 3 ((𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) Fn (dom 𝐹 ∪ {𝑧}) ↔ (Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) ∧ dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})))
2321, 22bitri 274 . 2 (𝐶 Fn (dom 𝐹 ∪ {𝑧}) ↔ (Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) ∧ dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})))
2416, 19, 23sylanblrc 590 1 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → 𝐶 Fn (dom 𝐹 ∪ {𝑧}))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 396   = wceq 1541  wcel 2106  cdif 3910  cun 3911  cin 3912  c0 4287  {csn 4591  cop 4597   Se wse 5591   We wwe 5592  dom cdm 5638  cres 5640  Predcpred 6257  Fun wfun 6495   Fn wfn 6496  cfv 6501  wrecscwrecs 8247
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1797  ax-4 1811  ax-5 1913  ax-6 1971  ax-7 2011  ax-8 2108  ax-9 2116  ax-10 2137  ax-11 2154  ax-12 2171  ax-ext 2702  ax-sep 5261  ax-nul 5268  ax-pr 5389  ax-un 7677
This theorem depends on definitions:  df-bi 206  df-an 397  df-or 846  df-3an 1089  df-tru 1544  df-fal 1554  df-ex 1782  df-nf 1786  df-sb 2068  df-mo 2533  df-eu 2562  df-clab 2709  df-cleq 2723  df-clel 2809  df-nfc 2884  df-ne 2940  df-ral 3061  df-rex 3070  df-rab 3406  df-v 3448  df-sbc 3743  df-csb 3859  df-dif 3916  df-un 3918  df-in 3920  df-ss 3930  df-nul 4288  df-if 4492  df-pw 4567  df-sn 4592  df-pr 4594  df-op 4598  df-uni 4871  df-iun 4961  df-br 5111  df-opab 5173  df-mpt 5194  df-id 5536  df-po 5550  df-so 5551  df-fr 5593  df-se 5594  df-we 5595  df-xp 5644  df-rel 5645  df-cnv 5646  df-co 5647  df-dm 5648  df-rn 5649  df-res 5650  df-ima 5651  df-pred 6258  df-iota 6453  df-fun 6503  df-fn 6504  df-f 6505  df-fo 6507  df-fv 6509  df-ov 7365  df-2nd 7927  df-frecs 8217  df-wrecs 8248
This theorem is referenced by:  wfrlem14OLD  8273  wfrlem15OLD  8274
  Copyright terms: Public domain W3C validator