![]() |
Metamath
Proof Explorer Theorem List (p. 158 of 429) | < 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-27903) |
![]() (27904-29428) |
![]() (29429-42879) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | 4sqlem3 15701* | Lemma for 4sq 15715. Sufficient condition to be in 𝑆. (Contributed by Mario Carneiro, 14-Jul-2014.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} ⇒ ⊢ (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝐶 ∈ ℤ ∧ 𝐷 ∈ ℤ)) → (((𝐴↑2) + (𝐵↑2)) + ((𝐶↑2) + (𝐷↑2))) ∈ 𝑆) | ||
Theorem | 4sqlem4a 15702* | Lemma for 4sqlem4 15703. (Contributed by Mario Carneiro, 14-Jul-2014.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} ⇒ ⊢ ((𝐴 ∈ ℤ[i] ∧ 𝐵 ∈ ℤ[i]) → (((abs‘𝐴)↑2) + ((abs‘𝐵)↑2)) ∈ 𝑆) | ||
Theorem | 4sqlem4 15703* | Lemma for 4sq 15715. We can express the four-square property more compactly in terms of gaussian integers, because the norms of gaussian integers are exactly sums of two squares. (Contributed by Mario Carneiro, 14-Jul-2014.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} ⇒ ⊢ (𝐴 ∈ 𝑆 ↔ ∃𝑢 ∈ ℤ[i] ∃𝑣 ∈ ℤ[i] 𝐴 = (((abs‘𝑢)↑2) + ((abs‘𝑣)↑2))) | ||
Theorem | mul4sqlem 15704* | Lemma for mul4sq 15705: algebraic manipulations. The extra assumptions involving 𝑀 are for a part of 4sqlem17 15712 which needs to know not just that the product is a sum of squares, but also that it preserves divisibility by 𝑀. (Contributed by Mario Carneiro, 14-Jul-2014.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} & ⊢ (𝜑 → 𝐴 ∈ ℤ[i]) & ⊢ (𝜑 → 𝐵 ∈ ℤ[i]) & ⊢ (𝜑 → 𝐶 ∈ ℤ[i]) & ⊢ (𝜑 → 𝐷 ∈ ℤ[i]) & ⊢ 𝑋 = (((abs‘𝐴)↑2) + ((abs‘𝐵)↑2)) & ⊢ 𝑌 = (((abs‘𝐶)↑2) + ((abs‘𝐷)↑2)) & ⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → ((𝐴 − 𝐶) / 𝑀) ∈ ℤ[i]) & ⊢ (𝜑 → ((𝐵 − 𝐷) / 𝑀) ∈ ℤ[i]) & ⊢ (𝜑 → (𝑋 / 𝑀) ∈ ℕ0) ⇒ ⊢ (𝜑 → ((𝑋 / 𝑀) · (𝑌 / 𝑀)) ∈ 𝑆) | ||
Theorem | mul4sq 15705* | Euler's four-square identity: The product of two sums of four squares is also a sum of four squares. This is usually quoted as an explicit formula involving eight real variables; we save some time by working with complex numbers (gaussian integers) instead, so that we only have to work with four variables, and also hiding the actual formula for the product in the proof of mul4sqlem 15704. (For the curious, the explicit formula that is used is ( ∣ 𝑎 ∣ ↑2 + ∣ 𝑏 ∣ ↑2)( ∣ 𝑐 ∣ ↑2 + ∣ 𝑑 ∣ ↑2) = ∣ 𝑎∗ · 𝑐 + 𝑏 · 𝑑∗ ∣ ↑2 + ∣ 𝑎∗ · 𝑑 − 𝑏 · 𝑐∗ ∣ ↑2.) (Contributed by Mario Carneiro, 14-Jul-2014.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} ⇒ ⊢ ((𝐴 ∈ 𝑆 ∧ 𝐵 ∈ 𝑆) → (𝐴 · 𝐵) ∈ 𝑆) | ||
Theorem | 4sqlem11 15706* | Lemma for 4sq 15715. Use the pigeonhole principle to show that the sets {𝑚↑2 ∣ 𝑚 ∈ (0...𝑁)} and {-1 − 𝑛↑2 ∣ 𝑛 ∈ (0...𝑁)} have a common element, mod 𝑃. (Contributed by Mario Carneiro, 15-Jul-2014.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝑃 = ((2 · 𝑁) + 1)) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ 𝐴 = {𝑢 ∣ ∃𝑚 ∈ (0...𝑁)𝑢 = ((𝑚↑2) mod 𝑃)} & ⊢ 𝐹 = (𝑣 ∈ 𝐴 ↦ ((𝑃 − 1) − 𝑣)) ⇒ ⊢ (𝜑 → (𝐴 ∩ ran 𝐹) ≠ ∅) | ||
Theorem | 4sqlem12 15707* | Lemma for 4sq 15715. For any odd prime 𝑃, there is a 𝑘 < 𝑃 such that 𝑘𝑃 − 1 is a sum of two squares. (Contributed by Mario Carneiro, 15-Jul-2014.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝑃 = ((2 · 𝑁) + 1)) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ 𝐴 = {𝑢 ∣ ∃𝑚 ∈ (0...𝑁)𝑢 = ((𝑚↑2) mod 𝑃)} & ⊢ 𝐹 = (𝑣 ∈ 𝐴 ↦ ((𝑃 − 1) − 𝑣)) ⇒ ⊢ (𝜑 → ∃𝑘 ∈ (1...(𝑃 − 1))∃𝑢 ∈ ℤ[i] (((abs‘𝑢)↑2) + 1) = (𝑘 · 𝑃)) | ||
Theorem | 4sqlem13 15708* | Lemma for 4sq 15715. (Contributed by Mario Carneiro, 16-Jul-2014.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝑃 = ((2 · 𝑁) + 1)) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → (0...(2 · 𝑁)) ⊆ 𝑆) & ⊢ 𝑇 = {𝑖 ∈ ℕ ∣ (𝑖 · 𝑃) ∈ 𝑆} & ⊢ 𝑀 = inf(𝑇, ℝ, < ) ⇒ ⊢ (𝜑 → (𝑇 ≠ ∅ ∧ 𝑀 < 𝑃)) | ||
Theorem | 4sqlem14 15709* | Lemma for 4sq 15715. (Contributed by Mario Carneiro, 16-Jul-2014.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝑃 = ((2 · 𝑁) + 1)) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → (0...(2 · 𝑁)) ⊆ 𝑆) & ⊢ 𝑇 = {𝑖 ∈ ℕ ∣ (𝑖 · 𝑃) ∈ 𝑆} & ⊢ 𝑀 = inf(𝑇, ℝ, < ) & ⊢ (𝜑 → 𝑀 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐴 ∈ ℤ) & ⊢ (𝜑 → 𝐵 ∈ ℤ) & ⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ (𝜑 → 𝐷 ∈ ℤ) & ⊢ 𝐸 = (((𝐴 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐹 = (((𝐵 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐺 = (((𝐶 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐻 = (((𝐷 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝑅 = ((((𝐸↑2) + (𝐹↑2)) + ((𝐺↑2) + (𝐻↑2))) / 𝑀) & ⊢ (𝜑 → (𝑀 · 𝑃) = (((𝐴↑2) + (𝐵↑2)) + ((𝐶↑2) + (𝐷↑2)))) ⇒ ⊢ (𝜑 → 𝑅 ∈ ℕ0) | ||
Theorem | 4sqlem15 15710* | Lemma for 4sq 15715. (Contributed by Mario Carneiro, 16-Jul-2014.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝑃 = ((2 · 𝑁) + 1)) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → (0...(2 · 𝑁)) ⊆ 𝑆) & ⊢ 𝑇 = {𝑖 ∈ ℕ ∣ (𝑖 · 𝑃) ∈ 𝑆} & ⊢ 𝑀 = inf(𝑇, ℝ, < ) & ⊢ (𝜑 → 𝑀 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐴 ∈ ℤ) & ⊢ (𝜑 → 𝐵 ∈ ℤ) & ⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ (𝜑 → 𝐷 ∈ ℤ) & ⊢ 𝐸 = (((𝐴 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐹 = (((𝐵 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐺 = (((𝐶 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐻 = (((𝐷 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝑅 = ((((𝐸↑2) + (𝐹↑2)) + ((𝐺↑2) + (𝐻↑2))) / 𝑀) & ⊢ (𝜑 → (𝑀 · 𝑃) = (((𝐴↑2) + (𝐵↑2)) + ((𝐶↑2) + (𝐷↑2)))) ⇒ ⊢ ((𝜑 ∧ 𝑅 = 𝑀) → ((((((𝑀↑2) / 2) / 2) − (𝐸↑2)) = 0 ∧ ((((𝑀↑2) / 2) / 2) − (𝐹↑2)) = 0) ∧ (((((𝑀↑2) / 2) / 2) − (𝐺↑2)) = 0 ∧ ((((𝑀↑2) / 2) / 2) − (𝐻↑2)) = 0))) | ||
Theorem | 4sqlem16 15711* | Lemma for 4sq 15715. (Contributed by Mario Carneiro, 16-Jul-2014.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝑃 = ((2 · 𝑁) + 1)) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → (0...(2 · 𝑁)) ⊆ 𝑆) & ⊢ 𝑇 = {𝑖 ∈ ℕ ∣ (𝑖 · 𝑃) ∈ 𝑆} & ⊢ 𝑀 = inf(𝑇, ℝ, < ) & ⊢ (𝜑 → 𝑀 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐴 ∈ ℤ) & ⊢ (𝜑 → 𝐵 ∈ ℤ) & ⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ (𝜑 → 𝐷 ∈ ℤ) & ⊢ 𝐸 = (((𝐴 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐹 = (((𝐵 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐺 = (((𝐶 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐻 = (((𝐷 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝑅 = ((((𝐸↑2) + (𝐹↑2)) + ((𝐺↑2) + (𝐻↑2))) / 𝑀) & ⊢ (𝜑 → (𝑀 · 𝑃) = (((𝐴↑2) + (𝐵↑2)) + ((𝐶↑2) + (𝐷↑2)))) ⇒ ⊢ (𝜑 → (𝑅 ≤ 𝑀 ∧ ((𝑅 = 0 ∨ 𝑅 = 𝑀) → (𝑀↑2) ∥ (𝑀 · 𝑃)))) | ||
Theorem | 4sqlem17 15712* | Lemma for 4sq 15715. (Contributed by Mario Carneiro, 16-Jul-2014.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝑃 = ((2 · 𝑁) + 1)) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → (0...(2 · 𝑁)) ⊆ 𝑆) & ⊢ 𝑇 = {𝑖 ∈ ℕ ∣ (𝑖 · 𝑃) ∈ 𝑆} & ⊢ 𝑀 = inf(𝑇, ℝ, < ) & ⊢ (𝜑 → 𝑀 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐴 ∈ ℤ) & ⊢ (𝜑 → 𝐵 ∈ ℤ) & ⊢ (𝜑 → 𝐶 ∈ ℤ) & ⊢ (𝜑 → 𝐷 ∈ ℤ) & ⊢ 𝐸 = (((𝐴 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐹 = (((𝐵 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐺 = (((𝐶 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝐻 = (((𝐷 + (𝑀 / 2)) mod 𝑀) − (𝑀 / 2)) & ⊢ 𝑅 = ((((𝐸↑2) + (𝐹↑2)) + ((𝐺↑2) + (𝐻↑2))) / 𝑀) & ⊢ (𝜑 → (𝑀 · 𝑃) = (((𝐴↑2) + (𝐵↑2)) + ((𝐶↑2) + (𝐷↑2)))) ⇒ ⊢ ¬ 𝜑 | ||
Theorem | 4sqlem18 15713* | Lemma for 4sq 15715. Inductive step, odd prime case. (Contributed by Mario Carneiro, 16-Jul-2014.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝑃 = ((2 · 𝑁) + 1)) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → (0...(2 · 𝑁)) ⊆ 𝑆) & ⊢ 𝑇 = {𝑖 ∈ ℕ ∣ (𝑖 · 𝑃) ∈ 𝑆} & ⊢ 𝑀 = inf(𝑇, ℝ, < ) ⇒ ⊢ (𝜑 → 𝑃 ∈ 𝑆) | ||
Theorem | 4sqlem19 15714* | Lemma for 4sq 15715. The proof is by strong induction - we show that if all the integers less than 𝑘 are in 𝑆, then 𝑘 is as well. In this part of the proof we do the induction argument and dispense with all the cases except the odd prime case, which is sent to 4sqlem18 15713. If 𝑘 is 0, 1, 2, we show 𝑘 ∈ 𝑆 directly; otherwise if 𝑘 is composite, 𝑘 is the product of two numbers less than it (and hence in 𝑆 by assumption), so by mul4sq 15705 𝑘 ∈ 𝑆. (Contributed by Mario Carneiro, 14-Jul-2014.) (Revised by Mario Carneiro, 20-Jun-2015.) |
⊢ 𝑆 = {𝑛 ∣ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ ∃𝑧 ∈ ℤ ∃𝑤 ∈ ℤ 𝑛 = (((𝑥↑2) + (𝑦↑2)) + ((𝑧↑2) + (𝑤↑2)))} ⇒ ⊢ ℕ0 = 𝑆 | ||
Theorem | 4sq 15715* | Lagrange's four-square theorem, or Bachet's conjecture: every nonnegative integer is expressible as a sum of four squares. This is Metamath 100 proof #19. (Contributed by Mario Carneiro, 16-Jul-2014.) |
⊢ (𝐴 ∈ ℕ0 ↔ ∃𝑎 ∈ ℤ ∃𝑏 ∈ ℤ ∃𝑐 ∈ ℤ ∃𝑑 ∈ ℤ 𝐴 = (((𝑎↑2) + (𝑏↑2)) + ((𝑐↑2) + (𝑑↑2)))) | ||
Syntax | cvdwa 15716 | The arithmetic progression function. |
class AP | ||
Syntax | cvdwm 15717 | The monochromatic arithmetic progression predicate. |
class MonoAP | ||
Syntax | cvdwp 15718 | The polychromatic arithmetic progression predicate. |
class PolyAP | ||
Definition | df-vdwap 15719* | Define the arithmetic progression function, which takes as input a length 𝑘, a start point 𝑎, and a step 𝑑 and outputs the set of points in this progression. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ AP = (𝑘 ∈ ℕ0 ↦ (𝑎 ∈ ℕ, 𝑑 ∈ ℕ ↦ ran (𝑚 ∈ (0...(𝑘 − 1)) ↦ (𝑎 + (𝑚 · 𝑑))))) | ||
Definition | df-vdwmc 15720* | Define the "contains a monochromatic AP" predicate. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ MonoAP = {〈𝑘, 𝑓〉 ∣ ∃𝑐(ran (AP‘𝑘) ∩ 𝒫 (◡𝑓 “ {𝑐})) ≠ ∅} | ||
Definition | df-vdwpc 15721* | Define the "contains a polychromatic collection of APs" predicate. See vdwpc 15731 for more information. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ PolyAP = {〈〈𝑚, 𝑘〉, 𝑓〉 ∣ ∃𝑎 ∈ ℕ ∃𝑑 ∈ (ℕ ↑𝑚 (1...𝑚))(∀𝑖 ∈ (1...𝑚)((𝑎 + (𝑑‘𝑖))(AP‘𝑘)(𝑑‘𝑖)) ⊆ (◡𝑓 “ {(𝑓‘(𝑎 + (𝑑‘𝑖)))}) ∧ (#‘ran (𝑖 ∈ (1...𝑚) ↦ (𝑓‘(𝑎 + (𝑑‘𝑖))))) = 𝑚)} | ||
Theorem | vdwapfval 15722* | Define the arithmetic progression function, which takes as input a length 𝑘, a start point 𝑎, and a step 𝑑 and outputs the set of points in this progression. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ (𝐾 ∈ ℕ0 → (AP‘𝐾) = (𝑎 ∈ ℕ, 𝑑 ∈ ℕ ↦ ran (𝑚 ∈ (0...(𝐾 − 1)) ↦ (𝑎 + (𝑚 · 𝑑))))) | ||
Theorem | vdwapf 15723 | The arithmetic progression function is a function. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ (𝐾 ∈ ℕ0 → (AP‘𝐾):(ℕ × ℕ)⟶𝒫 ℕ) | ||
Theorem | vdwapval 15724* | Value of the arithmetic progression function. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ ((𝐾 ∈ ℕ0 ∧ 𝐴 ∈ ℕ ∧ 𝐷 ∈ ℕ) → (𝑋 ∈ (𝐴(AP‘𝐾)𝐷) ↔ ∃𝑚 ∈ (0...(𝐾 − 1))𝑋 = (𝐴 + (𝑚 · 𝐷)))) | ||
Theorem | vdwapun 15725 | Remove the first element of an arithmetic progression. (Contributed by Mario Carneiro, 11-Sep-2014.) |
⊢ ((𝐾 ∈ ℕ0 ∧ 𝐴 ∈ ℕ ∧ 𝐷 ∈ ℕ) → (𝐴(AP‘(𝐾 + 1))𝐷) = ({𝐴} ∪ ((𝐴 + 𝐷)(AP‘𝐾)𝐷))) | ||
Theorem | vdwapid1 15726 | The first element of an arithmetic progression. (Contributed by Mario Carneiro, 12-Sep-2014.) |
⊢ ((𝐾 ∈ ℕ ∧ 𝐴 ∈ ℕ ∧ 𝐷 ∈ ℕ) → 𝐴 ∈ (𝐴(AP‘𝐾)𝐷)) | ||
Theorem | vdwap0 15727 | Value of a length-1 arithmetic progression. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ ((𝐴 ∈ ℕ ∧ 𝐷 ∈ ℕ) → (𝐴(AP‘0)𝐷) = ∅) | ||
Theorem | vdwap1 15728 | Value of a length-1 arithmetic progression. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ ((𝐴 ∈ ℕ ∧ 𝐷 ∈ ℕ) → (𝐴(AP‘1)𝐷) = {𝐴}) | ||
Theorem | vdwmc 15729* | The predicate " The 〈𝑅, 𝑁〉-coloring 𝐹 contains a monochromatic AP of length 𝐾". (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ 𝑋 ∈ V & ⊢ (𝜑 → 𝐾 ∈ ℕ0) & ⊢ (𝜑 → 𝐹:𝑋⟶𝑅) ⇒ ⊢ (𝜑 → (𝐾 MonoAP 𝐹 ↔ ∃𝑐∃𝑎 ∈ ℕ ∃𝑑 ∈ ℕ (𝑎(AP‘𝐾)𝑑) ⊆ (◡𝐹 “ {𝑐}))) | ||
Theorem | vdwmc2 15730* | Expand out the definition of an arithmetic progression. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ 𝑋 ∈ V & ⊢ (𝜑 → 𝐾 ∈ ℕ0) & ⊢ (𝜑 → 𝐹:𝑋⟶𝑅) & ⊢ (𝜑 → 𝐴 ∈ 𝑋) ⇒ ⊢ (𝜑 → (𝐾 MonoAP 𝐹 ↔ ∃𝑐 ∈ 𝑅 ∃𝑎 ∈ ℕ ∃𝑑 ∈ ℕ ∀𝑚 ∈ (0...(𝐾 − 1))(𝑎 + (𝑚 · 𝑑)) ∈ (◡𝐹 “ {𝑐}))) | ||
Theorem | vdwpc 15731* | The predicate " The coloring 𝐹 contains a polychromatic 𝑀-tuple of AP's of length 𝐾". A polychromatic 𝑀-tuple of AP's is a set of AP's with the same base point but different step lengths, such that each individual AP is monochromatic, but the AP's all have mutually distinct colors. (The common basepoint is not required to have the same color as any of the AP's.) (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ 𝑋 ∈ V & ⊢ (𝜑 → 𝐾 ∈ ℕ0) & ⊢ (𝜑 → 𝐹:𝑋⟶𝑅) & ⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ 𝐽 = (1...𝑀) ⇒ ⊢ (𝜑 → (〈𝑀, 𝐾〉 PolyAP 𝐹 ↔ ∃𝑎 ∈ ℕ ∃𝑑 ∈ (ℕ ↑𝑚 𝐽)(∀𝑖 ∈ 𝐽 ((𝑎 + (𝑑‘𝑖))(AP‘𝐾)(𝑑‘𝑖)) ⊆ (◡𝐹 “ {(𝐹‘(𝑎 + (𝑑‘𝑖)))}) ∧ (#‘ran (𝑖 ∈ 𝐽 ↦ (𝐹‘(𝑎 + (𝑑‘𝑖))))) = 𝑀))) | ||
Theorem | vdwlem1 15732* | Lemma for vdw 15745. (Contributed by Mario Carneiro, 12-Sep-2014.) |
⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐾 ∈ ℕ) & ⊢ (𝜑 → 𝑊 ∈ ℕ) & ⊢ (𝜑 → 𝐹:(1...𝑊)⟶𝑅) & ⊢ (𝜑 → 𝐴 ∈ ℕ) & ⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝐷:(1...𝑀)⟶ℕ) & ⊢ (𝜑 → ∀𝑖 ∈ (1...𝑀)((𝐴 + (𝐷‘𝑖))(AP‘𝐾)(𝐷‘𝑖)) ⊆ (◡𝐹 “ {(𝐹‘(𝐴 + (𝐷‘𝑖)))})) & ⊢ (𝜑 → 𝐼 ∈ (1...𝑀)) & ⊢ (𝜑 → (𝐹‘𝐴) = (𝐹‘(𝐴 + (𝐷‘𝐼)))) ⇒ ⊢ (𝜑 → (𝐾 + 1) MonoAP 𝐹) | ||
Theorem | vdwlem2 15733* | Lemma for vdw 15745. (Contributed by Mario Carneiro, 12-Sep-2014.) |
⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐾 ∈ ℕ0) & ⊢ (𝜑 → 𝑊 ∈ ℕ) & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝐹:(1...𝑀)⟶𝑅) & ⊢ (𝜑 → 𝑀 ∈ (ℤ≥‘(𝑊 + 𝑁))) & ⊢ 𝐺 = (𝑥 ∈ (1...𝑊) ↦ (𝐹‘(𝑥 + 𝑁))) ⇒ ⊢ (𝜑 → (𝐾 MonoAP 𝐺 → 𝐾 MonoAP 𝐹)) | ||
Theorem | vdwlem3 15734 | Lemma for vdw 15745. (Contributed by Mario Carneiro, 13-Sep-2014.) |
⊢ (𝜑 → 𝑉 ∈ ℕ) & ⊢ (𝜑 → 𝑊 ∈ ℕ) & ⊢ (𝜑 → 𝐴 ∈ (1...𝑉)) & ⊢ (𝜑 → 𝐵 ∈ (1...𝑊)) ⇒ ⊢ (𝜑 → (𝐵 + (𝑊 · ((𝐴 − 1) + 𝑉))) ∈ (1...(𝑊 · (2 · 𝑉)))) | ||
Theorem | vdwlem4 15735* | Lemma for vdw 15745. (Contributed by Mario Carneiro, 12-Sep-2014.) |
⊢ (𝜑 → 𝑉 ∈ ℕ) & ⊢ (𝜑 → 𝑊 ∈ ℕ) & ⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐻:(1...(𝑊 · (2 · 𝑉)))⟶𝑅) & ⊢ 𝐹 = (𝑥 ∈ (1...𝑉) ↦ (𝑦 ∈ (1...𝑊) ↦ (𝐻‘(𝑦 + (𝑊 · ((𝑥 − 1) + 𝑉)))))) ⇒ ⊢ (𝜑 → 𝐹:(1...𝑉)⟶(𝑅 ↑𝑚 (1...𝑊))) | ||
Theorem | vdwlem5 15736* | Lemma for vdw 15745. (Contributed by Mario Carneiro, 12-Sep-2014.) |
⊢ (𝜑 → 𝑉 ∈ ℕ) & ⊢ (𝜑 → 𝑊 ∈ ℕ) & ⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐻:(1...(𝑊 · (2 · 𝑉)))⟶𝑅) & ⊢ 𝐹 = (𝑥 ∈ (1...𝑉) ↦ (𝑦 ∈ (1...𝑊) ↦ (𝐻‘(𝑦 + (𝑊 · ((𝑥 − 1) + 𝑉)))))) & ⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝐺:(1...𝑊)⟶𝑅) & ⊢ (𝜑 → 𝐾 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐴 ∈ ℕ) & ⊢ (𝜑 → 𝐷 ∈ ℕ) & ⊢ (𝜑 → (𝐴(AP‘𝐾)𝐷) ⊆ (◡𝐹 “ {𝐺})) & ⊢ (𝜑 → 𝐵 ∈ ℕ) & ⊢ (𝜑 → 𝐸:(1...𝑀)⟶ℕ) & ⊢ (𝜑 → ∀𝑖 ∈ (1...𝑀)((𝐵 + (𝐸‘𝑖))(AP‘𝐾)(𝐸‘𝑖)) ⊆ (◡𝐺 “ {(𝐺‘(𝐵 + (𝐸‘𝑖)))})) & ⊢ 𝐽 = (𝑖 ∈ (1...𝑀) ↦ (𝐺‘(𝐵 + (𝐸‘𝑖)))) & ⊢ (𝜑 → (#‘ran 𝐽) = 𝑀) & ⊢ 𝑇 = (𝐵 + (𝑊 · ((𝐴 + (𝑉 − 𝐷)) − 1))) & ⊢ 𝑃 = (𝑗 ∈ (1...(𝑀 + 1)) ↦ (if(𝑗 = (𝑀 + 1), 0, (𝐸‘𝑗)) + (𝑊 · 𝐷))) ⇒ ⊢ (𝜑 → 𝑇 ∈ ℕ) | ||
Theorem | vdwlem6 15737* | Lemma for vdw 15745. (Contributed by Mario Carneiro, 13-Sep-2014.) |
⊢ (𝜑 → 𝑉 ∈ ℕ) & ⊢ (𝜑 → 𝑊 ∈ ℕ) & ⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐻:(1...(𝑊 · (2 · 𝑉)))⟶𝑅) & ⊢ 𝐹 = (𝑥 ∈ (1...𝑉) ↦ (𝑦 ∈ (1...𝑊) ↦ (𝐻‘(𝑦 + (𝑊 · ((𝑥 − 1) + 𝑉)))))) & ⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝐺:(1...𝑊)⟶𝑅) & ⊢ (𝜑 → 𝐾 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐴 ∈ ℕ) & ⊢ (𝜑 → 𝐷 ∈ ℕ) & ⊢ (𝜑 → (𝐴(AP‘𝐾)𝐷) ⊆ (◡𝐹 “ {𝐺})) & ⊢ (𝜑 → 𝐵 ∈ ℕ) & ⊢ (𝜑 → 𝐸:(1...𝑀)⟶ℕ) & ⊢ (𝜑 → ∀𝑖 ∈ (1...𝑀)((𝐵 + (𝐸‘𝑖))(AP‘𝐾)(𝐸‘𝑖)) ⊆ (◡𝐺 “ {(𝐺‘(𝐵 + (𝐸‘𝑖)))})) & ⊢ 𝐽 = (𝑖 ∈ (1...𝑀) ↦ (𝐺‘(𝐵 + (𝐸‘𝑖)))) & ⊢ (𝜑 → (#‘ran 𝐽) = 𝑀) & ⊢ 𝑇 = (𝐵 + (𝑊 · ((𝐴 + (𝑉 − 𝐷)) − 1))) & ⊢ 𝑃 = (𝑗 ∈ (1...(𝑀 + 1)) ↦ (if(𝑗 = (𝑀 + 1), 0, (𝐸‘𝑗)) + (𝑊 · 𝐷))) ⇒ ⊢ (𝜑 → (〈(𝑀 + 1), 𝐾〉 PolyAP 𝐻 ∨ (𝐾 + 1) MonoAP 𝐺)) | ||
Theorem | vdwlem7 15738* | Lemma for vdw 15745. (Contributed by Mario Carneiro, 12-Sep-2014.) |
⊢ (𝜑 → 𝑉 ∈ ℕ) & ⊢ (𝜑 → 𝑊 ∈ ℕ) & ⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐻:(1...(𝑊 · (2 · 𝑉)))⟶𝑅) & ⊢ 𝐹 = (𝑥 ∈ (1...𝑉) ↦ (𝑦 ∈ (1...𝑊) ↦ (𝐻‘(𝑦 + (𝑊 · ((𝑥 − 1) + 𝑉)))))) & ⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝐺:(1...𝑊)⟶𝑅) & ⊢ (𝜑 → 𝐾 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐴 ∈ ℕ) & ⊢ (𝜑 → 𝐷 ∈ ℕ) & ⊢ (𝜑 → (𝐴(AP‘𝐾)𝐷) ⊆ (◡𝐹 “ {𝐺})) ⇒ ⊢ (𝜑 → (〈𝑀, 𝐾〉 PolyAP 𝐺 → (〈(𝑀 + 1), 𝐾〉 PolyAP 𝐻 ∨ (𝐾 + 1) MonoAP 𝐺))) | ||
Theorem | vdwlem8 15739* | Lemma for vdw 15745. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐾 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝑊 ∈ ℕ) & ⊢ (𝜑 → 𝐹:(1...(2 · 𝑊))⟶𝑅) & ⊢ 𝐶 ∈ V & ⊢ (𝜑 → 𝐴 ∈ ℕ) & ⊢ (𝜑 → 𝐷 ∈ ℕ) & ⊢ (𝜑 → (𝐴(AP‘𝐾)𝐷) ⊆ (◡𝐺 “ {𝐶})) & ⊢ 𝐺 = (𝑥 ∈ (1...𝑊) ↦ (𝐹‘(𝑥 + 𝑊))) ⇒ ⊢ (𝜑 → 〈1, 𝐾〉 PolyAP 𝐹) | ||
Theorem | vdwlem9 15740* | Lemma for vdw 15745. (Contributed by Mario Carneiro, 12-Sep-2014.) |
⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐾 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → ∀𝑠 ∈ Fin ∃𝑛 ∈ ℕ ∀𝑓 ∈ (𝑠 ↑𝑚 (1...𝑛))𝐾 MonoAP 𝑓) & ⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝑊 ∈ ℕ) & ⊢ (𝜑 → ∀𝑔 ∈ (𝑅 ↑𝑚 (1...𝑊))(〈𝑀, 𝐾〉 PolyAP 𝑔 ∨ (𝐾 + 1) MonoAP 𝑔)) & ⊢ (𝜑 → 𝑉 ∈ ℕ) & ⊢ (𝜑 → ∀𝑓 ∈ ((𝑅 ↑𝑚 (1...𝑊)) ↑𝑚 (1...𝑉))𝐾 MonoAP 𝑓) & ⊢ (𝜑 → 𝐻:(1...(𝑊 · (2 · 𝑉)))⟶𝑅) & ⊢ 𝐹 = (𝑥 ∈ (1...𝑉) ↦ (𝑦 ∈ (1...𝑊) ↦ (𝐻‘(𝑦 + (𝑊 · ((𝑥 − 1) + 𝑉)))))) ⇒ ⊢ (𝜑 → (〈(𝑀 + 1), 𝐾〉 PolyAP 𝐻 ∨ (𝐾 + 1) MonoAP 𝐻)) | ||
Theorem | vdwlem10 15741* | Lemma for vdw 15745. Set up secondary induction on 𝑀. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐾 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → ∀𝑠 ∈ Fin ∃𝑛 ∈ ℕ ∀𝑓 ∈ (𝑠 ↑𝑚 (1...𝑛))𝐾 MonoAP 𝑓) & ⊢ (𝜑 → 𝑀 ∈ ℕ) ⇒ ⊢ (𝜑 → ∃𝑛 ∈ ℕ ∀𝑓 ∈ (𝑅 ↑𝑚 (1...𝑛))(〈𝑀, 𝐾〉 PolyAP 𝑓 ∨ (𝐾 + 1) MonoAP 𝑓)) | ||
Theorem | vdwlem11 15742* | Lemma for vdw 15745. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐾 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → ∀𝑠 ∈ Fin ∃𝑛 ∈ ℕ ∀𝑓 ∈ (𝑠 ↑𝑚 (1...𝑛))𝐾 MonoAP 𝑓) ⇒ ⊢ (𝜑 → ∃𝑛 ∈ ℕ ∀𝑓 ∈ (𝑅 ↑𝑚 (1...𝑛))(𝐾 + 1) MonoAP 𝑓) | ||
Theorem | vdwlem12 15743 | Lemma for vdw 15745. 𝐾 = 2 base case of induction. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐹:(1...((#‘𝑅) + 1))⟶𝑅) & ⊢ (𝜑 → ¬ 2 MonoAP 𝐹) ⇒ ⊢ ¬ 𝜑 | ||
Theorem | vdwlem13 15744* | Lemma for vdw 15745. Main induction on 𝐾; 𝐾 = 0, 𝐾 = 1 base cases. (Contributed by Mario Carneiro, 18-Aug-2014.) |
⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐾 ∈ ℕ0) ⇒ ⊢ (𝜑 → ∃𝑛 ∈ ℕ ∀𝑓 ∈ (𝑅 ↑𝑚 (1...𝑛))𝐾 MonoAP 𝑓) | ||
Theorem | vdw 15745* | Van der Waerden's theorem. For any finite coloring 𝑅 and integer 𝐾, there is an 𝑁 such that every coloring function from 1...𝑁 to 𝑅 contains a monochromatic arithmetic progression (which written out in full means that there is a color 𝑐 and base, increment values 𝑎, 𝑑 such that all the numbers 𝑎, 𝑎 + 𝑑, ..., 𝑎 + (𝑘 − 1)𝑑 lie in the preimage of {𝑐}, i.e. they are all in 1...𝑁 and 𝑓 evaluated at each one yields 𝑐). (Contributed by Mario Carneiro, 13-Sep-2014.) |
⊢ ((𝑅 ∈ Fin ∧ 𝐾 ∈ ℕ0) → ∃𝑛 ∈ ℕ ∀𝑓 ∈ (𝑅 ↑𝑚 (1...𝑛))∃𝑐 ∈ 𝑅 ∃𝑎 ∈ ℕ ∃𝑑 ∈ ℕ ∀𝑚 ∈ (0...(𝐾 − 1))(𝑎 + (𝑚 · 𝑑)) ∈ (◡𝑓 “ {𝑐})) | ||
Theorem | vdwnnlem1 15746* | Corollary of vdw 15745, and lemma for vdwnn 15749. If 𝐹 is a coloring of the integers, then there are arbitrarily long monochromatic APs in 𝐹. (Contributed by Mario Carneiro, 13-Sep-2014.) |
⊢ ((𝑅 ∈ Fin ∧ 𝐹:ℕ⟶𝑅 ∧ 𝐾 ∈ ℕ0) → ∃𝑐 ∈ 𝑅 ∃𝑎 ∈ ℕ ∃𝑑 ∈ ℕ ∀𝑚 ∈ (0...(𝐾 − 1))(𝑎 + (𝑚 · 𝑑)) ∈ (◡𝐹 “ {𝑐})) | ||
Theorem | vdwnnlem2 15747* | Lemma for vdwnn 15749. The set of all "bad" 𝑘 for the theorem is upwards-closed, because a long AP implies a short AP. (Contributed by Mario Carneiro, 13-Sep-2014.) |
⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐹:ℕ⟶𝑅) & ⊢ 𝑆 = {𝑘 ∈ ℕ ∣ ¬ ∃𝑎 ∈ ℕ ∃𝑑 ∈ ℕ ∀𝑚 ∈ (0...(𝑘 − 1))(𝑎 + (𝑚 · 𝑑)) ∈ (◡𝐹 “ {𝑐})} ⇒ ⊢ ((𝜑 ∧ 𝐵 ∈ (ℤ≥‘𝐴)) → (𝐴 ∈ 𝑆 → 𝐵 ∈ 𝑆)) | ||
Theorem | vdwnnlem3 15748* | Lemma for vdwnn 15749. (Contributed by Mario Carneiro, 13-Sep-2014.) (Proof shortened by AV, 27-Sep-2020.) |
⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐹:ℕ⟶𝑅) & ⊢ 𝑆 = {𝑘 ∈ ℕ ∣ ¬ ∃𝑎 ∈ ℕ ∃𝑑 ∈ ℕ ∀𝑚 ∈ (0...(𝑘 − 1))(𝑎 + (𝑚 · 𝑑)) ∈ (◡𝐹 “ {𝑐})} & ⊢ (𝜑 → ∀𝑐 ∈ 𝑅 𝑆 ≠ ∅) ⇒ ⊢ ¬ 𝜑 | ||
Theorem | vdwnn 15749* | Van der Waerden's theorem, infinitary version. For any finite coloring 𝐹 of the positive integers, there is a color 𝑐 that contains arbitrarily long arithmetic progressions. (Contributed by Mario Carneiro, 13-Sep-2014.) |
⊢ ((𝑅 ∈ Fin ∧ 𝐹:ℕ⟶𝑅) → ∃𝑐 ∈ 𝑅 ∀𝑘 ∈ ℕ ∃𝑎 ∈ ℕ ∃𝑑 ∈ ℕ ∀𝑚 ∈ (0...(𝑘 − 1))(𝑎 + (𝑚 · 𝑑)) ∈ (◡𝐹 “ {𝑐})) | ||
Syntax | cram 15750 | Extend class notation with the Ramsey number function. |
class Ramsey | ||
Theorem | ramtlecl 15751* | The set 𝑇 of numbers with the Ramsey number property is upward-closed. (Contributed by Mario Carneiro, 21-Apr-2015.) |
⊢ 𝑇 = {𝑛 ∈ ℕ0 ∣ ∀𝑠(𝑛 ≤ (#‘𝑠) → 𝜑)} ⇒ ⊢ (𝑀 ∈ 𝑇 → (ℤ≥‘𝑀) ⊆ 𝑇) | ||
Definition | df-ram 15752* | Define the Ramsey number function. The input is a number 𝑚 for the size of the edges of the hypergraph, and a tuple 𝑟 from the finite color set to lower bounds for each color. The Ramsey number (𝑀 Ramsey 𝑅) is the smallest number such that for any set 𝑆 with (𝑀 Ramsey 𝑅) ≤ #𝑆 and any coloring 𝐹 of the set of 𝑀-element subsets of 𝑆 (with color set dom 𝑅), there is a color 𝑐 ∈ dom 𝑅 and a subset 𝑥 ⊆ 𝑆 such that 𝑅(𝑐) ≤ #𝑥 and all the hyperedges of 𝑥 (that is, subsets of 𝑥 of size 𝑀) have color 𝑐. (Contributed by Mario Carneiro, 20-Apr-2015.) (Revised by AV, 14-Sep-2020.) |
⊢ Ramsey = (𝑚 ∈ ℕ0, 𝑟 ∈ V ↦ inf({𝑛 ∈ ℕ0 ∣ ∀𝑠(𝑛 ≤ (#‘𝑠) → ∀𝑓 ∈ (dom 𝑟 ↑𝑚 {𝑦 ∈ 𝒫 𝑠 ∣ (#‘𝑦) = 𝑚})∃𝑐 ∈ dom 𝑟∃𝑥 ∈ 𝒫 𝑠((𝑟‘𝑐) ≤ (#‘𝑥) ∧ ∀𝑦 ∈ 𝒫 𝑥((#‘𝑦) = 𝑚 → (𝑓‘𝑦) = 𝑐)))}, ℝ*, < )) | ||
Theorem | hashbcval 15753* | Value of the "binomial set", the set of all 𝑁-element subsets of 𝐴. (Contributed by Mario Carneiro, 20-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) ⇒ ⊢ ((𝐴 ∈ 𝑉 ∧ 𝑁 ∈ ℕ0) → (𝐴𝐶𝑁) = {𝑥 ∈ 𝒫 𝐴 ∣ (#‘𝑥) = 𝑁}) | ||
Theorem | hashbccl 15754* | The binomial set is a finite set. (Contributed by Mario Carneiro, 20-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) ⇒ ⊢ ((𝐴 ∈ Fin ∧ 𝑁 ∈ ℕ0) → (𝐴𝐶𝑁) ∈ Fin) | ||
Theorem | hashbcss 15755* | Subset relation for the binomial set. (Contributed by Mario Carneiro, 20-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) ⇒ ⊢ ((𝐴 ∈ 𝑉 ∧ 𝐵 ⊆ 𝐴 ∧ 𝑁 ∈ ℕ0) → (𝐵𝐶𝑁) ⊆ (𝐴𝐶𝑁)) | ||
Theorem | hashbc0 15756* | The set of subsets of size zero is the singleton of the empty set. (Contributed by Mario Carneiro, 22-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) ⇒ ⊢ (𝐴 ∈ 𝑉 → (𝐴𝐶0) = {∅}) | ||
Theorem | hashbc2 15757* | The size of the binomial set is the binomial coefficient. (Contributed by Mario Carneiro, 20-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) ⇒ ⊢ ((𝐴 ∈ Fin ∧ 𝑁 ∈ ℕ0) → (#‘(𝐴𝐶𝑁)) = ((#‘𝐴)C𝑁)) | ||
Theorem | 0hashbc 15758* | There are no subsets of the empty set with size greater than zero. (Contributed by Mario Carneiro, 22-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) ⇒ ⊢ (𝑁 ∈ ℕ → (∅𝐶𝑁) = ∅) | ||
Theorem | ramval 15759* | The value of the Ramsey number function. (Contributed by Mario Carneiro, 21-Apr-2015.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ 𝑇 = {𝑛 ∈ ℕ0 ∣ ∀𝑠(𝑛 ≤ (#‘𝑠) → ∀𝑓 ∈ (𝑅 ↑𝑚 (𝑠𝐶𝑀))∃𝑐 ∈ 𝑅 ∃𝑥 ∈ 𝒫 𝑠((𝐹‘𝑐) ≤ (#‘𝑥) ∧ (𝑥𝐶𝑀) ⊆ (◡𝑓 “ {𝑐})))} ⇒ ⊢ ((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ 𝑉 ∧ 𝐹:𝑅⟶ℕ0) → (𝑀 Ramsey 𝐹) = inf(𝑇, ℝ*, < )) | ||
Theorem | ramcl2lem 15760* | Lemma for extended real closure of the Ramsey number function. (Contributed by Mario Carneiro, 20-Apr-2015.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ 𝑇 = {𝑛 ∈ ℕ0 ∣ ∀𝑠(𝑛 ≤ (#‘𝑠) → ∀𝑓 ∈ (𝑅 ↑𝑚 (𝑠𝐶𝑀))∃𝑐 ∈ 𝑅 ∃𝑥 ∈ 𝒫 𝑠((𝐹‘𝑐) ≤ (#‘𝑥) ∧ (𝑥𝐶𝑀) ⊆ (◡𝑓 “ {𝑐})))} ⇒ ⊢ ((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ 𝑉 ∧ 𝐹:𝑅⟶ℕ0) → (𝑀 Ramsey 𝐹) = if(𝑇 = ∅, +∞, inf(𝑇, ℝ, < ))) | ||
Theorem | ramtcl 15761* | The Ramsey number has the Ramsey number property if any number does. (Contributed by Mario Carneiro, 20-Apr-2015.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ 𝑇 = {𝑛 ∈ ℕ0 ∣ ∀𝑠(𝑛 ≤ (#‘𝑠) → ∀𝑓 ∈ (𝑅 ↑𝑚 (𝑠𝐶𝑀))∃𝑐 ∈ 𝑅 ∃𝑥 ∈ 𝒫 𝑠((𝐹‘𝑐) ≤ (#‘𝑥) ∧ (𝑥𝐶𝑀) ⊆ (◡𝑓 “ {𝑐})))} ⇒ ⊢ ((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ 𝑉 ∧ 𝐹:𝑅⟶ℕ0) → ((𝑀 Ramsey 𝐹) ∈ 𝑇 ↔ 𝑇 ≠ ∅)) | ||
Theorem | ramtcl2 15762* | The Ramsey number is an integer iff there is a number with the Ramsey number property. (Contributed by Mario Carneiro, 20-Apr-2015.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ 𝑇 = {𝑛 ∈ ℕ0 ∣ ∀𝑠(𝑛 ≤ (#‘𝑠) → ∀𝑓 ∈ (𝑅 ↑𝑚 (𝑠𝐶𝑀))∃𝑐 ∈ 𝑅 ∃𝑥 ∈ 𝒫 𝑠((𝐹‘𝑐) ≤ (#‘𝑥) ∧ (𝑥𝐶𝑀) ⊆ (◡𝑓 “ {𝑐})))} ⇒ ⊢ ((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ 𝑉 ∧ 𝐹:𝑅⟶ℕ0) → ((𝑀 Ramsey 𝐹) ∈ ℕ0 ↔ 𝑇 ≠ ∅)) | ||
Theorem | ramtub 15763* | The Ramsey number is a lower bound on the set of all numbers with the Ramsey number property. (Contributed by Mario Carneiro, 20-Apr-2015.) (Revised by AV, 14-Sep-2020.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ 𝑇 = {𝑛 ∈ ℕ0 ∣ ∀𝑠(𝑛 ≤ (#‘𝑠) → ∀𝑓 ∈ (𝑅 ↑𝑚 (𝑠𝐶𝑀))∃𝑐 ∈ 𝑅 ∃𝑥 ∈ 𝒫 𝑠((𝐹‘𝑐) ≤ (#‘𝑥) ∧ (𝑥𝐶𝑀) ⊆ (◡𝑓 “ {𝑐})))} ⇒ ⊢ (((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ 𝑉 ∧ 𝐹:𝑅⟶ℕ0) ∧ 𝐴 ∈ 𝑇) → (𝑀 Ramsey 𝐹) ≤ 𝐴) | ||
Theorem | ramub 15764* | The Ramsey number is a lower bound on the set of all numbers with the Ramsey number property. (Contributed by Mario Carneiro, 22-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑅 ∈ 𝑉) & ⊢ (𝜑 → 𝐹:𝑅⟶ℕ0) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ ((𝜑 ∧ (𝑁 ≤ (#‘𝑠) ∧ 𝑓:(𝑠𝐶𝑀)⟶𝑅)) → ∃𝑐 ∈ 𝑅 ∃𝑥 ∈ 𝒫 𝑠((𝐹‘𝑐) ≤ (#‘𝑥) ∧ (𝑥𝐶𝑀) ⊆ (◡𝑓 “ {𝑐}))) ⇒ ⊢ (𝜑 → (𝑀 Ramsey 𝐹) ≤ 𝑁) | ||
Theorem | ramub2 15765* | It is sufficient to check the Ramsey property on finite sets of size equal to the upper bound. (Contributed by Mario Carneiro, 23-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑅 ∈ 𝑉) & ⊢ (𝜑 → 𝐹:𝑅⟶ℕ0) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ ((𝜑 ∧ ((#‘𝑠) = 𝑁 ∧ 𝑓:(𝑠𝐶𝑀)⟶𝑅)) → ∃𝑐 ∈ 𝑅 ∃𝑥 ∈ 𝒫 𝑠((𝐹‘𝑐) ≤ (#‘𝑥) ∧ (𝑥𝐶𝑀) ⊆ (◡𝑓 “ {𝑐}))) ⇒ ⊢ (𝜑 → (𝑀 Ramsey 𝐹) ≤ 𝑁) | ||
Theorem | rami 15766* | The defining property of a Ramsey number. (Contributed by Mario Carneiro, 22-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑅 ∈ 𝑉) & ⊢ (𝜑 → 𝐹:𝑅⟶ℕ0) & ⊢ (𝜑 → (𝑀 Ramsey 𝐹) ∈ ℕ0) & ⊢ (𝜑 → 𝑆 ∈ 𝑊) & ⊢ (𝜑 → (𝑀 Ramsey 𝐹) ≤ (#‘𝑆)) & ⊢ (𝜑 → 𝐺:(𝑆𝐶𝑀)⟶𝑅) ⇒ ⊢ (𝜑 → ∃𝑐 ∈ 𝑅 ∃𝑥 ∈ 𝒫 𝑆((𝐹‘𝑐) ≤ (#‘𝑥) ∧ (𝑥𝐶𝑀) ⊆ (◡𝐺 “ {𝑐}))) | ||
Theorem | ramcl2 15767 | The Ramsey number is either a nonnegative integer or plus infinity. (Contributed by Mario Carneiro, 20-Apr-2015.) (Revised by AV, 14-Sep-2020.) |
⊢ ((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ 𝑉 ∧ 𝐹:𝑅⟶ℕ0) → (𝑀 Ramsey 𝐹) ∈ (ℕ0 ∪ {+∞})) | ||
Theorem | ramxrcl 15768 | The Ramsey number is an extended real number. (This theorem does not imply Ramsey's theorem, unlike ramcl 15780.) (Contributed by Mario Carneiro, 20-Apr-2015.) |
⊢ ((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ 𝑉 ∧ 𝐹:𝑅⟶ℕ0) → (𝑀 Ramsey 𝐹) ∈ ℝ*) | ||
Theorem | ramubcl 15769 | If the Ramsey number is upper bounded, then it is an integer. (Contributed by Mario Carneiro, 20-Apr-2015.) |
⊢ (((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ 𝑉 ∧ 𝐹:𝑅⟶ℕ0) ∧ (𝐴 ∈ ℕ0 ∧ (𝑀 Ramsey 𝐹) ≤ 𝐴)) → (𝑀 Ramsey 𝐹) ∈ ℕ0) | ||
Theorem | ramlb 15770* | Establish a lower bound on a Ramsey number. (Contributed by Mario Carneiro, 22-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑅 ∈ 𝑉) & ⊢ (𝜑 → 𝐹:𝑅⟶ℕ0) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → 𝐺:((1...𝑁)𝐶𝑀)⟶𝑅) & ⊢ ((𝜑 ∧ (𝑐 ∈ 𝑅 ∧ 𝑥 ⊆ (1...𝑁))) → ((𝑥𝐶𝑀) ⊆ (◡𝐺 “ {𝑐}) → (#‘𝑥) < (𝐹‘𝑐))) ⇒ ⊢ (𝜑 → 𝑁 < (𝑀 Ramsey 𝐹)) | ||
Theorem | 0ram 15771* | The Ramsey number when 𝑀 = 0. (Contributed by Mario Carneiro, 22-Apr-2015.) |
⊢ (((𝑅 ∈ 𝑉 ∧ 𝑅 ≠ ∅ ∧ 𝐹:𝑅⟶ℕ0) ∧ ∃𝑥 ∈ ℤ ∀𝑦 ∈ ran 𝐹 𝑦 ≤ 𝑥) → (0 Ramsey 𝐹) = sup(ran 𝐹, ℝ, < )) | ||
Theorem | 0ram2 15772 | The Ramsey number when 𝑀 = 0. (Contributed by Mario Carneiro, 22-Apr-2015.) |
⊢ ((𝑅 ∈ Fin ∧ 𝑅 ≠ ∅ ∧ 𝐹:𝑅⟶ℕ0) → (0 Ramsey 𝐹) = sup(ran 𝐹, ℝ, < )) | ||
Theorem | ram0 15773 | The Ramsey number when 𝑅 = ∅. (Contributed by Mario Carneiro, 22-Apr-2015.) |
⊢ (𝑀 ∈ ℕ0 → (𝑀 Ramsey ∅) = 𝑀) | ||
Theorem | 0ramcl 15774 | Lemma for ramcl 15780: Existence of the Ramsey number when 𝑀 = 0. (Contributed by Mario Carneiro, 23-Apr-2015.) |
⊢ ((𝑅 ∈ Fin ∧ 𝐹:𝑅⟶ℕ0) → (0 Ramsey 𝐹) ∈ ℕ0) | ||
Theorem | ramz2 15775 | The Ramsey number when 𝐹 has value zero for some color 𝐶. (Contributed by Mario Carneiro, 22-Apr-2015.) |
⊢ (((𝑀 ∈ ℕ ∧ 𝑅 ∈ 𝑉 ∧ 𝐹:𝑅⟶ℕ0) ∧ (𝐶 ∈ 𝑅 ∧ (𝐹‘𝐶) = 0)) → (𝑀 Ramsey 𝐹) = 0) | ||
Theorem | ramz 15776 | The Ramsey number when 𝐹 is the zero function. (Contributed by Mario Carneiro, 22-Apr-2015.) |
⊢ ((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ 𝑉 ∧ 𝑅 ≠ ∅) → (𝑀 Ramsey (𝑅 × {0})) = 0) | ||
Theorem | ramub1lem1 15777* | Lemma for ramub1 15779. (Contributed by Mario Carneiro, 23-Apr-2015.) |
⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐹:𝑅⟶ℕ) & ⊢ 𝐺 = (𝑥 ∈ 𝑅 ↦ (𝑀 Ramsey (𝑦 ∈ 𝑅 ↦ if(𝑦 = 𝑥, ((𝐹‘𝑥) − 1), (𝐹‘𝑦))))) & ⊢ (𝜑 → 𝐺:𝑅⟶ℕ0) & ⊢ (𝜑 → ((𝑀 − 1) Ramsey 𝐺) ∈ ℕ0) & ⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ (𝜑 → 𝑆 ∈ Fin) & ⊢ (𝜑 → (#‘𝑆) = (((𝑀 − 1) Ramsey 𝐺) + 1)) & ⊢ (𝜑 → 𝐾:(𝑆𝐶𝑀)⟶𝑅) & ⊢ (𝜑 → 𝑋 ∈ 𝑆) & ⊢ 𝐻 = (𝑢 ∈ ((𝑆 ∖ {𝑋})𝐶(𝑀 − 1)) ↦ (𝐾‘(𝑢 ∪ {𝑋}))) & ⊢ (𝜑 → 𝐷 ∈ 𝑅) & ⊢ (𝜑 → 𝑊 ⊆ (𝑆 ∖ {𝑋})) & ⊢ (𝜑 → (𝐺‘𝐷) ≤ (#‘𝑊)) & ⊢ (𝜑 → (𝑊𝐶(𝑀 − 1)) ⊆ (◡𝐻 “ {𝐷})) & ⊢ (𝜑 → 𝐸 ∈ 𝑅) & ⊢ (𝜑 → 𝑉 ⊆ 𝑊) & ⊢ (𝜑 → if(𝐸 = 𝐷, ((𝐹‘𝐷) − 1), (𝐹‘𝐸)) ≤ (#‘𝑉)) & ⊢ (𝜑 → (𝑉𝐶𝑀) ⊆ (◡𝐾 “ {𝐸})) ⇒ ⊢ (𝜑 → ∃𝑧 ∈ 𝒫 𝑆((𝐹‘𝐸) ≤ (#‘𝑧) ∧ (𝑧𝐶𝑀) ⊆ (◡𝐾 “ {𝐸}))) | ||
Theorem | ramub1lem2 15778* | Lemma for ramub1 15779. (Contributed by Mario Carneiro, 23-Apr-2015.) |
⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐹:𝑅⟶ℕ) & ⊢ 𝐺 = (𝑥 ∈ 𝑅 ↦ (𝑀 Ramsey (𝑦 ∈ 𝑅 ↦ if(𝑦 = 𝑥, ((𝐹‘𝑥) − 1), (𝐹‘𝑦))))) & ⊢ (𝜑 → 𝐺:𝑅⟶ℕ0) & ⊢ (𝜑 → ((𝑀 − 1) Ramsey 𝐺) ∈ ℕ0) & ⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) & ⊢ (𝜑 → 𝑆 ∈ Fin) & ⊢ (𝜑 → (#‘𝑆) = (((𝑀 − 1) Ramsey 𝐺) + 1)) & ⊢ (𝜑 → 𝐾:(𝑆𝐶𝑀)⟶𝑅) & ⊢ (𝜑 → 𝑋 ∈ 𝑆) & ⊢ 𝐻 = (𝑢 ∈ ((𝑆 ∖ {𝑋})𝐶(𝑀 − 1)) ↦ (𝐾‘(𝑢 ∪ {𝑋}))) ⇒ ⊢ (𝜑 → ∃𝑐 ∈ 𝑅 ∃𝑧 ∈ 𝒫 𝑆((𝐹‘𝑐) ≤ (#‘𝑧) ∧ (𝑧𝐶𝑀) ⊆ (◡𝐾 “ {𝑐}))) | ||
Theorem | ramub1 15779* | Inductive step for Ramsey's theorem, in the form of an explicit upper bound. (Contributed by Mario Carneiro, 23-Apr-2015.) |
⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝑅 ∈ Fin) & ⊢ (𝜑 → 𝐹:𝑅⟶ℕ) & ⊢ 𝐺 = (𝑥 ∈ 𝑅 ↦ (𝑀 Ramsey (𝑦 ∈ 𝑅 ↦ if(𝑦 = 𝑥, ((𝐹‘𝑥) − 1), (𝐹‘𝑦))))) & ⊢ (𝜑 → 𝐺:𝑅⟶ℕ0) & ⊢ (𝜑 → ((𝑀 − 1) Ramsey 𝐺) ∈ ℕ0) ⇒ ⊢ (𝜑 → (𝑀 Ramsey 𝐹) ≤ (((𝑀 − 1) Ramsey 𝐺) + 1)) | ||
Theorem | ramcl 15780 | Ramsey's theorem: the Ramsey number is an integer for every finite coloring and set of upper bounds. (Contributed by Mario Carneiro, 23-Apr-2015.) |
⊢ ((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ Fin ∧ 𝐹:𝑅⟶ℕ0) → (𝑀 Ramsey 𝐹) ∈ ℕ0) | ||
Theorem | ramsey 15781* | Ramsey's theorem with the definition Ramsey eliminated. If 𝑀 is an integer, 𝑅 is a specified finite set of colors, and 𝐹:𝑅⟶ℕ0 is a set of lower bounds for each color, then there is an 𝑛 such that for every set 𝑠 of size greater than 𝑛 and every coloring 𝑓 of the set (𝑠𝐶𝑀) of all 𝑀-element subsets of 𝑠, there is a color 𝑐 and a subset 𝑥 ⊆ 𝑠 such that 𝑥 is larger than 𝐹(𝑐) and the 𝑀-element subsets of 𝑥 are monochromatic with color 𝑐. This is the hypergraph version of Ramsey's theorem; the version for simple graphs is the case 𝑀 = 2. This is Metamath 100 proof #31. (Contributed by Mario Carneiro, 23-Apr-2015.) |
⊢ 𝐶 = (𝑎 ∈ V, 𝑖 ∈ ℕ0 ↦ {𝑏 ∈ 𝒫 𝑎 ∣ (#‘𝑏) = 𝑖}) ⇒ ⊢ ((𝑀 ∈ ℕ0 ∧ 𝑅 ∈ Fin ∧ 𝐹:𝑅⟶ℕ0) → ∃𝑛 ∈ ℕ0 ∀𝑠(𝑛 ≤ (#‘𝑠) → ∀𝑓 ∈ (𝑅 ↑𝑚 (𝑠𝐶𝑀))∃𝑐 ∈ 𝑅 ∃𝑥 ∈ 𝒫 𝑠((𝐹‘𝑐) ≤ (#‘𝑥) ∧ (𝑥𝐶𝑀) ⊆ (◡𝑓 “ {𝑐})))) | ||
According to Wikipedia "Primorial", https://en.wikipedia.org/wiki/Primorial (28-Aug-2020): "In mathematics, and more particularly in number theory, primorial, denoted by "#", is a function from natural numbers to natural numbers similar to the factorial function, but rather than successively multiplying [all] positive integers [less than or equal to a given number], the function only multiplies [the] prime numbers [less than or equal to the given number]. The name "primorial", coined by Harvey Dubner, draws an analogy to primes similar to the way the name "factorial" relates to factors." | ||
Syntax | cprmo 15782 | Extend class notation to include the primorial of nonnegative integers. |
class #p | ||
Definition | df-prmo 15783* |
Define the primorial function on nonnegative integers as the product of
all prime numbers less than or equal to the integer. For example,
(#p‘10) = 2 · 3 · 5
· 7 = 210 (see ex-prmo 27446).
In the literature, the primorial function is written as a postscript hash: 6# = 30. In contrast to prmorcht 24949, where the primorial function is defined by using the sequence builder (𝑃 = seq1( · , 𝐹)), the more specialized definition of a product of a series is used here. (Contributed by AV, 28-Aug-2020.) |
⊢ #p = (𝑛 ∈ ℕ0 ↦ ∏𝑘 ∈ (1...𝑛)if(𝑘 ∈ ℙ, 𝑘, 1)) | ||
Theorem | prmoval 15784* | Value of the primorial function for nonnegative integers. (Contributed by AV, 28-Aug-2020.) |
⊢ (𝑁 ∈ ℕ0 → (#p‘𝑁) = ∏𝑘 ∈ (1...𝑁)if(𝑘 ∈ ℙ, 𝑘, 1)) | ||
Theorem | prmocl 15785 | Closure of the primorial function. (Contributed by AV, 28-Aug-2020.) |
⊢ (𝑁 ∈ ℕ0 → (#p‘𝑁) ∈ ℕ) | ||
Theorem | prmone0 15786 | The primorial function is nonzero. (Contributed by AV, 28-Aug-2020.) |
⊢ (𝑁 ∈ ℕ0 → (#p‘𝑁) ≠ 0) | ||
Theorem | prmo0 15787 | The primorial of 0. (Contributed by AV, 28-Aug-2020.) |
⊢ (#p‘0) = 1 | ||
Theorem | prmo1 15788 | The primorial of 1. (Contributed by AV, 28-Aug-2020.) |
⊢ (#p‘1) = 1 | ||
Theorem | prmop1 15789 | The primorial of a successor. (Contributed by AV, 28-Aug-2020.) |
⊢ (𝑁 ∈ ℕ0 → (#p‘(𝑁 + 1)) = if((𝑁 + 1) ∈ ℙ, ((#p‘𝑁) · (𝑁 + 1)), (#p‘𝑁))) | ||
Theorem | prmonn2 15790 | Value of the primorial function expressed recursively. (Contributed by AV, 28-Aug-2020.) |
⊢ (𝑁 ∈ ℕ → (#p‘𝑁) = if(𝑁 ∈ ℙ, ((#p‘(𝑁 − 1)) · 𝑁), (#p‘(𝑁 − 1)))) | ||
Theorem | prmo2 15791 | The primorial of 2. (Contributed by AV, 28-Aug-2020.) |
⊢ (#p‘2) = 2 | ||
Theorem | prmo3 15792 | The primorial of 3. (Contributed by AV, 28-Aug-2020.) |
⊢ (#p‘3) = 6 | ||
Theorem | prmdvdsprmo 15793* | The primorial of a number is divisible by each prime less then or equal to the number. (Contributed by AV, 15-Aug-2020.) (Revised by AV, 28-Aug-2020.) |
⊢ (𝑁 ∈ ℕ → ∀𝑝 ∈ ℙ (𝑝 ≤ 𝑁 → 𝑝 ∥ (#p‘𝑁))) | ||
Theorem | prmdvdsprmop 15794* | The primorial of a number plus an integer greater than 1 and less then or equal to the number is divisible by a prime less then or equal to the number. (Contributed by AV, 15-Aug-2020.) (Revised by AV, 28-Aug-2020.) |
⊢ ((𝑁 ∈ ℕ ∧ 𝐼 ∈ (2...𝑁)) → ∃𝑝 ∈ ℙ (𝑝 ≤ 𝑁 ∧ 𝑝 ∥ 𝐼 ∧ 𝑝 ∥ ((#p‘𝑁) + 𝐼))) | ||
Theorem | fvprmselelfz 15795* | The value of the prime selection function is in a finite sequence of integers if the argument is in this finite sequence of integers. (Contributed by AV, 19-Aug-2020.) |
⊢ 𝐹 = (𝑚 ∈ ℕ ↦ if(𝑚 ∈ ℙ, 𝑚, 1)) ⇒ ⊢ ((𝑁 ∈ ℕ ∧ 𝑋 ∈ (1...𝑁)) → (𝐹‘𝑋) ∈ (1...𝑁)) | ||
Theorem | fvprmselgcd1 15796* | The greatest common divisor of two values of the prime selection function for different arguments is 1. (Contributed by AV, 19-Aug-2020.) |
⊢ 𝐹 = (𝑚 ∈ ℕ ↦ if(𝑚 ∈ ℙ, 𝑚, 1)) ⇒ ⊢ ((𝑋 ∈ (1...𝑁) ∧ 𝑌 ∈ (1...𝑁) ∧ 𝑋 ≠ 𝑌) → ((𝐹‘𝑋) gcd (𝐹‘𝑌)) = 1) | ||
Theorem | prmolefac 15797 | The primorial of a positive integer is less than or equal to the factorial of the integer. (Contributed by AV, 15-Aug-2020.) (Revised by AV, 29-Aug-2020.) |
⊢ (𝑁 ∈ ℕ0 → (#p‘𝑁) ≤ (!‘𝑁)) | ||
Theorem | prmodvdslcmf 15798 | The primorial of a nonnegative integer divides the least common multiple of all positive integers less than or equal to the integer. (Contributed by AV, 19-Aug-2020.) (Revised by AV, 29-Aug-2020.) |
⊢ (𝑁 ∈ ℕ0 → (#p‘𝑁) ∥ (lcm‘(1...𝑁))) | ||
Theorem | prmolelcmf 15799 | The primorial of a positive integer is less than or equal to the least common multiple of all positive integers less than or equal to the integer. (Contributed by AV, 19-Aug-2020.) (Revised by AV, 29-Aug-2020.) |
⊢ (𝑁 ∈ ℕ0 → (#p‘𝑁) ≤ (lcm‘(1...𝑁))) | ||
According to Wikipedia "Prime gap", https://en.wikipedia.org/wiki/Prime_gap (16-Aug-2020): "A prime gap is the difference between two successive prime numbers. The n-th prime gap, denoted gn or g(pn) is the difference between the (n+1)-th and the n-th prime numbers, i.e. gn = pn+1 - pn . We have g1 = 1, g2 = g3 = 2, and g4 = 4." It can be proven that there are arbitrary large gaps, usually by showing that "in the sequence n!+2, n!+3, ..., n!+n the first term is divisible by 2, the second term is divisible by 3, and so on. Thus, this is a sequence of n-1 consecutive composite integers, and it must belong to a gap between primes having length at least n.", see prmgap 15810. Instead of using the factorial of n (see df-fac 13101), any function can be chosen for which f(n) is not relatively prime to the integers 2, 3, ..., n. For example, the least common multiple of the integers 2, 3, ..., n, see prmgaplcm 15811, or the primorial n# (the product of all prime numbers less than or equal to n), see prmgapprmo 15813, are such functions, which provide smaller values than the factorial function (see lcmflefac 15408 and prmolefac 15797 resp. prmolelcmf 15799): "For instance, the first prime gap of size larger than 14 occurs between the primes 523 and 541, while 15! is the vastly larger number 1307674368000." But the least common multiple of the integers 2, 3, ..., 15 is 360360, and 15# is 30030 (p3248 = 30029 and P3249 = 30047, so g3248 = 18). | ||
Theorem | prmgaplem1 15800 | Lemma for prmgap 15810: The factorial of a number plus an integer greater than 1 and less then or equal to the number is divisible by that integer. (Contributed by AV, 13-Aug-2020.) |
⊢ ((𝑁 ∈ ℕ ∧ 𝐼 ∈ (2...𝑁)) → 𝐼 ∥ ((!‘𝑁) + 𝐼)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |