![]() |
Intuitionistic Logic Explorer Theorem List (p. 95 of 108) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > ILE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | frec2uz0d 9401* | The mapping 𝐺 is a one-to-one mapping from ω onto upper integers that will be used to construct a recursive definition generator. Ordinal natural number 0 maps to complex number 𝐶 (normally 0 for the upper integers ℕ0 or 1 for the upper integers ℕ), 1 maps to 𝐶 + 1, etc. This theorem shows the value of 𝐺 at ordinal natural number zero. (Contributed by Jim Kingdon, 16-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) ⇒ ⊢ (𝜑 → (𝐺‘∅) = 𝐶) | ||
Theorem | frec2uzzd 9402* | The value of 𝐺 (see frec2uz0d 9401) is an integer. (Contributed by Jim Kingdon, 16-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝐴 ∈ ω) ⇒ ⊢ (𝜑 → (𝐺‘𝐴) ∈ ℤ) | ||
Theorem | frec2uzsucd 9403* | The value of 𝐺 (see frec2uz0d 9401) at a successor. (Contributed by Jim Kingdon, 16-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝐴 ∈ ω) ⇒ ⊢ (𝜑 → (𝐺‘suc 𝐴) = ((𝐺‘𝐴) + 1)) | ||
Theorem | frec2uzuzd 9404* | The value 𝐺 (see frec2uz0d 9401) at an ordinal natural number is in the upper integers. (Contributed by Jim Kingdon, 16-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝐴 ∈ ω) ⇒ ⊢ (𝜑 → (𝐺‘𝐴) ∈ (ℤ≥‘𝐶)) | ||
Theorem | frec2uzltd 9405* | Less-than relation for 𝐺 (see frec2uz0d 9401). (Contributed by Jim Kingdon, 16-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝐴 ∈ ω) & ⊢ (𝜑 → 𝐵 ∈ ω) ⇒ ⊢ (𝜑 → (𝐴 ∈ 𝐵 → (𝐺‘𝐴) < (𝐺‘𝐵))) | ||
Theorem | frec2uzlt2d 9406* | The mapping 𝐺 (see frec2uz0d 9401) preserves order. (Contributed by Jim Kingdon, 16-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝐴 ∈ ω) & ⊢ (𝜑 → 𝐵 ∈ ω) ⇒ ⊢ (𝜑 → (𝐴 ∈ 𝐵 ↔ (𝐺‘𝐴) < (𝐺‘𝐵))) | ||
Theorem | frec2uzrand 9407* | Range of 𝐺 (see frec2uz0d 9401). (Contributed by Jim Kingdon, 17-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) ⇒ ⊢ (𝜑 → ran 𝐺 = (ℤ≥‘𝐶)) | ||
Theorem | frec2uzf1od 9408* | 𝐺 (see frec2uz0d 9401) is a one-to-one onto mapping. (Contributed by Jim Kingdon, 17-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) ⇒ ⊢ (𝜑 → 𝐺:ω–1-1-onto→(ℤ≥‘𝐶)) | ||
Theorem | frec2uzisod 9409* | 𝐺 (see frec2uz0d 9401) is an isomorphism from natural ordinals to upper integers. (Contributed by Jim Kingdon, 17-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) ⇒ ⊢ (𝜑 → 𝐺 Isom E , < (ω, (ℤ≥‘𝐶))) | ||
Theorem | frecuzrdgrrn 9410* | The function 𝑅 (used in the definition of the recursive definition generator on upper integers) yields ordered pairs of integers and elements of 𝑆. (Contributed by Jim Kingdon, 27-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝐴 ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ (ℤ≥‘𝐶) ∧ 𝑦 ∈ 𝑆)) → (𝑥𝐹𝑦) ∈ 𝑆) & ⊢ 𝑅 = frec((𝑥 ∈ (ℤ≥‘𝐶), 𝑦 ∈ 𝑆 ↦ 〈(𝑥 + 1), (𝑥𝐹𝑦)〉), 〈𝐶, 𝐴〉) ⇒ ⊢ ((𝜑 ∧ 𝐷 ∈ ω) → (𝑅‘𝐷) ∈ ((ℤ≥‘𝐶) × 𝑆)) | ||
Theorem | frec2uzrdg 9411* | A helper lemma for the value of a recursive definition generator on upper integers (typically either ℕ or ℕ0) with characteristic function 𝐹(𝑥, 𝑦) and initial value 𝐴. This lemma shows that evaluating 𝑅 at an element of ω gives an ordered pair whose first element is the index (translated from ω to (ℤ≥‘𝐶)). See comment in frec2uz0d 9401 which describes 𝐺 and the index translation. (Contributed by Jim Kingdon, 24-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝐴 ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ (ℤ≥‘𝐶) ∧ 𝑦 ∈ 𝑆)) → (𝑥𝐹𝑦) ∈ 𝑆) & ⊢ 𝑅 = frec((𝑥 ∈ (ℤ≥‘𝐶), 𝑦 ∈ 𝑆 ↦ 〈(𝑥 + 1), (𝑥𝐹𝑦)〉), 〈𝐶, 𝐴〉) & ⊢ (𝜑 → 𝐵 ∈ ω) ⇒ ⊢ (𝜑 → (𝑅‘𝐵) = 〈(𝐺‘𝐵), (2nd ‘(𝑅‘𝐵))〉) | ||
Theorem | frecuzrdgrom 9412* | The function 𝑅 (used in the definition of the recursive definition generator on upper integers) is a function defined for all natural numbers. (Contributed by Jim Kingdon, 26-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝐴 ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ (ℤ≥‘𝐶) ∧ 𝑦 ∈ 𝑆)) → (𝑥𝐹𝑦) ∈ 𝑆) & ⊢ 𝑅 = frec((𝑥 ∈ (ℤ≥‘𝐶), 𝑦 ∈ 𝑆 ↦ 〈(𝑥 + 1), (𝑥𝐹𝑦)〉), 〈𝐶, 𝐴〉) ⇒ ⊢ (𝜑 → 𝑅 Fn ω) | ||
Theorem | frecuzrdglem 9413* | A helper lemma for the value of a recursive definition generator on upper integers. (Contributed by Jim Kingdon, 26-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝐴 ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ (ℤ≥‘𝐶) ∧ 𝑦 ∈ 𝑆)) → (𝑥𝐹𝑦) ∈ 𝑆) & ⊢ 𝑅 = frec((𝑥 ∈ (ℤ≥‘𝐶), 𝑦 ∈ 𝑆 ↦ 〈(𝑥 + 1), (𝑥𝐹𝑦)〉), 〈𝐶, 𝐴〉) & ⊢ (𝜑 → 𝐵 ∈ (ℤ≥‘𝐶)) ⇒ ⊢ (𝜑 → 〈𝐵, (2nd ‘(𝑅‘(◡𝐺‘𝐵)))〉 ∈ ran 𝑅) | ||
Theorem | frecuzrdgfn 9414* | The recursive definition generator on upper integers is a function. See comment in frec2uz0d 9401 for the description of 𝐺 as the mapping from ω to (ℤ≥‘𝐶). (Contributed by Jim Kingdon, 26-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝐴 ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ (ℤ≥‘𝐶) ∧ 𝑦 ∈ 𝑆)) → (𝑥𝐹𝑦) ∈ 𝑆) & ⊢ 𝑅 = frec((𝑥 ∈ (ℤ≥‘𝐶), 𝑦 ∈ 𝑆 ↦ 〈(𝑥 + 1), (𝑥𝐹𝑦)〉), 〈𝐶, 𝐴〉) & ⊢ (𝜑 → 𝑇 = ran 𝑅) ⇒ ⊢ (𝜑 → 𝑇 Fn (ℤ≥‘𝐶)) | ||
Theorem | frecuzrdgcl 9415* | Closure law for the recursive definition generator on upper integers. See comment in frec2uz0d 9401 for the description of 𝐺 as the mapping from ω to (ℤ≥‘𝐶). (Contributed by Jim Kingdon, 31-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝐴 ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ (ℤ≥‘𝐶) ∧ 𝑦 ∈ 𝑆)) → (𝑥𝐹𝑦) ∈ 𝑆) & ⊢ 𝑅 = frec((𝑥 ∈ (ℤ≥‘𝐶), 𝑦 ∈ 𝑆 ↦ 〈(𝑥 + 1), (𝑥𝐹𝑦)〉), 〈𝐶, 𝐴〉) & ⊢ (𝜑 → 𝑇 = ran 𝑅) ⇒ ⊢ ((𝜑 ∧ 𝐵 ∈ (ℤ≥‘𝐶)) → (𝑇‘𝐵) ∈ 𝑆) | ||
Theorem | frecuzrdg0 9416* | Initial value of a recursive definition generator on upper integers. See comment in frec2uz0d 9401 for the description of 𝐺 as the mapping from ω to (ℤ≥‘𝐶). (Contributed by Jim Kingdon, 27-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝐴 ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ (ℤ≥‘𝐶) ∧ 𝑦 ∈ 𝑆)) → (𝑥𝐹𝑦) ∈ 𝑆) & ⊢ 𝑅 = frec((𝑥 ∈ (ℤ≥‘𝐶), 𝑦 ∈ 𝑆 ↦ 〈(𝑥 + 1), (𝑥𝐹𝑦)〉), 〈𝐶, 𝐴〉) & ⊢ (𝜑 → 𝑇 = ran 𝑅) ⇒ ⊢ (𝜑 → (𝑇‘𝐶) = 𝐴) | ||
Theorem | frecuzrdgsuc 9417* | Successor value of a recursive definition generator on upper integers. See comment in frec2uz0d 9401 for the description of 𝐺 as the mapping from ω to (ℤ≥‘𝐶). (Contributed by Jim Kingdon, 28-May-2020.) |
⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 𝐶) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝐴 ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ (ℤ≥‘𝐶) ∧ 𝑦 ∈ 𝑆)) → (𝑥𝐹𝑦) ∈ 𝑆) & ⊢ 𝑅 = frec((𝑥 ∈ (ℤ≥‘𝐶), 𝑦 ∈ 𝑆 ↦ 〈(𝑥 + 1), (𝑥𝐹𝑦)〉), 〈𝐶, 𝐴〉) & ⊢ (𝜑 → 𝑇 = ran 𝑅) ⇒ ⊢ ((𝜑 ∧ 𝐵 ∈ (ℤ≥‘𝐶)) → (𝑇‘(𝐵 + 1)) = (𝐵𝐹(𝑇‘𝐵))) | ||
Theorem | uzenom 9418 | An upper integer set is denumerable. (Contributed by Mario Carneiro, 15-Oct-2015.) |
⊢ 𝑍 = (ℤ≥‘𝑀) ⇒ ⊢ (𝑀 ∈ ℤ → 𝑍 ≈ ω) | ||
Theorem | frecfzennn 9419 | The cardinality of a finite set of sequential integers. (See frec2uz0d 9401 for a description of the hypothesis.) (Contributed by Jim Kingdon, 18-May-2020.) |
⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) ⇒ ⊢ (𝑁 ∈ ℕ0 → (1...𝑁) ≈ (◡𝐺‘𝑁)) | ||
Theorem | frecfzen2 9420 | The cardinality of a finite set of sequential integers with arbitrary endpoints. (Contributed by Jim Kingdon, 18-May-2020.) |
⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) ⇒ ⊢ (𝑁 ∈ (ℤ≥‘𝑀) → (𝑀...𝑁) ≈ (◡𝐺‘((𝑁 + 1) − 𝑀))) | ||
Theorem | frechashgf1o 9421 | 𝐺 maps ω one-to-one onto ℕ0. (Contributed by Jim Kingdon, 19-May-2020.) |
⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) ⇒ ⊢ 𝐺:ω–1-1-onto→ℕ0 | ||
Theorem | fzfig 9422 | A finite interval of integers is finite. (Contributed by Jim Kingdon, 19-May-2020.) |
⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀...𝑁) ∈ Fin) | ||
Theorem | fzfigd 9423 | Deduction form of fzfig 9422. (Contributed by Jim Kingdon, 21-May-2020.) |
⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑁 ∈ ℤ) ⇒ ⊢ (𝜑 → (𝑀...𝑁) ∈ Fin) | ||
Theorem | fzofig 9424 | Half-open integer sets are finite. (Contributed by Jim Kingdon, 21-May-2020.) |
⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀..^𝑁) ∈ Fin) | ||
Theorem | nn0ennn 9425 | The nonnegative integers are equinumerous to the positive integers. (Contributed by NM, 19-Jul-2004.) |
⊢ ℕ0 ≈ ℕ | ||
Theorem | nnenom 9426 | The set of positive integers (as a subset of complex numbers) is equinumerous to omega (the set of finite ordinal numbers). (Contributed by NM, 31-Jul-2004.) (Revised by Mario Carneiro, 15-Sep-2013.) |
⊢ ℕ ≈ ω | ||
Theorem | nnct 9427 | ℕ is countable. (Contributed by Thierry Arnoux, 29-Dec-2016.) |
⊢ ℕ ≼ ω | ||
Theorem | uzsinds 9428* | Strong (or "total") induction principle over an upper set of integers. (Contributed by Scott Fenton, 16-May-2014.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝑁 → (𝜑 ↔ 𝜒)) & ⊢ (𝑥 ∈ (ℤ≥‘𝑀) → (∀𝑦 ∈ (𝑀...(𝑥 − 1))𝜓 → 𝜑)) ⇒ ⊢ (𝑁 ∈ (ℤ≥‘𝑀) → 𝜒) | ||
Theorem | nnsinds 9429* | Strong (or "total") induction principle over the naturals. (Contributed by Scott Fenton, 16-May-2014.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝑁 → (𝜑 ↔ 𝜒)) & ⊢ (𝑥 ∈ ℕ → (∀𝑦 ∈ (1...(𝑥 − 1))𝜓 → 𝜑)) ⇒ ⊢ (𝑁 ∈ ℕ → 𝜒) | ||
Theorem | nn0sinds 9430* | Strong (or "total") induction principle over the nonnegative integers. (Contributed by Scott Fenton, 16-May-2014.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝑁 → (𝜑 ↔ 𝜒)) & ⊢ (𝑥 ∈ ℕ0 → (∀𝑦 ∈ (0...(𝑥 − 1))𝜓 → 𝜑)) ⇒ ⊢ (𝑁 ∈ ℕ0 → 𝜒) | ||
Syntax | cseq 9431 | Extend class notation with recursive sequence builder. |
class seq𝑀( + , 𝐹, 𝑆) | ||
Definition | df-iseq 9432* |
Define a general-purpose operation that builds a recursive sequence
(i.e. a function on the positive integers ℕ or some other upper
integer set) whose value at an index is a function of its previous value
and the value of an input sequence at that index. This definition is
complicated, but fortunately it is not intended to be used directly.
Instead, the only purpose of this definition is to provide us with an
object that has the properties expressed by iseq1 9442 and iseqp1 9445.
Typically, those are the main theorems that would be used in practice.
The first operand in the parentheses is the operation that is applied to the previous value and the value of the input sequence (second operand). The operand to the left of the parenthesis is the integer to start from. For example, for the operation +, an input sequence 𝐹 with values 1, 1/2, 1/4, 1/8,... would be transformed into the output sequence seq1( + , 𝐹, ℚ) with values 1, 3/2, 7/4, 15/8,.., so that (seq1( + , 𝐹, ℚ)‘1) = 1, (seq1( + , 𝐹, ℚ)‘2) = 3/2, etc. In other words, seq𝑀( + , 𝐹, ℚ) transforms a sequence 𝐹 into an infinite series. Internally, the frec function generates as its values a set of ordered pairs starting at 〈𝑀, (𝐹‘𝑀)〉, with the first member of each pair incremented by one in each successive value. So, the range of frec is exactly the sequence we want, and we just extract the range and throw away the domain. (Contributed by Jim Kingdon, 29-May-2020.) |
⊢ seq𝑀( + , 𝐹, 𝑆) = ran frec((𝑥 ∈ (ℤ≥‘𝑀), 𝑦 ∈ 𝑆 ↦ 〈(𝑥 + 1), (𝑦 + (𝐹‘(𝑥 + 1)))〉), 〈𝑀, (𝐹‘𝑀)〉) | ||
Theorem | iseqex 9433 | Existence of the sequence builder operation. (Contributed by Jim Kingdon, 20-Aug-2021.) |
⊢ seq𝑀( + , 𝐹, 𝑆) ∈ V | ||
Theorem | iseqeq1 9434 | Equality theorem for the sequence builder operation. (Contributed by Jim Kingdon, 30-May-2020.) |
⊢ (𝑀 = 𝑁 → seq𝑀( + , 𝐹, 𝑆) = seq𝑁( + , 𝐹, 𝑆)) | ||
Theorem | iseqeq2 9435 | Equality theorem for the sequence builder operation. (Contributed by Jim Kingdon, 30-May-2020.) |
⊢ ( + = 𝑄 → seq𝑀( + , 𝐹, 𝑆) = seq𝑀(𝑄, 𝐹, 𝑆)) | ||
Theorem | iseqeq3 9436 | Equality theorem for the sequence builder operation. (Contributed by Jim Kingdon, 30-May-2020.) |
⊢ (𝐹 = 𝐺 → seq𝑀( + , 𝐹, 𝑆) = seq𝑀( + , 𝐺, 𝑆)) | ||
Theorem | iseqeq4 9437 | Equality theorem for the sequence builder operation. (Contributed by Jim Kingdon, 30-May-2020.) |
⊢ (𝑆 = 𝑇 → seq𝑀( + , 𝐹, 𝑆) = seq𝑀( + , 𝐹, 𝑇)) | ||
Theorem | nfiseq 9438 | Hypothesis builder for the sequence builder operation. (Contributed by Jim Kingdon, 30-May-2020.) |
⊢ Ⅎ𝑥𝑀 & ⊢ Ⅎ𝑥 + & ⊢ Ⅎ𝑥𝐹 & ⊢ Ⅎ𝑥𝑆 ⇒ ⊢ Ⅎ𝑥seq𝑀( + , 𝐹, 𝑆) | ||
Theorem | iseqovex 9439* | Closure of a function used in proving sequence builder theorems. This can be thought of as a lemma for the small number of sequence builder theorems which need it. (Contributed by Jim Kingdon, 31-May-2020.) |
⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ ((𝜑 ∧ (𝑥 ∈ (ℤ≥‘𝑀) ∧ 𝑦 ∈ 𝑆)) → (𝑥(𝑧 ∈ (ℤ≥‘𝑀), 𝑤 ∈ 𝑆 ↦ (𝑤 + (𝐹‘(𝑧 + 1))))𝑦) ∈ 𝑆) | ||
Theorem | iseqval 9440* | Value of the sequence builder function. (Contributed by Jim Kingdon, 30-May-2020.) |
⊢ 𝑅 = frec((𝑥 ∈ (ℤ≥‘𝑀), 𝑦 ∈ 𝑆 ↦ 〈(𝑥 + 1), (𝑥(𝑧 ∈ (ℤ≥‘𝑀), 𝑤 ∈ 𝑆 ↦ (𝑤 + (𝐹‘(𝑧 + 1))))𝑦)〉), 〈𝑀, (𝐹‘𝑀)〉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → seq𝑀( + , 𝐹, 𝑆) = ran 𝑅) | ||
Theorem | iseqfn 9441* | The sequence builder function is a function. (Contributed by Jim Kingdon, 30-May-2020.) |
⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → seq𝑀( + , 𝐹, 𝑆) Fn (ℤ≥‘𝑀)) | ||
Theorem | iseq1 9442* | Value of the sequence builder function at its initial value. (Contributed by Jim Kingdon, 31-May-2020.) |
⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝑀) = (𝐹‘𝑀)) | ||
Theorem | iseqcl 9443* | Closure properties of the recursive sequence builder. (Contributed by Jim Kingdon, 1-Jun-2020.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝑁) ∈ 𝑆) | ||
Theorem | iseqf 9444* | Range of the recursive sequence builder. (Contributed by Jim Kingdon, 23-Jul-2021.) |
⊢ 𝑍 = (ℤ≥‘𝑀) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝑍) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → seq𝑀( + , 𝐹, 𝑆):𝑍⟶𝑆) | ||
Theorem | iseqp1 9445* | Value of the sequence builder function at a successor. (Contributed by Jim Kingdon, 31-May-2020.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘(𝑁 + 1)) = ((seq𝑀( + , 𝐹, 𝑆)‘𝑁) + (𝐹‘(𝑁 + 1)))) | ||
Theorem | iseqss 9446* | Specifying a larger universe for seq. As long as 𝐹 and + are closed over 𝑆, then any set which contains 𝑆 can be used as the last argument to seq. This theorem does not allow 𝑇 to be a proper class, however. It also currently requires that + be closed over 𝑇 (as well as 𝑆). (Contributed by Jim Kingdon, 18-Aug-2021.) |
⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑇 ∈ 𝑉) & ⊢ (𝜑 → 𝑆 ⊆ 𝑇) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑇 ∧ 𝑦 ∈ 𝑇)) → (𝑥 + 𝑦) ∈ 𝑇) ⇒ ⊢ (𝜑 → seq𝑀( + , 𝐹, 𝑆) = seq𝑀( + , 𝐹, 𝑇)) | ||
Theorem | iseqm1 9447* | Value of the sequence builder function at a successor. (Contributed by Mario Carneiro, 24-Jun-2013.) |
⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘(𝑀 + 1))) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝑁) = ((seq𝑀( + , 𝐹, 𝑆)‘(𝑁 − 1)) + (𝐹‘𝑁))) | ||
Theorem | iseqfveq2 9448* | Equality of sequences. (Contributed by Jim Kingdon, 3-Jun-2020.) |
⊢ (𝜑 → 𝐾 ∈ (ℤ≥‘𝑀)) & ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝐾) = (𝐺‘𝐾)) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝐾)) → (𝐺‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝐾)) & ⊢ ((𝜑 ∧ 𝑘 ∈ ((𝐾 + 1)...𝑁)) → (𝐹‘𝑘) = (𝐺‘𝑘)) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝑁) = (seq𝐾( + , 𝐺, 𝑆)‘𝑁)) | ||
Theorem | iseqfeq2 9449* | Equality of sequences. (Contributed by Jim Kingdon, 3-Jun-2020.) |
⊢ (𝜑 → 𝐾 ∈ (ℤ≥‘𝑀)) & ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝐾) = (𝐺‘𝐾)) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝐾)) → (𝐺‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘(𝐾 + 1))) → (𝐹‘𝑘) = (𝐺‘𝑘)) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆) ↾ (ℤ≥‘𝐾)) = seq𝐾( + , 𝐺, 𝑆)) | ||
Theorem | iseqfveq 9450* | Equality of sequences. (Contributed by Jim Kingdon, 4-Jun-2020.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → (𝐹‘𝑘) = (𝐺‘𝑘)) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝑁) = (seq𝑀( + , 𝐺, 𝑆)‘𝑁)) | ||
Theorem | iseqfeq 9451* | Equality of sequences. (Contributed by Jim Kingdon, 15-Aug-2021.) |
⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) = (𝐺‘𝑘)) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → seq𝑀( + , 𝐹, 𝑆) = seq𝑀( + , 𝐺, 𝑆)) | ||
Theorem | iseqshft2 9452* | Shifting the index set of a sequence. (Contributed by Jim Kingdon, 15-Aug-2021.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ (𝜑 → 𝐾 ∈ ℤ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → (𝐹‘𝑘) = (𝐺‘(𝑘 + 𝐾))) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘(𝑀 + 𝐾))) → (𝐺‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝑁) = (seq(𝑀 + 𝐾)( + , 𝐺, 𝑆)‘(𝑁 + 𝐾))) | ||
Theorem | iserf 9453* | An infinite series of complex terms is a function from ℕ to ℂ. (Contributed by Jim Kingdon, 15-Aug-2021.) |
⊢ 𝑍 = (ℤ≥‘𝑀) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝑍) → (𝐹‘𝑘) ∈ ℂ) ⇒ ⊢ (𝜑 → seq𝑀( + , 𝐹, ℂ):𝑍⟶ℂ) | ||
Theorem | iserfre 9454* | An infinite series of real numbers is a function from ℕ to ℝ. (Contributed by NM, 18-Apr-2005.) (Revised by Mario Carneiro, 27-May-2014.) |
⊢ 𝑍 = (ℤ≥‘𝑀) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝑍) → (𝐹‘𝑘) ∈ ℝ) ⇒ ⊢ (𝜑 → seq𝑀( + , 𝐹, ℝ):𝑍⟶ℝ) | ||
Theorem | monoord 9455* | Ordering relation for a monotonic sequence, increasing case. (Contributed by NM, 13-Mar-2005.) (Revised by Mario Carneiro, 9-Feb-2014.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → (𝐹‘𝑘) ∈ ℝ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...(𝑁 − 1))) → (𝐹‘𝑘) ≤ (𝐹‘(𝑘 + 1))) ⇒ ⊢ (𝜑 → (𝐹‘𝑀) ≤ (𝐹‘𝑁)) | ||
Theorem | monoord2 9456* | Ordering relation for a monotonic sequence, decreasing case. (Contributed by Mario Carneiro, 18-Jul-2014.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → (𝐹‘𝑘) ∈ ℝ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...(𝑁 − 1))) → (𝐹‘(𝑘 + 1)) ≤ (𝐹‘𝑘)) ⇒ ⊢ (𝜑 → (𝐹‘𝑁) ≤ (𝐹‘𝑀)) | ||
Theorem | isermono 9457* | The partial sums in an infinite series of positive terms form a monotonic sequence. (Contributed by Jim Kingdon, 15-Aug-2021.) |
⊢ (𝜑 → 𝐾 ∈ (ℤ≥‘𝑀)) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝐾)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ ((𝜑 ∧ 𝑥 ∈ ((𝐾 + 1)...𝑁)) → 0 ≤ (𝐹‘𝑥)) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, ℝ)‘𝐾) ≤ (seq𝑀( + , 𝐹, ℝ)‘𝑁)) | ||
Theorem | iseqsplit 9458* | Split a sequence into two sequences. (Contributed by Jim Kingdon, 16-Aug-2021.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆 ∧ 𝑧 ∈ 𝑆)) → ((𝑥 + 𝑦) + 𝑧) = (𝑥 + (𝑦 + 𝑧))) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘(𝑀 + 1))) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝑀 ∈ (ℤ≥‘𝐾)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝐾)) → (𝐹‘𝑥) ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝐾( + , 𝐹, 𝑆)‘𝑁) = ((seq𝐾( + , 𝐹, 𝑆)‘𝑀) + (seq(𝑀 + 1)( + , 𝐹, 𝑆)‘𝑁))) | ||
Theorem | iseq1p 9459* | Removing the first term from a sequence. (Contributed by Jim Kingdon, 16-Aug-2021.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆 ∧ 𝑧 ∈ 𝑆)) → ((𝑥 + 𝑦) + 𝑧) = (𝑥 + (𝑦 + 𝑧))) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘(𝑀 + 1))) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝑁) = ((𝐹‘𝑀) + (seq(𝑀 + 1)( + , 𝐹, 𝑆)‘𝑁))) | ||
Theorem | iseqcaopr3 9460* | Lemma for iseqcaopr2 . (Contributed by Jim Kingdon, 16-Aug-2021.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥𝑄𝑦) ∈ 𝑆) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑘) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐻‘𝑘) = ((𝐹‘𝑘)𝑄(𝐺‘𝑘))) & ⊢ ((𝜑 ∧ 𝑛 ∈ (𝑀..^𝑁)) → (((seq𝑀( + , 𝐹, 𝑆)‘𝑛)𝑄(seq𝑀( + , 𝐺, 𝑆)‘𝑛)) + ((𝐹‘(𝑛 + 1))𝑄(𝐺‘(𝑛 + 1)))) = (((seq𝑀( + , 𝐹, 𝑆)‘𝑛) + (𝐹‘(𝑛 + 1)))𝑄((seq𝑀( + , 𝐺, 𝑆)‘𝑛) + (𝐺‘(𝑛 + 1))))) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐻, 𝑆)‘𝑁) = ((seq𝑀( + , 𝐹, 𝑆)‘𝑁)𝑄(seq𝑀( + , 𝐺, 𝑆)‘𝑁))) | ||
Theorem | iseqcaopr2 9461* | The sum of two infinite series (generalized to an arbitrary commutative and associative operation). (Contributed by Mario Carneiro, 30-May-2014.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥𝑄𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ ((𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆) ∧ (𝑧 ∈ 𝑆 ∧ 𝑤 ∈ 𝑆))) → ((𝑥𝑄𝑧) + (𝑦𝑄𝑤)) = ((𝑥 + 𝑦)𝑄(𝑧 + 𝑤))) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑘) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐻‘𝑘) = ((𝐹‘𝑘)𝑄(𝐺‘𝑘))) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐻, 𝑆)‘𝑁) = ((seq𝑀( + , 𝐹, 𝑆)‘𝑁)𝑄(seq𝑀( + , 𝐺, 𝑆)‘𝑁))) | ||
Theorem | iseqcaopr 9462* | The sum of two infinite series (generalized to an arbitrary commutative and associative operation). (Contributed by Jim Kingdon, 17-Aug-2021.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) = (𝑦 + 𝑥)) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆 ∧ 𝑧 ∈ 𝑆)) → ((𝑥 + 𝑦) + 𝑧) = (𝑥 + (𝑦 + 𝑧))) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑘) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐻‘𝑘) = ((𝐹‘𝑘) + (𝐺‘𝑘))) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐻, 𝑆)‘𝑁) = ((seq𝑀( + , 𝐹, 𝑆)‘𝑁) + (seq𝑀( + , 𝐺, 𝑆)‘𝑁))) | ||
Theorem | iseradd 9463* | The sum of two infinite series. (Contributed by NM, 17-Mar-2005.) (Revised by Mario Carneiro, 26-May-2014.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑘) ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐻‘𝑘) = ((𝐹‘𝑘) + (𝐺‘𝑘))) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐻, ℂ)‘𝑁) = ((seq𝑀( + , 𝐹, ℂ)‘𝑁) + (seq𝑀( + , 𝐺, ℂ)‘𝑁))) | ||
Theorem | isersub 9464* | The difference of two infinite series. (Contributed by NM, 17-Mar-2005.) (Revised by Mario Carneiro, 27-May-2014.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑘) ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐻‘𝑘) = ((𝐹‘𝑘) − (𝐺‘𝑘))) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐻, ℂ)‘𝑁) = ((seq𝑀( + , 𝐹, ℂ)‘𝑁) − (seq𝑀( + , 𝐺, ℂ)‘𝑁))) | ||
Theorem | iseqid3 9465* | A sequence that consists entirely of zeroes (or whatever the identity 𝑍 is for operation +) sums to zero. (Contributed by Jim Kingdon, 18-Aug-2021.) |
⊢ (𝜑 → (𝑍 + 𝑍) = 𝑍) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) = 𝑍) & ⊢ (𝜑 → 𝑍 ∈ 𝑉) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, {𝑍})‘𝑁) = 𝑍) | ||
Theorem | iseqid3s 9466* | A sequence that consists of zeroes up to 𝑁 sums to zero at 𝑁. In this case by "zero" we mean whatever the identity 𝑍 is for the operation +). (Contributed by Jim Kingdon, 18-Aug-2021.) |
⊢ (𝜑 → (𝑍 + 𝑍) = 𝑍) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝑀...𝑁)) → (𝐹‘𝑥) = 𝑍) & ⊢ (𝜑 → 𝑍 ∈ 𝑆) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝑁) = 𝑍) | ||
Theorem | iseqid 9467* | Discard the first few terms of a sequence that starts with all zeroes (or whatever the identity 𝑍 is for operation +). (Contributed by Mario Carneiro, 13-Jul-2013.) (Revised by Mario Carneiro, 27-May-2014.) |
⊢ ((𝜑 ∧ 𝑥 ∈ 𝑆) → (𝑍 + 𝑥) = 𝑥) & ⊢ (𝜑 → 𝑍 ∈ 𝑆) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ (𝜑 → (𝐹‘𝑁) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝑀...(𝑁 − 1))) → (𝐹‘𝑥) = 𝑍) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆) ↾ (ℤ≥‘𝑁)) = seq𝑁( + , 𝐹, 𝑆)) | ||
Theorem | iseqhomo 9468* | Apply a homomorphism to a sequence. (Contributed by Jim Kingdon, 21-Aug-2021.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝐻‘(𝑥 + 𝑦)) = ((𝐻‘𝑥)𝑄(𝐻‘𝑦))) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐻‘(𝐹‘𝑥)) = (𝐺‘𝑥)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥𝑄𝑦) ∈ 𝑆) ⇒ ⊢ (𝜑 → (𝐻‘(seq𝑀( + , 𝐹, 𝑆)‘𝑁)) = (seq𝑀(𝑄, 𝐺, 𝑆)‘𝑁)) | ||
Theorem | iseqz 9469* | If the operation + has an absorbing element 𝑍 (a.k.a. zero element), then any sequence containing a 𝑍 evaluates to 𝑍. (Contributed by Mario Carneiro, 27-May-2014.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝑆) → (𝑍 + 𝑥) = 𝑍) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝑆) → (𝑥 + 𝑍) = 𝑍) & ⊢ (𝜑 → 𝐾 ∈ (𝑀...𝑁)) & ⊢ (𝜑 → 𝑁 ∈ 𝑉) & ⊢ (𝜑 → (𝐹‘𝐾) = 𝑍) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝑁) = 𝑍) | ||
Theorem | iseqdistr 9470* | The distributive property for series. (Contributed by Jim Kingdon, 21-Aug-2021.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 + 𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝐶𝑇(𝑥 + 𝑦)) = ((𝐶𝑇𝑥) + (𝐶𝑇𝑦))) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑥) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) = (𝐶𝑇(𝐺‘𝑥))) & ⊢ (𝜑 → 𝑆 ∈ 𝑉) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥𝑇𝑦) ∈ 𝑆) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑥) ∈ 𝑆) & ⊢ (𝜑 → 𝐶 ∈ 𝑆) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, 𝑆)‘𝑁) = (𝐶𝑇(seq𝑀( + , 𝐺, 𝑆)‘𝑁))) | ||
Theorem | iser0 9471 | The value of the partial sums in a zero-valued infinite series. (Contributed by Jim Kingdon, 19-Aug-2021.) |
⊢ 𝑍 = (ℤ≥‘𝑀) ⇒ ⊢ (𝑁 ∈ 𝑍 → (seq𝑀( + , (𝑍 × {0}), ℂ)‘𝑁) = 0) | ||
Theorem | iser0f 9472 | A zero-valued infinite series is equal to the constant zero function. (Contributed by Jim Kingdon, 19-Aug-2021.) |
⊢ 𝑍 = (ℤ≥‘𝑀) ⇒ ⊢ (𝑀 ∈ ℤ → seq𝑀( + , (𝑍 × {0}), ℂ) = (𝑍 × {0})) | ||
Theorem | serige0 9473* | A finite sum of nonnegative terms is nonnegative. (Contributed by Jim Kingdon, 22-Aug-2021.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ ℝ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → 0 ≤ (𝐹‘𝑘)) ⇒ ⊢ (𝜑 → 0 ≤ (seq𝑀( + , 𝐹, ℂ)‘𝑁)) | ||
Theorem | serile 9474* | Comparison of partial sums of two infinite series of reals. (Contributed by Jim Kingdon, 22-Aug-2021.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ ℝ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑘) ∈ ℝ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ≤ (𝐺‘𝑘)) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹, ℂ)‘𝑁) ≤ (seq𝑀( + , 𝐺, ℂ)‘𝑁)) | ||
Syntax | cexp 9475 | Extend class notation to include exponentiation of a complex number to an integer power. |
class ↑ | ||
Definition | df-iexp 9476* | Define exponentiation to nonnegative integer powers. This definition is not meant to be used directly; instead, exp0 9480 and expp1 9483 provide the standard recursive definition. The up-arrow notation is used by Donald Knuth for iterated exponentiation (Science 194, 1235-1242, 1976) and is convenient for us since we don't have superscripts. 10-Jun-2005: The definition was extended to include zero exponents, so that 0↑0 = 1 per the convention of Definition 10-4.1 of [Gleason] p. 134. 4-Jun-2014: The definition was extended to include negative integer exponents. The case 𝑥 = 0, 𝑦 < 0 gives the value (1 / 0), so we will avoid this case in our theorems. (Contributed by Jim Kingdon, 7-Jun-2020.) |
⊢ ↑ = (𝑥 ∈ ℂ, 𝑦 ∈ ℤ ↦ if(𝑦 = 0, 1, if(0 < 𝑦, (seq1( · , (ℕ × {𝑥}), ℂ)‘𝑦), (1 / (seq1( · , (ℕ × {𝑥}), ℂ)‘-𝑦))))) | ||
Theorem | expivallem 9477 | Lemma for expival 9478. If we take a complex number apart from zero and raise it to a positive integer power, the result is apart from zero. (Contributed by Jim Kingdon, 7-Jun-2020.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝐴 # 0 ∧ 𝑁 ∈ ℕ) → (seq1( · , (ℕ × {𝐴}), ℂ)‘𝑁) # 0) | ||
Theorem | expival 9478 | Value of exponentiation to integer powers. (Contributed by Jim Kingdon, 7-Jun-2020.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝑁 ∈ ℤ ∧ (𝐴 # 0 ∨ 0 ≤ 𝑁)) → (𝐴↑𝑁) = if(𝑁 = 0, 1, if(0 < 𝑁, (seq1( · , (ℕ × {𝐴}), ℂ)‘𝑁), (1 / (seq1( · , (ℕ × {𝐴}), ℂ)‘-𝑁))))) | ||
Theorem | expinnval 9479 | Value of exponentiation to positive integer powers. (Contributed by Jim Kingdon, 8-Jun-2020.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝑁 ∈ ℕ) → (𝐴↑𝑁) = (seq1( · , (ℕ × {𝐴}), ℂ)‘𝑁)) | ||
Theorem | exp0 9480 | Value of a complex number raised to the 0th power. Note that under our definition, 0↑0 = 1, following the convention used by Gleason. Part of Definition 10-4.1 of [Gleason] p. 134. (Contributed by NM, 20-May-2004.) (Revised by Mario Carneiro, 4-Jun-2014.) |
⊢ (𝐴 ∈ ℂ → (𝐴↑0) = 1) | ||
Theorem | 0exp0e1 9481 | 0↑0 = 1 (common case). This is our convention. It follows the convention used by Gleason; see Part of Definition 10-4.1 of [Gleason] p. 134. (Contributed by David A. Wheeler, 8-Dec-2018.) |
⊢ (0↑0) = 1 | ||
Theorem | exp1 9482 | Value of a complex number raised to the first power. (Contributed by NM, 20-Oct-2004.) (Revised by Mario Carneiro, 2-Jul-2013.) |
⊢ (𝐴 ∈ ℂ → (𝐴↑1) = 𝐴) | ||
Theorem | expp1 9483 | Value of a complex number raised to a nonnegative integer power plus one. Part of Definition 10-4.1 of [Gleason] p. 134. (Contributed by NM, 20-May-2005.) (Revised by Mario Carneiro, 2-Jul-2013.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝑁 ∈ ℕ0) → (𝐴↑(𝑁 + 1)) = ((𝐴↑𝑁) · 𝐴)) | ||
Theorem | expnegap0 9484 | Value of a complex number raised to a negative integer power. (Contributed by Jim Kingdon, 8-Jun-2020.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝐴 # 0 ∧ 𝑁 ∈ ℕ0) → (𝐴↑-𝑁) = (1 / (𝐴↑𝑁))) | ||
Theorem | expineg2 9485 | Value of a complex number raised to a negative integer power. (Contributed by Jim Kingdon, 8-Jun-2020.) |
⊢ (((𝐴 ∈ ℂ ∧ 𝐴 # 0) ∧ (𝑁 ∈ ℂ ∧ -𝑁 ∈ ℕ0)) → (𝐴↑𝑁) = (1 / (𝐴↑-𝑁))) | ||
Theorem | expn1ap0 9486 | A number to the negative one power is the reciprocal. (Contributed by Jim Kingdon, 8-Jun-2020.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝐴 # 0) → (𝐴↑-1) = (1 / 𝐴)) | ||
Theorem | expcllem 9487* | Lemma for proving nonnegative integer exponentiation closure laws. (Contributed by NM, 14-Dec-2005.) |
⊢ 𝐹 ⊆ ℂ & ⊢ ((𝑥 ∈ 𝐹 ∧ 𝑦 ∈ 𝐹) → (𝑥 · 𝑦) ∈ 𝐹) & ⊢ 1 ∈ 𝐹 ⇒ ⊢ ((𝐴 ∈ 𝐹 ∧ 𝐵 ∈ ℕ0) → (𝐴↑𝐵) ∈ 𝐹) | ||
Theorem | expcl2lemap 9488* | Lemma for proving integer exponentiation closure laws. (Contributed by Jim Kingdon, 8-Jun-2020.) |
⊢ 𝐹 ⊆ ℂ & ⊢ ((𝑥 ∈ 𝐹 ∧ 𝑦 ∈ 𝐹) → (𝑥 · 𝑦) ∈ 𝐹) & ⊢ 1 ∈ 𝐹 & ⊢ ((𝑥 ∈ 𝐹 ∧ 𝑥 # 0) → (1 / 𝑥) ∈ 𝐹) ⇒ ⊢ ((𝐴 ∈ 𝐹 ∧ 𝐴 # 0 ∧ 𝐵 ∈ ℤ) → (𝐴↑𝐵) ∈ 𝐹) | ||
Theorem | nnexpcl 9489 | Closure of exponentiation of nonnegative integers. (Contributed by NM, 16-Dec-2005.) |
⊢ ((𝐴 ∈ ℕ ∧ 𝑁 ∈ ℕ0) → (𝐴↑𝑁) ∈ ℕ) | ||
Theorem | nn0expcl 9490 | Closure of exponentiation of nonnegative integers. (Contributed by NM, 14-Dec-2005.) |
⊢ ((𝐴 ∈ ℕ0 ∧ 𝑁 ∈ ℕ0) → (𝐴↑𝑁) ∈ ℕ0) | ||
Theorem | zexpcl 9491 | Closure of exponentiation of integers. (Contributed by NM, 16-Dec-2005.) |
⊢ ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ0) → (𝐴↑𝑁) ∈ ℤ) | ||
Theorem | qexpcl 9492 | Closure of exponentiation of rationals. (Contributed by NM, 16-Dec-2005.) |
⊢ ((𝐴 ∈ ℚ ∧ 𝑁 ∈ ℕ0) → (𝐴↑𝑁) ∈ ℚ) | ||
Theorem | reexpcl 9493 | Closure of exponentiation of reals. (Contributed by NM, 14-Dec-2005.) |
⊢ ((𝐴 ∈ ℝ ∧ 𝑁 ∈ ℕ0) → (𝐴↑𝑁) ∈ ℝ) | ||
Theorem | expcl 9494 | Closure law for nonnegative integer exponentiation. (Contributed by NM, 26-May-2005.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝑁 ∈ ℕ0) → (𝐴↑𝑁) ∈ ℂ) | ||
Theorem | rpexpcl 9495 | Closure law for exponentiation of positive reals. (Contributed by NM, 24-Feb-2008.) (Revised by Mario Carneiro, 9-Sep-2014.) |
⊢ ((𝐴 ∈ ℝ+ ∧ 𝑁 ∈ ℤ) → (𝐴↑𝑁) ∈ ℝ+) | ||
Theorem | reexpclzap 9496 | Closure of exponentiation of reals. (Contributed by Jim Kingdon, 9-Jun-2020.) |
⊢ ((𝐴 ∈ ℝ ∧ 𝐴 # 0 ∧ 𝑁 ∈ ℤ) → (𝐴↑𝑁) ∈ ℝ) | ||
Theorem | qexpclz 9497 | Closure of exponentiation of rational numbers. (Contributed by Mario Carneiro, 9-Sep-2014.) |
⊢ ((𝐴 ∈ ℚ ∧ 𝐴 ≠ 0 ∧ 𝑁 ∈ ℤ) → (𝐴↑𝑁) ∈ ℚ) | ||
Theorem | m1expcl2 9498 | Closure of exponentiation of negative one. (Contributed by Mario Carneiro, 18-Jun-2015.) |
⊢ (𝑁 ∈ ℤ → (-1↑𝑁) ∈ {-1, 1}) | ||
Theorem | m1expcl 9499 | Closure of exponentiation of negative one. (Contributed by Mario Carneiro, 18-Jun-2015.) |
⊢ (𝑁 ∈ ℤ → (-1↑𝑁) ∈ ℤ) | ||
Theorem | expclzaplem 9500* | Closure law for integer exponentiation. Lemma for expclzap 9501 and expap0i 9508. (Contributed by Jim Kingdon, 9-Jun-2020.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝐴 # 0 ∧ 𝑁 ∈ ℤ) → (𝐴↑𝑁) ∈ {𝑧 ∈ ℂ ∣ 𝑧 # 0}) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |