| Intuitionistic Logic Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > ILE Home > Th. List > tfrlem3a | GIF version | ||
| Description: Lemma for transfinite recursion. Let 𝐴 be the class of "acceptable" functions. The final thing we're interested in is the union of all these acceptable functions. This lemma just changes some bound variables in 𝐴 for later use. (Contributed by NM, 9-Apr-1995.) |
| Ref | Expression |
|---|---|
| tfrlem3.1 | ⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} |
| tfrlem3.2 | ⊢ 𝐺 ∈ V |
| Ref | Expression |
|---|---|
| tfrlem3a | ⊢ (𝐺 ∈ 𝐴 ↔ ∃𝑧 ∈ On (𝐺 Fn 𝑧 ∧ ∀𝑤 ∈ 𝑧 (𝐺‘𝑤) = (𝐹‘(𝐺 ↾ 𝑤)))) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | tfrlem3.2 | . 2 ⊢ 𝐺 ∈ V | |
| 2 | fneq12 5351 | . . . 4 ⊢ ((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) → (𝑓 Fn 𝑥 ↔ 𝐺 Fn 𝑧)) | |
| 3 | simpll 527 | . . . . . . 7 ⊢ (((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) ∧ 𝑦 = 𝑤) → 𝑓 = 𝐺) | |
| 4 | simpr 110 | . . . . . . 7 ⊢ (((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) ∧ 𝑦 = 𝑤) → 𝑦 = 𝑤) | |
| 5 | 3, 4 | fveq12d 5565 | . . . . . 6 ⊢ (((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) ∧ 𝑦 = 𝑤) → (𝑓‘𝑦) = (𝐺‘𝑤)) |
| 6 | 3, 4 | reseq12d 4947 | . . . . . . 7 ⊢ (((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) ∧ 𝑦 = 𝑤) → (𝑓 ↾ 𝑦) = (𝐺 ↾ 𝑤)) |
| 7 | 6 | fveq2d 5562 | . . . . . 6 ⊢ (((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) ∧ 𝑦 = 𝑤) → (𝐹‘(𝑓 ↾ 𝑦)) = (𝐹‘(𝐺 ↾ 𝑤))) |
| 8 | 5, 7 | eqeq12d 2211 | . . . . 5 ⊢ (((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) ∧ 𝑦 = 𝑤) → ((𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)) ↔ (𝐺‘𝑤) = (𝐹‘(𝐺 ↾ 𝑤)))) |
| 9 | simpr 110 | . . . . . 6 ⊢ ((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) → 𝑥 = 𝑧) | |
| 10 | 9 | adantr 276 | . . . . 5 ⊢ (((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) ∧ 𝑦 = 𝑤) → 𝑥 = 𝑧) |
| 11 | 8, 10 | cbvraldva2 2736 | . . . 4 ⊢ ((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) → (∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)) ↔ ∀𝑤 ∈ 𝑧 (𝐺‘𝑤) = (𝐹‘(𝐺 ↾ 𝑤)))) |
| 12 | 2, 11 | anbi12d 473 | . . 3 ⊢ ((𝑓 = 𝐺 ∧ 𝑥 = 𝑧) → ((𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦))) ↔ (𝐺 Fn 𝑧 ∧ ∀𝑤 ∈ 𝑧 (𝐺‘𝑤) = (𝐹‘(𝐺 ↾ 𝑤))))) |
| 13 | 12 | cbvrexdva 2739 | . 2 ⊢ (𝑓 = 𝐺 → (∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦))) ↔ ∃𝑧 ∈ On (𝐺 Fn 𝑧 ∧ ∀𝑤 ∈ 𝑧 (𝐺‘𝑤) = (𝐹‘(𝐺 ↾ 𝑤))))) |
| 14 | tfrlem3.1 | . 2 ⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} | |
| 15 | 1, 13, 14 | elab2 2912 | 1 ⊢ (𝐺 ∈ 𝐴 ↔ ∃𝑧 ∈ On (𝐺 Fn 𝑧 ∧ ∀𝑤 ∈ 𝑧 (𝐺‘𝑤) = (𝐹‘(𝐺 ↾ 𝑤)))) |
| Colors of variables: wff set class |
| Syntax hints: ∧ wa 104 ↔ wb 105 = wceq 1364 ∈ wcel 2167 {cab 2182 ∀wral 2475 ∃wrex 2476 Vcvv 2763 Oncon0 4398 ↾ cres 4665 Fn wfn 5253 ‘cfv 5258 |
| This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-ia1 106 ax-ia2 107 ax-ia3 108 ax-io 710 ax-5 1461 ax-7 1462 ax-gen 1463 ax-ie1 1507 ax-ie2 1508 ax-8 1518 ax-10 1519 ax-11 1520 ax-i12 1521 ax-bndl 1523 ax-4 1524 ax-17 1540 ax-i9 1544 ax-ial 1548 ax-i5r 1549 ax-ext 2178 |
| This theorem depends on definitions: df-bi 117 df-3an 982 df-tru 1367 df-nf 1475 df-sb 1777 df-clab 2183 df-cleq 2189 df-clel 2192 df-nfc 2328 df-ral 2480 df-rex 2481 df-v 2765 df-un 3161 df-in 3163 df-ss 3170 df-sn 3628 df-pr 3629 df-op 3631 df-uni 3840 df-br 4034 df-opab 4095 df-xp 4669 df-rel 4670 df-cnv 4671 df-co 4672 df-dm 4673 df-res 4675 df-iota 5219 df-fun 5260 df-fn 5261 df-fv 5266 |
| This theorem is referenced by: tfrlem3 6369 tfrlem5 6372 tfrlemisucaccv 6383 tfrlemibxssdm 6385 tfrlemi14d 6391 tfrexlem 6392 |
| Copyright terms: Public domain | W3C validator |