![]() |
Metamath
Proof Explorer Theorem List (p. 77 of 431) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > MPE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Color key: | ![]() (1-28084) |
![]() (28085-29609) |
![]() (29610-43064) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | wfr2a 7601 | A weak version of wfr2 7603 which is useful for proofs that avoid the Axiom of Replacement. (Contributed by Scott Fenton, 30-Jul-2020.) |
⊢ 𝑅 We 𝐴 & ⊢ 𝑅 Se 𝐴 & ⊢ 𝐹 = wrecs(𝑅, 𝐴, 𝐺) ⇒ ⊢ (𝑋 ∈ dom 𝐹 → (𝐹‘𝑋) = (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑋)))) | ||
Theorem | wfr1 7602 | The Principle of Well-Founded Recursion, part 1 of 3. We start with an arbitrary function 𝐺. Then, using a base class 𝐴 and a well-ordering 𝑅 of 𝐴, we define a function 𝐹. This function is said to be defined by "well-founded recursion." The purpose of these three theorems is to demonstrate the properties of 𝐹. We begin by showing that 𝐹 is a function over 𝐴. (Contributed by Scott Fenton, 22-Apr-2011.) (Revised by Mario Carneiro, 26-Jun-2015.) |
⊢ 𝑅 We 𝐴 & ⊢ 𝑅 Se 𝐴 & ⊢ 𝐹 = wrecs(𝑅, 𝐴, 𝐺) ⇒ ⊢ 𝐹 Fn 𝐴 | ||
Theorem | wfr2 7603 | The Principle of Well-Founded Recursion, part 2 of 3. Next, we show that the value of 𝐹 at any 𝑋 ∈ 𝐴 is 𝐺 recursively applied to all "previous" values of 𝐹. (Contributed by Scott Fenton, 18-Apr-2011.) (Revised by Mario Carneiro, 26-Jun-2015.) |
⊢ 𝑅 We 𝐴 & ⊢ 𝑅 Se 𝐴 & ⊢ 𝐹 = wrecs(𝑅, 𝐴, 𝐺) ⇒ ⊢ (𝑋 ∈ 𝐴 → (𝐹‘𝑋) = (𝐺‘(𝐹 ↾ Pred(𝑅, 𝐴, 𝑋)))) | ||
Theorem | wfr3 7604* | The principle of Well-Founded Recursion, part 3 of 3. Finally, we show that 𝐹 is unique. We do this by showing that any function 𝐻 with the same properties we proved of 𝐹 in wfr1 7602 and wfr2 7603 is identical to 𝐹. (Contributed by Scott Fenton, 18-Apr-2011.) (Revised by Mario Carneiro, 26-Jun-2015.) |
⊢ 𝑅 We 𝐴 & ⊢ 𝑅 Se 𝐴 & ⊢ 𝐹 = wrecs(𝑅, 𝐴, 𝐺) ⇒ ⊢ ((𝐻 Fn 𝐴 ∧ ∀𝑧 ∈ 𝐴 (𝐻‘𝑧) = (𝐺‘(𝐻 ↾ Pred(𝑅, 𝐴, 𝑧)))) → 𝐹 = 𝐻) | ||
Theorem | iunon 7605* | The indexed union of a set of ordinal numbers 𝐵(𝑥) is an ordinal number. (Contributed by NM, 13-Oct-2003.) (Revised by Mario Carneiro, 5-Dec-2016.) |
⊢ ((𝐴 ∈ 𝑉 ∧ ∀𝑥 ∈ 𝐴 𝐵 ∈ On) → ∪ 𝑥 ∈ 𝐴 𝐵 ∈ On) | ||
Theorem | iinon 7606* | The nonempty indexed intersection of a class of ordinal numbers 𝐵(𝑥) is an ordinal number. (Contributed by NM, 13-Oct-2003.) (Proof shortened by Mario Carneiro, 5-Dec-2016.) |
⊢ ((∀𝑥 ∈ 𝐴 𝐵 ∈ On ∧ 𝐴 ≠ ∅) → ∩ 𝑥 ∈ 𝐴 𝐵 ∈ On) | ||
Theorem | onfununi 7607* | A property of functions on ordinal numbers. Generalization of Theorem Schema 8E of [Enderton] p. 218. (Contributed by Eric Schmidt, 26-May-2009.) |
⊢ (Lim 𝑦 → (𝐹‘𝑦) = ∪ 𝑥 ∈ 𝑦 (𝐹‘𝑥)) & ⊢ ((𝑥 ∈ On ∧ 𝑦 ∈ On ∧ 𝑥 ⊆ 𝑦) → (𝐹‘𝑥) ⊆ (𝐹‘𝑦)) ⇒ ⊢ ((𝑆 ∈ 𝑇 ∧ 𝑆 ⊆ On ∧ 𝑆 ≠ ∅) → (𝐹‘∪ 𝑆) = ∪ 𝑥 ∈ 𝑆 (𝐹‘𝑥)) | ||
Theorem | onovuni 7608* | A variant of onfununi 7607 for operations. (Contributed by Eric Schmidt, 26-May-2009.) (Revised by Mario Carneiro, 11-Sep-2015.) |
⊢ (Lim 𝑦 → (𝐴𝐹𝑦) = ∪ 𝑥 ∈ 𝑦 (𝐴𝐹𝑥)) & ⊢ ((𝑥 ∈ On ∧ 𝑦 ∈ On ∧ 𝑥 ⊆ 𝑦) → (𝐴𝐹𝑥) ⊆ (𝐴𝐹𝑦)) ⇒ ⊢ ((𝑆 ∈ 𝑇 ∧ 𝑆 ⊆ On ∧ 𝑆 ≠ ∅) → (𝐴𝐹∪ 𝑆) = ∪ 𝑥 ∈ 𝑆 (𝐴𝐹𝑥)) | ||
Theorem | onoviun 7609* | A variant of onovuni 7608 with indexed unions. (Contributed by Eric Schmidt, 26-May-2009.) (Proof shortened by Mario Carneiro, 5-Dec-2016.) |
⊢ (Lim 𝑦 → (𝐴𝐹𝑦) = ∪ 𝑥 ∈ 𝑦 (𝐴𝐹𝑥)) & ⊢ ((𝑥 ∈ On ∧ 𝑦 ∈ On ∧ 𝑥 ⊆ 𝑦) → (𝐴𝐹𝑥) ⊆ (𝐴𝐹𝑦)) ⇒ ⊢ ((𝐾 ∈ 𝑇 ∧ ∀𝑧 ∈ 𝐾 𝐿 ∈ On ∧ 𝐾 ≠ ∅) → (𝐴𝐹∪ 𝑧 ∈ 𝐾 𝐿) = ∪ 𝑧 ∈ 𝐾 (𝐴𝐹𝐿)) | ||
Theorem | onnseq 7610* | There are no length ω decreasing sequences in the ordinals. See also noinfep 8730 for a stronger version assuming Regularity. (Contributed by Mario Carneiro, 19-May-2015.) |
⊢ ((𝐹‘∅) ∈ On → ∃𝑥 ∈ ω ¬ (𝐹‘suc 𝑥) ∈ (𝐹‘𝑥)) | ||
Syntax | wsmo 7611 | Introduce the strictly monotone ordinal function. A strictly monotone function is one that is constantly increasing across the ordinals. |
wff Smo 𝐴 | ||
Definition | df-smo 7612* | Definition of a strictly monotone ordinal function. Definition 7.46 in [TakeutiZaring] p. 50. (Contributed by Andrew Salmon, 15-Nov-2011.) |
⊢ (Smo 𝐴 ↔ (𝐴:dom 𝐴⟶On ∧ Ord dom 𝐴 ∧ ∀𝑥 ∈ dom 𝐴∀𝑦 ∈ dom 𝐴(𝑥 ∈ 𝑦 → (𝐴‘𝑥) ∈ (𝐴‘𝑦)))) | ||
Theorem | dfsmo2 7613* | Alternate definition of a strictly monotone ordinal function. (Contributed by Mario Carneiro, 4-Mar-2013.) |
⊢ (Smo 𝐹 ↔ (𝐹:dom 𝐹⟶On ∧ Ord dom 𝐹 ∧ ∀𝑥 ∈ dom 𝐹∀𝑦 ∈ 𝑥 (𝐹‘𝑦) ∈ (𝐹‘𝑥))) | ||
Theorem | issmo 7614* | Conditions for which 𝐴 is a strictly monotone ordinal function. (Contributed by Andrew Salmon, 15-Nov-2011.) |
⊢ 𝐴:𝐵⟶On & ⊢ Ord 𝐵 & ⊢ ((𝑥 ∈ 𝐵 ∧ 𝑦 ∈ 𝐵) → (𝑥 ∈ 𝑦 → (𝐴‘𝑥) ∈ (𝐴‘𝑦))) & ⊢ dom 𝐴 = 𝐵 ⇒ ⊢ Smo 𝐴 | ||
Theorem | issmo2 7615* | Alternate definition of a strictly monotone ordinal function. (Contributed by Mario Carneiro, 12-Mar-2013.) |
⊢ (𝐹:𝐴⟶𝐵 → ((𝐵 ⊆ On ∧ Ord 𝐴 ∧ ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝑥 (𝐹‘𝑦) ∈ (𝐹‘𝑥)) → Smo 𝐹)) | ||
Theorem | smoeq 7616 | Equality theorem for strictly monotone functions. (Contributed by Andrew Salmon, 16-Nov-2011.) |
⊢ (𝐴 = 𝐵 → (Smo 𝐴 ↔ Smo 𝐵)) | ||
Theorem | smodm 7617 | The domain of a strictly monotone function is an ordinal. (Contributed by Andrew Salmon, 16-Nov-2011.) |
⊢ (Smo 𝐴 → Ord dom 𝐴) | ||
Theorem | smores 7618 | A strictly monotone function restricted to an ordinal remains strictly monotone. (Contributed by Andrew Salmon, 16-Nov-2011.) (Proof shortened by Mario Carneiro, 5-Dec-2016.) |
⊢ ((Smo 𝐴 ∧ 𝐵 ∈ dom 𝐴) → Smo (𝐴 ↾ 𝐵)) | ||
Theorem | smores3 7619 | A strictly monotone function restricted to an ordinal remains strictly monotone. (Contributed by Andrew Salmon, 19-Nov-2011.) |
⊢ ((Smo (𝐴 ↾ 𝐵) ∧ 𝐶 ∈ (dom 𝐴 ∩ 𝐵) ∧ Ord 𝐵) → Smo (𝐴 ↾ 𝐶)) | ||
Theorem | smores2 7620 | A strictly monotone ordinal function restricted to an ordinal is still monotone. (Contributed by Mario Carneiro, 15-Mar-2013.) |
⊢ ((Smo 𝐹 ∧ Ord 𝐴) → Smo (𝐹 ↾ 𝐴)) | ||
Theorem | smodm2 7621 | The domain of a strictly monotone ordinal function is an ordinal. (Contributed by Mario Carneiro, 12-Mar-2013.) |
⊢ ((𝐹 Fn 𝐴 ∧ Smo 𝐹) → Ord 𝐴) | ||
Theorem | smofvon2 7622 | The function values of a strictly monotone ordinal function are ordinals. (Contributed by Mario Carneiro, 12-Mar-2013.) |
⊢ (Smo 𝐹 → (𝐹‘𝐵) ∈ On) | ||
Theorem | iordsmo 7623 | The identity relation restricted to the ordinals is a strictly monotone function. (Contributed by Andrew Salmon, 16-Nov-2011.) |
⊢ Ord 𝐴 ⇒ ⊢ Smo ( I ↾ 𝐴) | ||
Theorem | smo0 7624 | The null set is a strictly monotone ordinal function. (Contributed by Andrew Salmon, 20-Nov-2011.) |
⊢ Smo ∅ | ||
Theorem | smofvon 7625 | If 𝐵 is a strictly monotone ordinal function, and 𝐴 is in the domain of 𝐵, then the value of the function at 𝐴 is an ordinal. (Contributed by Andrew Salmon, 20-Nov-2011.) |
⊢ ((Smo 𝐵 ∧ 𝐴 ∈ dom 𝐵) → (𝐵‘𝐴) ∈ On) | ||
Theorem | smoel 7626 | If 𝑥 is less than 𝑦 then a strictly monotone function's value will be strictly less at 𝑥 than at 𝑦. (Contributed by Andrew Salmon, 22-Nov-2011.) |
⊢ ((Smo 𝐵 ∧ 𝐴 ∈ dom 𝐵 ∧ 𝐶 ∈ 𝐴) → (𝐵‘𝐶) ∈ (𝐵‘𝐴)) | ||
Theorem | smoiun 7627* | The value of a strictly monotone ordinal function contains its indexed union. (Contributed by Andrew Salmon, 22-Nov-2011.) |
⊢ ((Smo 𝐵 ∧ 𝐴 ∈ dom 𝐵) → ∪ 𝑥 ∈ 𝐴 (𝐵‘𝑥) ⊆ (𝐵‘𝐴)) | ||
Theorem | smoiso 7628 | If 𝐹 is an isomorphism from an ordinal 𝐴 onto 𝐵, which is a subset of the ordinals, then 𝐹 is a strictly monotonic function. Exercise 3 in [TakeutiZaring] p. 50. (Contributed by Andrew Salmon, 24-Nov-2011.) |
⊢ ((𝐹 Isom E , E (𝐴, 𝐵) ∧ Ord 𝐴 ∧ 𝐵 ⊆ On) → Smo 𝐹) | ||
Theorem | smoel2 7629 | A strictly monotone ordinal function preserves the epsilon relation. (Contributed by Mario Carneiro, 12-Mar-2013.) |
⊢ (((𝐹 Fn 𝐴 ∧ Smo 𝐹) ∧ (𝐵 ∈ 𝐴 ∧ 𝐶 ∈ 𝐵)) → (𝐹‘𝐶) ∈ (𝐹‘𝐵)) | ||
Theorem | smo11 7630 | A strictly monotone ordinal function is one-to-one. (Contributed by Mario Carneiro, 28-Feb-2013.) |
⊢ ((𝐹:𝐴⟶𝐵 ∧ Smo 𝐹) → 𝐹:𝐴–1-1→𝐵) | ||
Theorem | smoord 7631 | A strictly monotone ordinal function preserves strict ordering. (Contributed by Mario Carneiro, 4-Mar-2013.) |
⊢ (((𝐹 Fn 𝐴 ∧ Smo 𝐹) ∧ (𝐶 ∈ 𝐴 ∧ 𝐷 ∈ 𝐴)) → (𝐶 ∈ 𝐷 ↔ (𝐹‘𝐶) ∈ (𝐹‘𝐷))) | ||
Theorem | smoword 7632 | A strictly monotone ordinal function preserves weak ordering. (Contributed by Mario Carneiro, 4-Mar-2013.) |
⊢ (((𝐹 Fn 𝐴 ∧ Smo 𝐹) ∧ (𝐶 ∈ 𝐴 ∧ 𝐷 ∈ 𝐴)) → (𝐶 ⊆ 𝐷 ↔ (𝐹‘𝐶) ⊆ (𝐹‘𝐷))) | ||
Theorem | smogt 7633 | A strictly monotone ordinal function is greater than or equal to its argument. Exercise 1 in [TakeutiZaring] p. 50. (Contributed by Andrew Salmon, 23-Nov-2011.) (Revised by Mario Carneiro, 28-Feb-2013.) |
⊢ ((𝐹 Fn 𝐴 ∧ Smo 𝐹 ∧ 𝐶 ∈ 𝐴) → 𝐶 ⊆ (𝐹‘𝐶)) | ||
Theorem | smorndom 7634 | The range of a strictly monotone ordinal function dominates the domain. (Contributed by Mario Carneiro, 13-Mar-2013.) |
⊢ ((𝐹:𝐴⟶𝐵 ∧ Smo 𝐹 ∧ Ord 𝐵) → 𝐴 ⊆ 𝐵) | ||
Theorem | smoiso2 7635 | The strictly monotone ordinal functions are also epsilon isomorphisms of subclasses of On. (Contributed by Mario Carneiro, 20-Mar-2013.) |
⊢ ((Ord 𝐴 ∧ 𝐵 ⊆ On) → ((𝐹:𝐴–onto→𝐵 ∧ Smo 𝐹) ↔ 𝐹 Isom E , E (𝐴, 𝐵))) | ||
Syntax | crecs 7636 | Notation for a function defined by strong transfinite recursion. |
class recs(𝐹) | ||
Definition | df-recs 7637 | Define a function recs(𝐹) on On, the class of ordinal numbers, by transfinite recursion given a rule 𝐹 which sets the next value given all values so far. See df-rdg 7675 for more details on why this definition is desirable. Unlike df-rdg 7675 which restricts the update rule to use only the previous value, this version allows the update rule to use all previous values, which is why it is described as "strong", although it is actually more primitive. See recsfnon 7668 and recsval 7669 for the primary contract of this definition. (Contributed by Stefan O'Rear, 18-Jan-2015.) (Revised by Scott Fenton, 3-Aug-2020.) |
⊢ recs(𝐹) = wrecs( E , On, 𝐹) | ||
Theorem | dfrecs3 7638* | The old definition of transfinite recursion. This version is preferred for development, as it demonstrates the properties of transfinite recursion without relying on well-founded recursion. (Contributed by Scott Fenton, 3-Aug-2020.) |
⊢ recs(𝐹) = ∪ {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} | ||
Theorem | recseq 7639 | Equality theorem for recs. (Contributed by Stefan O'Rear, 18-Jan-2015.) |
⊢ (𝐹 = 𝐺 → recs(𝐹) = recs(𝐺)) | ||
Theorem | nfrecs 7640 | Bound-variable hypothesis builder for recs. (Contributed by Stefan O'Rear, 18-Jan-2015.) |
⊢ Ⅎ𝑥𝐹 ⇒ ⊢ Ⅎ𝑥recs(𝐹) | ||
Theorem | tfrlem1 7641* | A technical lemma for transfinite recursion. Compare Lemma 1 of [TakeutiZaring] p. 47. (Contributed by NM, 23-Mar-1995.) (Revised by Mario Carneiro, 24-May-2019.) |
⊢ (𝜑 → 𝐴 ∈ On) & ⊢ (𝜑 → (Fun 𝐹 ∧ 𝐴 ⊆ dom 𝐹)) & ⊢ (𝜑 → (Fun 𝐺 ∧ 𝐴 ⊆ dom 𝐺)) & ⊢ (𝜑 → ∀𝑥 ∈ 𝐴 (𝐹‘𝑥) = (𝐵‘(𝐹 ↾ 𝑥))) & ⊢ (𝜑 → ∀𝑥 ∈ 𝐴 (𝐺‘𝑥) = (𝐵‘(𝐺 ↾ 𝑥))) ⇒ ⊢ (𝜑 → ∀𝑥 ∈ 𝐴 (𝐹‘𝑥) = (𝐺‘𝑥)) | ||
Theorem | tfrlem3a 7642* | 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.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} & ⊢ 𝐺 ∈ V ⇒ ⊢ (𝐺 ∈ 𝐴 ↔ ∃𝑧 ∈ On (𝐺 Fn 𝑧 ∧ ∀𝑤 ∈ 𝑧 (𝐺‘𝑤) = (𝐹‘(𝐺 ↾ 𝑤)))) | ||
Theorem | tfrlem3 7643* | 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.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ 𝐴 = {𝑔 ∣ ∃𝑧 ∈ On (𝑔 Fn 𝑧 ∧ ∀𝑤 ∈ 𝑧 (𝑔‘𝑤) = (𝐹‘(𝑔 ↾ 𝑤)))} | ||
Theorem | tfrlem4 7644* | Lemma for transfinite recursion. 𝐴 is the class of all "acceptable" functions, and 𝐹 is their union. First we show that an acceptable function is in fact a function. (Contributed by NM, 9-Apr-1995.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ (𝑔 ∈ 𝐴 → Fun 𝑔) | ||
Theorem | tfrlem5 7645* | Lemma for transfinite recursion. The values of two acceptable functions are the same within their domains. (Contributed by NM, 9-Apr-1995.) (Revised by Mario Carneiro, 24-May-2019.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ ((𝑔 ∈ 𝐴 ∧ ℎ ∈ 𝐴) → ((𝑥𝑔𝑢 ∧ 𝑥ℎ𝑣) → 𝑢 = 𝑣)) | ||
Theorem | recsfval 7646* | Lemma for transfinite recursion. The definition recs is the union of all acceptable functions. (Contributed by Mario Carneiro, 9-May-2015.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ recs(𝐹) = ∪ 𝐴 | ||
Theorem | tfrlem6 7647* | Lemma for transfinite recursion. The union of all acceptable functions is a relation. (Contributed by NM, 8-Aug-1994.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ Rel recs(𝐹) | ||
Theorem | tfrlem7 7648* | Lemma for transfinite recursion. The union of all acceptable functions is a function. (Contributed by NM, 9-Aug-1994.) (Revised by Mario Carneiro, 24-May-2019.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ Fun recs(𝐹) | ||
Theorem | tfrlem8 7649* | Lemma for transfinite recursion. The domain of recs is an ordinal. (Contributed by NM, 14-Aug-1994.) (Proof shortened by Alan Sare, 11-Mar-2008.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ Ord dom recs(𝐹) | ||
Theorem | tfrlem9 7650* | Lemma for transfinite recursion. Here we compute the value of recs (the union of all acceptable functions). (Contributed by NM, 17-Aug-1994.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ (𝐵 ∈ dom recs(𝐹) → (recs(𝐹)‘𝐵) = (𝐹‘(recs(𝐹) ↾ 𝐵))) | ||
Theorem | tfrlem9a 7651* | Lemma for transfinite recursion. Without using ax-rep 4923, show that all the restrictions of recs are sets. (Contributed by Mario Carneiro, 16-Nov-2014.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ (𝐵 ∈ dom recs(𝐹) → (recs(𝐹) ↾ 𝐵) ∈ V) | ||
Theorem | tfrlem10 7652* | Lemma for transfinite recursion. We define class 𝐶 by extending recs with one ordered pair. We will assume, falsely, that domain of recs is a member of, and thus not equal to, On. Using this assumption we will prove facts about 𝐶 that will lead to a contradiction in tfrlem14 7656, thus showing the domain of recs does in fact equal On. Here we show (under the false assumption) that 𝐶 is a function extending the domain of recs(𝐹) by one. (Contributed by NM, 14-Aug-1994.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} & ⊢ 𝐶 = (recs(𝐹) ∪ {〈dom recs(𝐹), (𝐹‘recs(𝐹))〉}) ⇒ ⊢ (dom recs(𝐹) ∈ On → 𝐶 Fn suc dom recs(𝐹)) | ||
Theorem | tfrlem11 7653* | Lemma for transfinite recursion. Compute the value of 𝐶. (Contributed by NM, 18-Aug-1994.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} & ⊢ 𝐶 = (recs(𝐹) ∪ {〈dom recs(𝐹), (𝐹‘recs(𝐹))〉}) ⇒ ⊢ (dom recs(𝐹) ∈ On → (𝐵 ∈ suc dom recs(𝐹) → (𝐶‘𝐵) = (𝐹‘(𝐶 ↾ 𝐵)))) | ||
Theorem | tfrlem12 7654* | Lemma for transfinite recursion. Show 𝐶 is an acceptable function. (Contributed by NM, 15-Aug-1994.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} & ⊢ 𝐶 = (recs(𝐹) ∪ {〈dom recs(𝐹), (𝐹‘recs(𝐹))〉}) ⇒ ⊢ (recs(𝐹) ∈ V → 𝐶 ∈ 𝐴) | ||
Theorem | tfrlem13 7655* | Lemma for transfinite recursion. If recs is a set function, then 𝐶 is acceptable, and thus a subset of recs, but dom 𝐶 is bigger than dom recs. This is a contradiction, so recs must be a proper class function. (Contributed by NM, 14-Aug-1994.) (Revised by Mario Carneiro, 14-Nov-2014.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ ¬ recs(𝐹) ∈ V | ||
Theorem | tfrlem14 7656* | Lemma for transfinite recursion. Assuming ax-rep 4923, dom recs ∈ V ↔ recs ∈ V, so since dom recs is an ordinal, it must be equal to On. (Contributed by NM, 14-Aug-1994.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ dom recs(𝐹) = On | ||
Theorem | tfrlem15 7657* | Lemma for transfinite recursion. Without assuming ax-rep 4923, we can show that all proper initial subsets of recs are sets, while nothing larger is a set. (Contributed by Mario Carneiro, 14-Nov-2014.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ (𝐵 ∈ On → (𝐵 ∈ dom recs(𝐹) ↔ (recs(𝐹) ↾ 𝐵) ∈ V)) | ||
Theorem | tfrlem16 7658* | Lemma for finite recursion. Without assuming ax-rep 4923, we can show that the domain of the constructed function is a limit ordinal, and hence contains all the finite ordinals. (Contributed by Mario Carneiro, 14-Nov-2014.) |
⊢ 𝐴 = {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐹‘(𝑓 ↾ 𝑦)))} ⇒ ⊢ Lim dom recs(𝐹) | ||
Theorem | tfr1a 7659 | A weak version of tfr1 7662 which is useful for proofs that avoid the Axiom of Replacement. (Contributed by Mario Carneiro, 24-Jun-2015.) |
⊢ 𝐹 = recs(𝐺) ⇒ ⊢ (Fun 𝐹 ∧ Lim dom 𝐹) | ||
Theorem | tfr2a 7660 | A weak version of tfr2 7663 which is useful for proofs that avoid the Axiom of Replacement. (Contributed by Mario Carneiro, 24-Jun-2015.) |
⊢ 𝐹 = recs(𝐺) ⇒ ⊢ (𝐴 ∈ dom 𝐹 → (𝐹‘𝐴) = (𝐺‘(𝐹 ↾ 𝐴))) | ||
Theorem | tfr2b 7661 | Without assuming ax-rep 4923, we can show that all proper initial subsets of recs are sets, while nothing larger is a set. (Contributed by Mario Carneiro, 24-Jun-2015.) |
⊢ 𝐹 = recs(𝐺) ⇒ ⊢ (Ord 𝐴 → (𝐴 ∈ dom 𝐹 ↔ (𝐹 ↾ 𝐴) ∈ V)) | ||
Theorem | tfr1 7662 | Principle of Transfinite Recursion, part 1 of 3. Theorem 7.41(1) of [TakeutiZaring] p. 47. We start with an arbitrary class 𝐺, normally a function, and define a class 𝐴 of all "acceptable" functions. The final function we're interested in is the union 𝐹 = recs(𝐺) of them. 𝐹 is then said to be defined by transfinite recursion. The purpose of the 3 parts of this theorem is to demonstrate properties of 𝐹. In this first part we show that 𝐹 is a function whose domain is all ordinal numbers. (Contributed by NM, 17-Aug-1994.) (Revised by Mario Carneiro, 18-Jan-2015.) |
⊢ 𝐹 = recs(𝐺) ⇒ ⊢ 𝐹 Fn On | ||
Theorem | tfr2 7663 | Principle of Transfinite Recursion, part 2 of 3. Theorem 7.41(2) of [TakeutiZaring] p. 47. Here we show that the function 𝐹 has the property that for any function 𝐺 whatsoever, the "next" value of 𝐹 is 𝐺 recursively applied to all "previous" values of 𝐹. (Contributed by NM, 9-Apr-1995.) (Revised by Stefan O'Rear, 18-Jan-2015.) |
⊢ 𝐹 = recs(𝐺) ⇒ ⊢ (𝐴 ∈ On → (𝐹‘𝐴) = (𝐺‘(𝐹 ↾ 𝐴))) | ||
Theorem | tfr3 7664* | Principle of Transfinite Recursion, part 3 of 3. Theorem 7.41(3) of [TakeutiZaring] p. 47. Finally, we show that 𝐹 is unique. We do this by showing that any class 𝐵 with the same properties of 𝐹 that we showed in parts 1 and 2 is identical to 𝐹. (Contributed by NM, 18-Aug-1994.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ 𝐹 = recs(𝐺) ⇒ ⊢ ((𝐵 Fn On ∧ ∀𝑥 ∈ On (𝐵‘𝑥) = (𝐺‘(𝐵 ↾ 𝑥))) → 𝐵 = 𝐹) | ||
Theorem | tfr1ALT 7665 | Alternate proof of tfr1 7662 using well-founded recursion. (Contributed by Scott Fenton, 3-Aug-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ 𝐹 = recs(𝐺) ⇒ ⊢ 𝐹 Fn On | ||
Theorem | tfr2ALT 7666 | Alternate proof of tfr2 7663 using well-founded recursion. (Contributed by Scott Fenton, 3-Aug-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ 𝐹 = recs(𝐺) ⇒ ⊢ (𝐴 ∈ On → (𝐹‘𝐴) = (𝐺‘(𝐹 ↾ 𝐴))) | ||
Theorem | tfr3ALT 7667* | Alternate proof of tfr3 7664 using well-founded recursion. (Contributed by Scott Fenton, 3-Aug-2020.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ 𝐹 = recs(𝐺) ⇒ ⊢ ((𝐵 Fn On ∧ ∀𝑥 ∈ On (𝐵‘𝑥) = (𝐺‘(𝐵 ↾ 𝑥))) → 𝐵 = 𝐹) | ||
Theorem | recsfnon 7668 | Strong transfinite recursion defines a function on ordinals. (Contributed by Stefan O'Rear, 18-Jan-2015.) |
⊢ recs(𝐹) Fn On | ||
Theorem | recsval 7669 | Strong transfinite recursion in terms of all previous values. (Contributed by Stefan O'Rear, 18-Jan-2015.) |
⊢ (𝐴 ∈ On → (recs(𝐹)‘𝐴) = (𝐹‘(recs(𝐹) ↾ 𝐴))) | ||
Theorem | tz7.44lem1 7670* | 𝐺 is a function. Lemma for tz7.44-1 7671, tz7.44-2 7672, and tz7.44-3 7673. (Contributed by NM, 23-Apr-1995.) (Revised by David Abernethy, 19-Jun-2012.) |
⊢ 𝐺 = {〈𝑥, 𝑦〉 ∣ ((𝑥 = ∅ ∧ 𝑦 = 𝐴) ∨ (¬ (𝑥 = ∅ ∨ Lim dom 𝑥) ∧ 𝑦 = (𝐻‘(𝑥‘∪ dom 𝑥))) ∨ (Lim dom 𝑥 ∧ 𝑦 = ∪ ran 𝑥))} ⇒ ⊢ Fun 𝐺 | ||
Theorem | tz7.44-1 7671* | The value of 𝐹 at ∅. Part 1 of Theorem 7.44 of [TakeutiZaring] p. 49. (Contributed by NM, 23-Apr-1995.) (Revised by Mario Carneiro, 14-Nov-2014.) |
⊢ 𝐺 = (𝑥 ∈ V ↦ if(𝑥 = ∅, 𝐴, if(Lim dom 𝑥, ∪ ran 𝑥, (𝐻‘(𝑥‘∪ dom 𝑥))))) & ⊢ (𝑦 ∈ 𝑋 → (𝐹‘𝑦) = (𝐺‘(𝐹 ↾ 𝑦))) & ⊢ 𝐴 ∈ V ⇒ ⊢ (∅ ∈ 𝑋 → (𝐹‘∅) = 𝐴) | ||
Theorem | tz7.44-2 7672* | The value of 𝐹 at a successor ordinal. Part 2 of Theorem 7.44 of [TakeutiZaring] p. 49. (Unnecessary distinct variable restrictions were removed by David Abernethy, 19-Jun-2012.) (Contributed by NM, 23-Apr-1995.) (Revised by Mario Carneiro, 14-Nov-2014.) |
⊢ 𝐺 = (𝑥 ∈ V ↦ if(𝑥 = ∅, 𝐴, if(Lim dom 𝑥, ∪ ran 𝑥, (𝐻‘(𝑥‘∪ dom 𝑥))))) & ⊢ (𝑦 ∈ 𝑋 → (𝐹‘𝑦) = (𝐺‘(𝐹 ↾ 𝑦))) & ⊢ (𝑦 ∈ 𝑋 → (𝐹 ↾ 𝑦) ∈ V) & ⊢ 𝐹 Fn 𝑋 & ⊢ Ord 𝑋 ⇒ ⊢ (suc 𝐵 ∈ 𝑋 → (𝐹‘suc 𝐵) = (𝐻‘(𝐹‘𝐵))) | ||
Theorem | tz7.44-3 7673* | The value of 𝐹 at a limit ordinal. Part 3 of Theorem 7.44 of [TakeutiZaring] p. 49. (Contributed by NM, 23-Apr-1995.) (Revised by David Abernethy, 19-Jun-2012.) |
⊢ 𝐺 = (𝑥 ∈ V ↦ if(𝑥 = ∅, 𝐴, if(Lim dom 𝑥, ∪ ran 𝑥, (𝐻‘(𝑥‘∪ dom 𝑥))))) & ⊢ (𝑦 ∈ 𝑋 → (𝐹‘𝑦) = (𝐺‘(𝐹 ↾ 𝑦))) & ⊢ (𝑦 ∈ 𝑋 → (𝐹 ↾ 𝑦) ∈ V) & ⊢ 𝐹 Fn 𝑋 & ⊢ Ord 𝑋 ⇒ ⊢ ((𝐵 ∈ 𝑋 ∧ Lim 𝐵) → (𝐹‘𝐵) = ∪ (𝐹 “ 𝐵)) | ||
Syntax | crdg 7674 | Extend class notation with the recursive definition generator, with characteristic function 𝐹 and initial value 𝐼. |
class rec(𝐹, 𝐼) | ||
Definition | df-rdg 7675* |
Define a recursive definition generator on On (the
class of ordinal
numbers) with characteristic function 𝐹 and initial value 𝐼.
This combines functions 𝐹 in tfr1 7662
and 𝐺 in tz7.44-1 7671 into one
definition. This rather amazing operation allows us to define, with
compact direct definitions, functions that are usually defined in
textbooks only with indirect self-referencing recursive definitions. A
recursive definition requires advanced metalogic to justify - in
particular, eliminating a recursive definition is very difficult and
often not even shown in textbooks. On the other hand, the elimination
of a direct definition is a matter of simple mechanical substitution.
The price paid is the daunting complexity of our rec operation
(especially when df-recs 7637 that it is built on is also eliminated). But
once we get past this hurdle, definitions that would otherwise be
recursive become relatively simple, as in for example oav 7760,
from which
we prove the recursive textbook definition as theorems oa0 7765,
oasuc 7773,
and oalim 7781 (with the help of theorems rdg0 7686,
rdgsuc 7689, and
rdglim2a 7698). We can also restrict the rec operation to define
otherwise recursive functions on the natural numbers ω; see
fr0g 7700 and frsuc 7701. Our rec
operation apparently does not appear
in published literature, although closely related is Definition 25.2 of
[Quine] p. 177, which he uses to
"turn...a recursion into a genuine or
direct definition" (p. 174). Note that the if operations (see
df-if 4231) select cases based on whether the domain of
𝑔
is zero, a
successor, or a limit ordinal.
An important use of this definition is in the recursive sequence generator df-seq 12996 on the natural numbers (as a subset of the complex numbers), allowing us to define, with direct definitions, recursive infinite sequences such as the factorial function df-fac 13255 and integer powers df-exp 13055. Note: We introduce rec with the philosophical goal of being able to eliminate all definitions with direct mechanical substitution and to verify easily the soundness of definitions. Metamath itself has no built-in technical limitation that prevents multiple-part recursive definitions in the traditional textbook style. (Contributed by NM, 9-Apr-1995.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ rec(𝐹, 𝐼) = recs((𝑔 ∈ V ↦ if(𝑔 = ∅, 𝐼, if(Lim dom 𝑔, ∪ ran 𝑔, (𝐹‘(𝑔‘∪ dom 𝑔)))))) | ||
Theorem | rdgeq1 7676 | Equality theorem for the recursive definition generator. (Contributed by NM, 9-Apr-1995.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ (𝐹 = 𝐺 → rec(𝐹, 𝐴) = rec(𝐺, 𝐴)) | ||
Theorem | rdgeq2 7677 | Equality theorem for the recursive definition generator. (Contributed by NM, 9-Apr-1995.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ (𝐴 = 𝐵 → rec(𝐹, 𝐴) = rec(𝐹, 𝐵)) | ||
Theorem | rdgeq12 7678 | Equality theorem for the recursive definition generator. (Contributed by Scott Fenton, 28-Apr-2012.) |
⊢ ((𝐹 = 𝐺 ∧ 𝐴 = 𝐵) → rec(𝐹, 𝐴) = rec(𝐺, 𝐵)) | ||
Theorem | nfrdg 7679 | Bound-variable hypothesis builder for the recursive definition generator. (Contributed by NM, 14-Sep-2003.) (Revised by Mario Carneiro, 8-Sep-2013.) |
⊢ Ⅎ𝑥𝐹 & ⊢ Ⅎ𝑥𝐴 ⇒ ⊢ Ⅎ𝑥rec(𝐹, 𝐴) | ||
Theorem | rdglem1 7680* | Lemma used with the recursive definition generator. This is a trivial lemma that just changes bound variables for later use. (Contributed by NM, 9-Apr-1995.) |
⊢ {𝑓 ∣ ∃𝑥 ∈ On (𝑓 Fn 𝑥 ∧ ∀𝑦 ∈ 𝑥 (𝑓‘𝑦) = (𝐺‘(𝑓 ↾ 𝑦)))} = {𝑔 ∣ ∃𝑧 ∈ On (𝑔 Fn 𝑧 ∧ ∀𝑤 ∈ 𝑧 (𝑔‘𝑤) = (𝐺‘(𝑔 ↾ 𝑤)))} | ||
Theorem | rdgfun 7681 | The recursive definition generator is a function. (Contributed by Mario Carneiro, 16-Nov-2014.) |
⊢ Fun rec(𝐹, 𝐴) | ||
Theorem | rdgdmlim 7682 | The domain of the recursive definition generator is a limit ordinal. (Contributed by NM, 16-Nov-2014.) |
⊢ Lim dom rec(𝐹, 𝐴) | ||
Theorem | rdgfnon 7683 | The recursive definition generator is a function on ordinal numbers. (Contributed by NM, 9-Apr-1995.) (Revised by Mario Carneiro, 9-May-2015.) |
⊢ rec(𝐹, 𝐴) Fn On | ||
Theorem | rdgvalg 7684* | Value of the recursive definition generator. (Contributed by NM, 9-Apr-1995.) (Revised by Mario Carneiro, 8-Sep-2013.) |
⊢ (𝐵 ∈ dom rec(𝐹, 𝐴) → (rec(𝐹, 𝐴)‘𝐵) = ((𝑔 ∈ V ↦ if(𝑔 = ∅, 𝐴, if(Lim dom 𝑔, ∪ ran 𝑔, (𝐹‘(𝑔‘∪ dom 𝑔)))))‘(rec(𝐹, 𝐴) ↾ 𝐵))) | ||
Theorem | rdgval 7685* | Value of the recursive definition generator. (Contributed by NM, 9-Apr-1995.) (Revised by Mario Carneiro, 8-Sep-2013.) |
⊢ (𝐵 ∈ On → (rec(𝐹, 𝐴)‘𝐵) = ((𝑔 ∈ V ↦ if(𝑔 = ∅, 𝐴, if(Lim dom 𝑔, ∪ ran 𝑔, (𝐹‘(𝑔‘∪ dom 𝑔)))))‘(rec(𝐹, 𝐴) ↾ 𝐵))) | ||
Theorem | rdg0 7686 | The initial value of the recursive definition generator. (Contributed by NM, 23-Apr-1995.) (Revised by Mario Carneiro, 14-Nov-2014.) |
⊢ 𝐴 ∈ V ⇒ ⊢ (rec(𝐹, 𝐴)‘∅) = 𝐴 | ||
Theorem | rdgseg 7687 | The initial segments of the recursive definition generator are sets. (Contributed by Mario Carneiro, 16-Nov-2014.) |
⊢ (𝐵 ∈ dom rec(𝐹, 𝐴) → (rec(𝐹, 𝐴) ↾ 𝐵) ∈ V) | ||
Theorem | rdgsucg 7688 | The value of the recursive definition generator at a successor. (Contributed by NM, 16-Nov-2014.) |
⊢ (𝐵 ∈ dom rec(𝐹, 𝐴) → (rec(𝐹, 𝐴)‘suc 𝐵) = (𝐹‘(rec(𝐹, 𝐴)‘𝐵))) | ||
Theorem | rdgsuc 7689 | The value of the recursive definition generator at a successor. (Contributed by NM, 23-Apr-1995.) (Revised by Mario Carneiro, 14-Nov-2014.) |
⊢ (𝐵 ∈ On → (rec(𝐹, 𝐴)‘suc 𝐵) = (𝐹‘(rec(𝐹, 𝐴)‘𝐵))) | ||
Theorem | rdglimg 7690 | The value of the recursive definition generator at a limit ordinal. (Contributed by NM, 16-Nov-2014.) |
⊢ ((𝐵 ∈ dom rec(𝐹, 𝐴) ∧ Lim 𝐵) → (rec(𝐹, 𝐴)‘𝐵) = ∪ (rec(𝐹, 𝐴) “ 𝐵)) | ||
Theorem | rdglim 7691 | The value of the recursive definition generator at a limit ordinal. (Contributed by NM, 23-Apr-1995.) (Revised by Mario Carneiro, 14-Nov-2014.) |
⊢ ((𝐵 ∈ 𝐶 ∧ Lim 𝐵) → (rec(𝐹, 𝐴)‘𝐵) = ∪ (rec(𝐹, 𝐴) “ 𝐵)) | ||
Theorem | rdg0g 7692 | The initial value of the recursive definition generator. (Contributed by NM, 25-Apr-1995.) |
⊢ (𝐴 ∈ 𝐶 → (rec(𝐹, 𝐴)‘∅) = 𝐴) | ||
Theorem | rdgsucmptf 7693 | The value of the recursive definition generator at a successor (special case where the characteristic function uses the map operation). (Contributed by NM, 22-Oct-2003.) (Revised by Mario Carneiro, 15-Oct-2016.) |
⊢ Ⅎ𝑥𝐴 & ⊢ Ⅎ𝑥𝐵 & ⊢ Ⅎ𝑥𝐷 & ⊢ 𝐹 = rec((𝑥 ∈ V ↦ 𝐶), 𝐴) & ⊢ (𝑥 = (𝐹‘𝐵) → 𝐶 = 𝐷) ⇒ ⊢ ((𝐵 ∈ On ∧ 𝐷 ∈ 𝑉) → (𝐹‘suc 𝐵) = 𝐷) | ||
Theorem | rdgsucmptnf 7694 | The value of the recursive definition generator at a successor (special case where the characteristic function is an ordered-pair class abstraction and where the mapping class 𝐷 is a proper class). This is a technical lemma that can be used together with rdgsucmptf 7693 to help eliminate redundant sethood antecedents. (Contributed by NM, 22-Oct-2003.) (Revised by Mario Carneiro, 15-Oct-2016.) |
⊢ Ⅎ𝑥𝐴 & ⊢ Ⅎ𝑥𝐵 & ⊢ Ⅎ𝑥𝐷 & ⊢ 𝐹 = rec((𝑥 ∈ V ↦ 𝐶), 𝐴) & ⊢ (𝑥 = (𝐹‘𝐵) → 𝐶 = 𝐷) ⇒ ⊢ (¬ 𝐷 ∈ V → (𝐹‘suc 𝐵) = ∅) | ||
Theorem | rdgsucmpt2 7695* | This version of rdgsucmpt 7696 avoids the not-free hypothesis of rdgsucmptf 7693 by using two substitutions instead of one. (Contributed by Mario Carneiro, 11-Sep-2015.) |
⊢ 𝐹 = rec((𝑥 ∈ V ↦ 𝐶), 𝐴) & ⊢ (𝑦 = 𝑥 → 𝐸 = 𝐶) & ⊢ (𝑦 = (𝐹‘𝐵) → 𝐸 = 𝐷) ⇒ ⊢ ((𝐵 ∈ On ∧ 𝐷 ∈ 𝑉) → (𝐹‘suc 𝐵) = 𝐷) | ||
Theorem | rdgsucmpt 7696* | The value of the recursive definition generator at a successor (special case where the characteristic function uses the map operation). (Contributed by Mario Carneiro, 9-Sep-2013.) |
⊢ 𝐹 = rec((𝑥 ∈ V ↦ 𝐶), 𝐴) & ⊢ (𝑥 = (𝐹‘𝐵) → 𝐶 = 𝐷) ⇒ ⊢ ((𝐵 ∈ On ∧ 𝐷 ∈ 𝑉) → (𝐹‘suc 𝐵) = 𝐷) | ||
Theorem | rdglim2 7697* | The value of the recursive definition generator at a limit ordinal, in terms of the union of all smaller values. (Contributed by NM, 23-Apr-1995.) |
⊢ ((𝐵 ∈ 𝐶 ∧ Lim 𝐵) → (rec(𝐹, 𝐴)‘𝐵) = ∪ {𝑦 ∣ ∃𝑥 ∈ 𝐵 𝑦 = (rec(𝐹, 𝐴)‘𝑥)}) | ||
Theorem | rdglim2a 7698* | The value of the recursive definition generator at a limit ordinal, in terms of indexed union of all smaller values. (Contributed by NM, 28-Jun-1998.) |
⊢ ((𝐵 ∈ 𝐶 ∧ Lim 𝐵) → (rec(𝐹, 𝐴)‘𝐵) = ∪ 𝑥 ∈ 𝐵 (rec(𝐹, 𝐴)‘𝑥)) | ||
Theorem | frfnom 7699 | The function generated by finite recursive definition generation is a function on omega. (Contributed by NM, 15-Oct-1996.) (Revised by Mario Carneiro, 14-Nov-2014.) |
⊢ (rec(𝐹, 𝐴) ↾ ω) Fn ω | ||
Theorem | fr0g 7700 | The initial value resulting from finite recursive definition generation. (Contributed by NM, 15-Oct-1996.) |
⊢ (𝐴 ∈ 𝐵 → ((rec(𝐹, 𝐴) ↾ ω)‘∅) = 𝐴) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |