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

Theorem wfrlem13OLD 8152
Description: Lemma for well-ordered recursion. From here through wfrlem16OLD 8155, 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 8150 . . . 4 Fun 𝐹
5 vex 3436 . . . . 5 𝑧 ∈ V
6 fvex 6787 . . . . 5 (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧))) ∈ V
75, 6funsn 6487 . . . 4 Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}
84, 7pm3.2i 471 . . 3 (Fun 𝐹 ∧ Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
96dmsnop 6119 . . . . 5 dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩} = {𝑧}
109ineq2i 4143 . . . 4 (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∩ {𝑧})
11 eldifn 4062 . . . . 5 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → ¬ 𝑧 ∈ dom 𝐹)
12 disjsn 4647 . . . . 5 ((dom 𝐹 ∩ {𝑧}) = ∅ ↔ ¬ 𝑧 ∈ dom 𝐹)
1311, 12sylibr 233 . . . 4 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → (dom 𝐹 ∩ {𝑧}) = ∅)
1410, 13eqtrid 2790 . . 3 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = ∅)
15 funun 6480 . . 3 (((Fun 𝐹 ∧ Fun {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) ∧ (dom 𝐹 ∩ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = ∅) → Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}))
168, 14, 15sylancr 587 . 2 (𝑧 ∈ (𝐴 ∖ dom 𝐹) → Fun (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}))
17 dmun 5819 . . 3 dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
189uneq2i 4094 . . 3 (dom 𝐹 ∪ dom {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})
1917, 18eqtri 2766 . 2 dom (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) = (dom 𝐹 ∪ {𝑧})
20 wfrlem13OLD.4 . . . 4 𝐶 = (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩})
2120fneq1i 6530 . . 3 (𝐶 Fn (dom 𝐹 ∪ {𝑧}) ↔ (𝐹 ∪ {⟨𝑧, (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑧)))⟩}) Fn (dom 𝐹 ∪ {𝑧}))
22 df-fn 6436 . . 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 1539  wcel 2106  cdif 3884  cun 3885  cin 3886  c0 4256  {csn 4561  cop 4567   Se wse 5542   We wwe 5543  dom cdm 5589  cres 5591  Predcpred 6201  Fun wfun 6427   Fn wfn 6428  cfv 6433  wrecscwrecs 8127
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1798  ax-4 1812  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 2709  ax-sep 5223  ax-nul 5230  ax-pr 5352  ax-un 7588
This theorem depends on definitions:  df-bi 206  df-an 397  df-or 845  df-3an 1088  df-tru 1542  df-fal 1552  df-ex 1783  df-nf 1787  df-sb 2068  df-mo 2540  df-eu 2569  df-clab 2716  df-cleq 2730  df-clel 2816  df-nfc 2889  df-ne 2944  df-ral 3069  df-rex 3070  df-rab 3073  df-v 3434  df-sbc 3717  df-csb 3833  df-dif 3890  df-un 3892  df-in 3894  df-ss 3904  df-nul 4257  df-if 4460  df-pw 4535  df-sn 4562  df-pr 4564  df-op 4568  df-uni 4840  df-iun 4926  df-br 5075  df-opab 5137  df-mpt 5158  df-id 5489  df-po 5503  df-so 5504  df-fr 5544  df-se 5545  df-we 5546  df-xp 5595  df-rel 5596  df-cnv 5597  df-co 5598  df-dm 5599  df-rn 5600  df-res 5601  df-ima 5602  df-pred 6202  df-iota 6391  df-fun 6435  df-fn 6436  df-f 6437  df-fo 6439  df-fv 6441  df-ov 7278  df-2nd 7832  df-frecs 8097  df-wrecs 8128
This theorem is referenced by:  wfrlem14OLD  8153  wfrlem15OLD  8154
  Copyright terms: Public domain W3C validator