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

Theorem wfrlem13OLD 8123
Description: Lemma for well-ordered recursion. From here through wfrlem16OLD 8126, 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 8121 . . . 4 Fun 𝐹
5 vex 3426 . . . . 5 𝑧 ∈ V
6 fvex 6769 . . . . 5 (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧))) ∈ V
75, 6funsn 6471 . . . 4 Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}
84, 7pm3.2i 470 . . 3 (Fun 𝐹 ∧ Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
96dmsnop 6108 . . . . 5 dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩} = {𝑧}
109ineq2i 4140 . . . 4 (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∩ {𝑧})
11 eldifn 4058 . . . . 5 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → ¬ 𝑧 ∈ dom 𝐹)
12 disjsn 4644 . . . . 5 ((dom 𝐹 ∩ {𝑧}) = ∅ ↔ ¬ 𝑧 ∈ dom 𝐹)
1311, 12sylibr 233 . . . 4 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → (dom 𝐹 ∩ {𝑧}) = ∅)
1410, 13eqtrid 2790 . . 3 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = ∅)
15 funun 6464 . . 3 (((Fun 𝐹 ∧ Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) ∧ (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = ∅) → Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}))
168, 14, 15sylancr 586 . 2 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}))
17 dmun 5808 . . 3 dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
189uneq2i 4090 . . 3 (dom 𝐹 ∪ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})
1917, 18eqtri 2766 . 2 dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})
20 wfrlem13OLD.4 . . . 4 𝐶 = (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
2120fneq1i 6514 . . 3 (𝐶 Fn (dom 𝐹 ∪ {𝑧}) ↔ (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) Fn (dom 𝐹 ∪ {𝑧}))
22 df-fn 6421 . . 3 ((𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) Fn (dom 𝐹 ∪ {𝑧}) ↔ (Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) ∧ dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})))
2321, 22bitri 274 . 2 (𝐶 Fn (dom 𝐹 ∪ {𝑧}) ↔ (Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) ∧ dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})))
2416, 19, 23sylanblrc 589 1 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → 𝐶 Fn (dom 𝐹 ∪ {𝑧}))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 395   = wceq 1539  wcel 2108  cdif 3880  cun 3881  cin 3882  c0 4253  {csn 4558  cop 4564   Se wse 5533   We wwe 5534  dom cdm 5580  cres 5582  Predcpred 6190  Fun wfun 6412   Fn wfn 6413  cfv 6418  wrecscwrecs 8098
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1799  ax-4 1813  ax-5 1914  ax-6 1972  ax-7 2012  ax-8 2110  ax-9 2118  ax-10 2139  ax-11 2156  ax-12 2173  ax-ext 2709  ax-sep 5218  ax-nul 5225  ax-pr 5347  ax-un 7566
This theorem depends on definitions:  df-bi 206  df-an 396  df-or 844  df-3an 1087  df-tru 1542  df-fal 1552  df-ex 1784  df-nf 1788  df-sb 2069  df-mo 2540  df-eu 2569  df-clab 2716  df-cleq 2730  df-clel 2817  df-nfc 2888  df-ne 2943  df-ral 3068  df-rex 3069  df-rab 3072  df-v 3424  df-sbc 3712  df-csb 3829  df-dif 3886  df-un 3888  df-in 3890  df-ss 3900  df-nul 4254  df-if 4457  df-pw 4532  df-sn 4559  df-pr 4561  df-op 4565  df-uni 4837  df-iun 4923  df-br 5071  df-opab 5133  df-mpt 5154  df-id 5480  df-po 5494  df-so 5495  df-fr 5535  df-se 5536  df-we 5537  df-xp 5586  df-rel 5587  df-cnv 5588  df-co 5589  df-dm 5590  df-rn 5591  df-res 5592  df-ima 5593  df-pred 6191  df-iota 6376  df-fun 6420  df-fn 6421  df-f 6422  df-fo 6424  df-fv 6426  df-ov 7258  df-2nd 7805  df-frecs 8068  df-wrecs 8099
This theorem is referenced by:  wfrlem14OLD  8124  wfrlem15OLD  8125
  Copyright terms: Public domain W3C validator