Intuitionistic Logic Explorer Most Recent Proofs |
||
Mirrors > Home > ILE Home > Th. List > Recent | MPE Most Recent Other > MM 100 |
See the MPE Most Recent Proofs page for news and some useful links.
Color key: | Intuitionistic Logic Explorer | User Mathboxes |
Date | Label | Description |
---|---|---|
Theorem | ||
1-Oct-2024 | infex2g 6979 | Existence of infimum. (Contributed by Jim Kingdon, 1-Oct-2024.) |
⊢ (𝐴 ∈ 𝐶 → inf(𝐵, 𝐴, 𝑅) ∈ V) | ||
30-Sep-2024 | unbendc 12227 | An unbounded decidable set of positive integers is infinite. (Contributed by NM, 5-May-2005.) (Revised by Jim Kingdon, 30-Sep-2024.) |
⊢ ((𝐴 ⊆ ℕ ∧ ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴 ∧ ∀𝑚 ∈ ℕ ∃𝑛 ∈ 𝐴 𝑚 < 𝑛) → 𝐴 ≈ ℕ) | ||
30-Sep-2024 | prmdc 12011 | Primality is decidable. (Contributed by Jim Kingdon, 30-Sep-2024.) |
⊢ (𝑁 ∈ ℕ → DECID 𝑁 ∈ ℙ) | ||
30-Sep-2024 | dcfi 6926 | Decidability of a family of propositions indexed by a finite set. (Contributed by Jim Kingdon, 30-Sep-2024.) |
⊢ ((𝐴 ∈ Fin ∧ ∀𝑥 ∈ 𝐴 DECID 𝜑) → DECID ∀𝑥 ∈ 𝐴 𝜑) | ||
29-Sep-2024 | ssnnct 12218 | A decidable subset of ℕ is countable. (Contributed by Jim Kingdon, 29-Sep-2024.) |
⊢ ((𝐴 ⊆ ℕ ∧ ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴) → ∃𝑓 𝑓:ω–onto→(𝐴 ⊔ 1o)) | ||
29-Sep-2024 | ssnnctlemct 12217 | Lemma for ssnnct 12218. The result. (Contributed by Jim Kingdon, 29-Sep-2024.) |
⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 1) ⇒ ⊢ ((𝐴 ⊆ ℕ ∧ ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴) → ∃𝑓 𝑓:ω–onto→(𝐴 ⊔ 1o)) | ||
28-Sep-2024 | nninfdcex 11843 | A decidable set of natural numbers has an infimum. (Contributed by Jim Kingdon, 28-Sep-2024.) |
⊢ (𝜑 → 𝐴 ⊆ ℕ) & ⊢ (𝜑 → ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑦 𝑦 ∈ 𝐴) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℝ (∀𝑦 ∈ 𝐴 ¬ 𝑦 < 𝑥 ∧ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 → ∃𝑧 ∈ 𝐴 𝑧 < 𝑦))) | ||
27-Sep-2024 | infregelbex 9510 | Any lower bound of a set of real numbers with an infimum is less than or equal to the infimum. (Contributed by Jim Kingdon, 27-Sep-2024.) |
⊢ (𝜑 → ∃𝑥 ∈ ℝ (∀𝑦 ∈ 𝐴 ¬ 𝑦 < 𝑥 ∧ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 → ∃𝑧 ∈ 𝐴 𝑧 < 𝑦))) & ⊢ (𝜑 → 𝐴 ⊆ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) ⇒ ⊢ (𝜑 → (𝐵 ≤ inf(𝐴, ℝ, < ) ↔ ∀𝑧 ∈ 𝐴 𝐵 ≤ 𝑧)) | ||
26-Sep-2024 | nninfdclemp1 12223 | Lemma for nninfdc 12226. Each element of the sequence 𝐹 is greater than the previous element. (Contributed by Jim Kingdon, 26-Sep-2024.) |
⊢ (𝜑 → 𝐴 ⊆ ℕ) & ⊢ (𝜑 → ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑚 ∈ ℕ ∃𝑛 ∈ 𝐴 𝑚 < 𝑛) & ⊢ (𝜑 → (𝐽 ∈ 𝐴 ∧ 1 < 𝐽)) & ⊢ 𝐹 = seq1((𝑦 ∈ ℕ, 𝑧 ∈ ℕ ↦ inf((𝐴 ∩ (ℤ≥‘(𝑦 + 1))), ℝ, < )), (𝑖 ∈ ℕ ↦ 𝐽)) & ⊢ (𝜑 → 𝑈 ∈ ℕ) ⇒ ⊢ (𝜑 → (𝐹‘𝑈) < (𝐹‘(𝑈 + 1))) | ||
26-Sep-2024 | nnminle 12220 | The infimum of a decidable subset of the natural numbers is less than an element of the set. The infimum is also a minimum as shown at nnmindc 12219. (Contributed by Jim Kingdon, 26-Sep-2024.) |
⊢ ((𝐴 ⊆ ℕ ∧ ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴 ∧ 𝐵 ∈ 𝐴) → inf(𝐴, ℝ, < ) ≤ 𝐵) | ||
25-Sep-2024 | nninfdclemcl 12221 | Lemma for nninfdc 12226. (Contributed by Jim Kingdon, 25-Sep-2024.) |
⊢ (𝜑 → 𝐴 ⊆ ℕ) & ⊢ (𝜑 → ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑚 ∈ ℕ ∃𝑛 ∈ 𝐴 𝑚 < 𝑛) & ⊢ (𝜑 → 𝑃 ∈ 𝐴) & ⊢ (𝜑 → 𝑄 ∈ 𝐴) ⇒ ⊢ (𝜑 → (𝑃(𝑦 ∈ ℕ, 𝑧 ∈ ℕ ↦ inf((𝐴 ∩ (ℤ≥‘(𝑦 + 1))), ℝ, < ))𝑄) ∈ 𝐴) | ||
24-Sep-2024 | nninfdclemlt 12224 | Lemma for nninfdc 12226. The function from nninfdclemf 12222 is strictly monotonic. (Contributed by Jim Kingdon, 24-Sep-2024.) |
⊢ (𝜑 → 𝐴 ⊆ ℕ) & ⊢ (𝜑 → ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑚 ∈ ℕ ∃𝑛 ∈ 𝐴 𝑚 < 𝑛) & ⊢ (𝜑 → (𝐽 ∈ 𝐴 ∧ 1 < 𝐽)) & ⊢ 𝐹 = seq1((𝑦 ∈ ℕ, 𝑧 ∈ ℕ ↦ inf((𝐴 ∩ (ℤ≥‘(𝑦 + 1))), ℝ, < )), (𝑖 ∈ ℕ ↦ 𝐽)) & ⊢ (𝜑 → 𝑈 ∈ ℕ) & ⊢ (𝜑 → 𝑉 ∈ ℕ) & ⊢ (𝜑 → 𝑈 < 𝑉) ⇒ ⊢ (𝜑 → (𝐹‘𝑈) < (𝐹‘𝑉)) | ||
23-Sep-2024 | nninfdc 12226 | An unbounded decidable set of positive integers is infinite. (Contributed by Jim Kingdon, 23-Sep-2024.) |
⊢ ((𝐴 ⊆ ℕ ∧ ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴 ∧ ∀𝑚 ∈ ℕ ∃𝑛 ∈ 𝐴 𝑚 < 𝑛) → ω ≼ 𝐴) | ||
23-Sep-2024 | nninfdclemf1 12225 | Lemma for nninfdc 12226. The function from nninfdclemf 12222 is one-to-one. (Contributed by Jim Kingdon, 23-Sep-2024.) |
⊢ (𝜑 → 𝐴 ⊆ ℕ) & ⊢ (𝜑 → ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑚 ∈ ℕ ∃𝑛 ∈ 𝐴 𝑚 < 𝑛) & ⊢ (𝜑 → (𝐽 ∈ 𝐴 ∧ 1 < 𝐽)) & ⊢ 𝐹 = seq1((𝑦 ∈ ℕ, 𝑧 ∈ ℕ ↦ inf((𝐴 ∩ (ℤ≥‘(𝑦 + 1))), ℝ, < )), (𝑖 ∈ ℕ ↦ 𝐽)) ⇒ ⊢ (𝜑 → 𝐹:ℕ–1-1→𝐴) | ||
23-Sep-2024 | nninfdclemf 12222 | Lemma for nninfdc 12226. A function from the natural numbers into 𝐴. (Contributed by Jim Kingdon, 23-Sep-2024.) |
⊢ (𝜑 → 𝐴 ⊆ ℕ) & ⊢ (𝜑 → ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑚 ∈ ℕ ∃𝑛 ∈ 𝐴 𝑚 < 𝑛) & ⊢ (𝜑 → (𝐽 ∈ 𝐴 ∧ 1 < 𝐽)) & ⊢ 𝐹 = seq1((𝑦 ∈ ℕ, 𝑧 ∈ ℕ ↦ inf((𝐴 ∩ (ℤ≥‘(𝑦 + 1))), ℝ, < )), (𝑖 ∈ ℕ ↦ 𝐽)) ⇒ ⊢ (𝜑 → 𝐹:ℕ⟶𝐴) | ||
23-Sep-2024 | nnmindc 12219 | An inhabited decidable subset of the natural numbers has a minimum. (Contributed by Jim Kingdon, 23-Sep-2024.) |
⊢ ((𝐴 ⊆ ℕ ∧ ∀𝑥 ∈ ℕ DECID 𝑥 ∈ 𝐴 ∧ ∃𝑦 𝑦 ∈ 𝐴) → inf(𝐴, ℝ, < ) ∈ 𝐴) | ||
19-Sep-2024 | ssomct 12216 | A decidable subset of ω is countable. (Contributed by Jim Kingdon, 19-Sep-2024.) |
⊢ ((𝐴 ⊆ ω ∧ ∀𝑥 ∈ ω DECID 𝑥 ∈ 𝐴) → ∃𝑓 𝑓:ω–onto→(𝐴 ⊔ 1o)) | ||
14-Sep-2024 | nnpredlt 4584 | The predecessor (see nnpredcl 4583) of a nonzero natural number is less than (see df-iord 4327) that number. (Contributed by Jim Kingdon, 14-Sep-2024.) |
⊢ ((𝐴 ∈ ω ∧ 𝐴 ≠ ∅) → ∪ 𝐴 ∈ 𝐴) | ||
13-Sep-2024 | nninfisollemeq 7076 | Lemma for nninfisol 7077. The case where 𝑁 is a successor and 𝑁 and 𝑋 are equal. (Contributed by Jim Kingdon, 13-Sep-2024.) |
⊢ (𝜑 → 𝑋 ∈ ℕ∞) & ⊢ (𝜑 → (𝑋‘𝑁) = ∅) & ⊢ (𝜑 → 𝑁 ∈ ω) & ⊢ (𝜑 → 𝑁 ≠ ∅) & ⊢ (𝜑 → (𝑋‘∪ 𝑁) = 1o) ⇒ ⊢ (𝜑 → DECID (𝑖 ∈ ω ↦ if(𝑖 ∈ 𝑁, 1o, ∅)) = 𝑋) | ||
13-Sep-2024 | nninfisollemne 7075 | Lemma for nninfisol 7077. A case where 𝑁 is a successor and 𝑁 and 𝑋 are not equal. (Contributed by Jim Kingdon, 13-Sep-2024.) |
⊢ (𝜑 → 𝑋 ∈ ℕ∞) & ⊢ (𝜑 → (𝑋‘𝑁) = ∅) & ⊢ (𝜑 → 𝑁 ∈ ω) & ⊢ (𝜑 → 𝑁 ≠ ∅) & ⊢ (𝜑 → (𝑋‘∪ 𝑁) = ∅) ⇒ ⊢ (𝜑 → DECID (𝑖 ∈ ω ↦ if(𝑖 ∈ 𝑁, 1o, ∅)) = 𝑋) | ||
13-Sep-2024 | nninfisollem0 7074 | Lemma for nninfisol 7077. The case where 𝑁 is zero. (Contributed by Jim Kingdon, 13-Sep-2024.) |
⊢ (𝜑 → 𝑋 ∈ ℕ∞) & ⊢ (𝜑 → (𝑋‘𝑁) = ∅) & ⊢ (𝜑 → 𝑁 ∈ ω) & ⊢ (𝜑 → 𝑁 = ∅) ⇒ ⊢ (𝜑 → DECID (𝑖 ∈ ω ↦ if(𝑖 ∈ 𝑁, 1o, ∅)) = 𝑋) | ||
12-Sep-2024 | nninfisol 7077 | Finite elements of ℕ∞ are isolated. That is, given a natural number and any element of ℕ∞, it is decidable whether the natural number (when converted to an element of ℕ∞) is equal to the given element of ℕ∞. Stated in an online post by Martin Escardo. One way to understand this theorem is that you do not need to look at an unbounded number of elements of the sequence 𝑋 to decide whether it is equal to 𝑁 (in fact, you only need to look at two elements and 𝑁 tells you where to look). (Contributed by BJ and Jim Kingdon, 12-Sep-2024.) |
⊢ ((𝑁 ∈ ω ∧ 𝑋 ∈ ℕ∞) → DECID (𝑖 ∈ ω ↦ if(𝑖 ∈ 𝑁, 1o, ∅)) = 𝑋) | ||
7-Sep-2024 | eulerthlemfi 12107 | Lemma for eulerth 12112. The set 𝑆 is finite. (Contributed by Mario Carneiro, 28-Feb-2014.) (Revised by Jim Kingdon, 7-Sep-2024.) |
⊢ (𝜑 → (𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ (𝐴 gcd 𝑁) = 1)) & ⊢ 𝑆 = {𝑦 ∈ (0..^𝑁) ∣ (𝑦 gcd 𝑁) = 1} ⇒ ⊢ (𝜑 → 𝑆 ∈ Fin) | ||
7-Sep-2024 | modqexp 10548 | Exponentiation property of the modulo operation, see theorem 5.2(c) in [ApostolNT] p. 107. (Contributed by Mario Carneiro, 28-Feb-2014.) (Revised by Jim Kingdon, 7-Sep-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℤ) & ⊢ (𝜑 → 𝐵 ∈ ℤ) & ⊢ (𝜑 → 𝐶 ∈ ℕ0) & ⊢ (𝜑 → 𝐷 ∈ ℚ) & ⊢ (𝜑 → 0 < 𝐷) & ⊢ (𝜑 → (𝐴 mod 𝐷) = (𝐵 mod 𝐷)) ⇒ ⊢ (𝜑 → ((𝐴↑𝐶) mod 𝐷) = ((𝐵↑𝐶) mod 𝐷)) | ||
5-Sep-2024 | eulerthlemh 12110 | Lemma for eulerth 12112. A permutation of (1...(ϕ‘𝑁)). (Contributed by Mario Carneiro, 28-Feb-2014.) (Revised by Jim Kingdon, 5-Sep-2024.) |
⊢ (𝜑 → (𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ (𝐴 gcd 𝑁) = 1)) & ⊢ 𝑆 = {𝑦 ∈ (0..^𝑁) ∣ (𝑦 gcd 𝑁) = 1} & ⊢ (𝜑 → 𝐹:(1...(ϕ‘𝑁))–1-1-onto→𝑆) & ⊢ 𝐻 = (◡𝐹 ∘ (𝑦 ∈ (1...(ϕ‘𝑁)) ↦ ((𝐴 · (𝐹‘𝑦)) mod 𝑁))) ⇒ ⊢ (𝜑 → 𝐻:(1...(ϕ‘𝑁))–1-1-onto→(1...(ϕ‘𝑁))) | ||
2-Sep-2024 | eulerthlemth 12111 | Lemma for eulerth 12112. The result. (Contributed by Mario Carneiro, 28-Feb-2014.) (Revised by Jim Kingdon, 2-Sep-2024.) |
⊢ (𝜑 → (𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ (𝐴 gcd 𝑁) = 1)) & ⊢ 𝑆 = {𝑦 ∈ (0..^𝑁) ∣ (𝑦 gcd 𝑁) = 1} & ⊢ (𝜑 → 𝐹:(1...(ϕ‘𝑁))–1-1-onto→𝑆) ⇒ ⊢ (𝜑 → ((𝐴↑(ϕ‘𝑁)) mod 𝑁) = (1 mod 𝑁)) | ||
2-Sep-2024 | eulerthlema 12109 | Lemma for eulerth 12112. (Contributed by Mario Carneiro, 28-Feb-2014.) (Revised by Jim Kingdon, 2-Sep-2024.) |
⊢ (𝜑 → (𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ (𝐴 gcd 𝑁) = 1)) & ⊢ 𝑆 = {𝑦 ∈ (0..^𝑁) ∣ (𝑦 gcd 𝑁) = 1} & ⊢ (𝜑 → 𝐹:(1...(ϕ‘𝑁))–1-1-onto→𝑆) ⇒ ⊢ (𝜑 → (((𝐴↑(ϕ‘𝑁)) · ∏𝑥 ∈ (1...(ϕ‘𝑁))(𝐹‘𝑥)) mod 𝑁) = (∏𝑥 ∈ (1...(ϕ‘𝑁))((𝐴 · (𝐹‘𝑥)) mod 𝑁) mod 𝑁)) | ||
2-Sep-2024 | eulerthlemrprm 12108 | Lemma for eulerth 12112. 𝑁 and ∏𝑥 ∈ (1...(ϕ‘𝑁))(𝐹‘𝑥) are relatively prime. (Contributed by Mario Carneiro, 28-Feb-2014.) (Revised by Jim Kingdon, 2-Sep-2024.) |
⊢ (𝜑 → (𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ (𝐴 gcd 𝑁) = 1)) & ⊢ 𝑆 = {𝑦 ∈ (0..^𝑁) ∣ (𝑦 gcd 𝑁) = 1} & ⊢ (𝜑 → 𝐹:(1...(ϕ‘𝑁))–1-1-onto→𝑆) ⇒ ⊢ (𝜑 → (𝑁 gcd ∏𝑥 ∈ (1...(ϕ‘𝑁))(𝐹‘𝑥)) = 1) | ||
30-Aug-2024 | fprodap0f 11537 | A finite product of terms apart from zero is apart from zero. A version of fprodap0 11522 using bound-variable hypotheses instead of distinct variable conditions. (Contributed by Glauco Siliprandi, 5-Apr-2020.) (Revised by Jim Kingdon, 30-Aug-2024.) |
⊢ Ⅎ𝑘𝜑 & ⊢ (𝜑 → 𝐴 ∈ Fin) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐵 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐵 # 0) ⇒ ⊢ (𝜑 → ∏𝑘 ∈ 𝐴 𝐵 # 0) | ||
28-Aug-2024 | fprodrec 11530 | The finite product of reciprocals is the reciprocal of the product. (Contributed by Jim Kingdon, 28-Aug-2024.) |
⊢ (𝜑 → 𝐴 ∈ Fin) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐵 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐵 # 0) ⇒ ⊢ (𝜑 → ∏𝑘 ∈ 𝐴 (1 / 𝐵) = (1 / ∏𝑘 ∈ 𝐴 𝐵)) | ||
26-Aug-2024 | exmidontri2or 7179 | Ordinal trichotomy is equivalent to excluded middle. (Contributed by Jim Kingdon, 26-Aug-2024.) |
⊢ (EXMID ↔ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥)) | ||
26-Aug-2024 | exmidontri 7175 | Ordinal trichotomy is equivalent to excluded middle. (Contributed by Jim Kingdon, 26-Aug-2024.) |
⊢ (EXMID ↔ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥)) | ||
26-Aug-2024 | ontri2orexmidim 4532 | Ordinal trichotomy implies excluded middle. Closed form of ordtri2or2exmid 4531. (Contributed by Jim Kingdon, 26-Aug-2024.) |
⊢ (∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥) → DECID 𝜑) | ||
26-Aug-2024 | ontriexmidim 4482 | Ordinal trichotomy implies excluded middle. Closed form of ordtriexmid 4481. (Contributed by Jim Kingdon, 26-Aug-2024.) |
⊢ (∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥) → DECID 𝜑) | ||
25-Aug-2024 | onntri2or 7182 | Double negated ordinal trichotomy. (Contributed by Jim Kingdon, 25-Aug-2024.) |
⊢ (¬ ¬ EXMID ↔ ∀𝑥 ∈ On ∀𝑦 ∈ On ¬ ¬ (𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥)) | ||
25-Aug-2024 | onntri3or 7181 | Double negated ordinal trichotomy. (Contributed by Jim Kingdon, 25-Aug-2024.) |
⊢ (¬ ¬ EXMID ↔ ∀𝑥 ∈ On ∀𝑦 ∈ On ¬ ¬ (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥)) | ||
25-Aug-2024 | csbcow 3042 | Composition law for chained substitutions into a class. Version of csbco 3041 with a disjoint variable condition, which requires fewer axioms. (Contributed by NM, 10-Nov-2005.) (Revised by Gino Giotto, 25-Aug-2024.) |
⊢ ⦋𝐴 / 𝑦⦌⦋𝑦 / 𝑥⦌𝐵 = ⦋𝐴 / 𝑥⦌𝐵 | ||
25-Aug-2024 | cbvreuvw 2686 | Version of cbvreuv 2682 with a disjoint variable condition. (Contributed by Gino Giotto, 10-Jan-2024.) Reduce axiom usage. (Revised by Gino Giotto, 25-Aug-2024.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃!𝑥 ∈ 𝐴 𝜑 ↔ ∃!𝑦 ∈ 𝐴 𝜓) | ||
25-Aug-2024 | cbvrexvw 2685 | Version of cbvrexv 2681 with a disjoint variable condition. (Contributed by Gino Giotto, 10-Jan-2024.) Reduce axiom usage. (Revised by Gino Giotto, 25-Aug-2024.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑥 ∈ 𝐴 𝜑 ↔ ∃𝑦 ∈ 𝐴 𝜓) | ||
25-Aug-2024 | cbvralvw 2684 | Version of cbvralv 2680 with a disjoint variable condition. (Contributed by Gino Giotto, 10-Jan-2024.) Reduce axiom usage. (Revised by Gino Giotto, 25-Aug-2024.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝜑 ↔ ∀𝑦 ∈ 𝐴 𝜓) | ||
25-Aug-2024 | cbvabw 2280 | Version of cbvab 2281 with a disjoint variable condition. (Contributed by Gino Giotto, 10-Jan-2024.) Reduce axiom usage. (Revised by Gino Giotto, 25-Aug-2024.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ {𝑥 ∣ 𝜑} = {𝑦 ∣ 𝜓} | ||
25-Aug-2024 | nfsbv 1927 | If 𝑧 is not free in 𝜑, it is not free in [𝑦 / 𝑥]𝜑 when 𝑧 is distinct from 𝑥 and 𝑦. Version of nfsb 1926 requiring more disjoint variables. (Contributed by Wolf Lammen, 7-Feb-2023.) Remove disjoint variable condition on 𝑥, 𝑦. (Revised by Steven Nguyen, 13-Aug-2023.) Reduce axiom usage. (Revised by Gino Giotto, 25-Aug-2024.) |
⊢ Ⅎ𝑧𝜑 ⇒ ⊢ Ⅎ𝑧[𝑦 / 𝑥]𝜑 | ||
25-Aug-2024 | cbvexvw 1900 | Change bound variable. See cbvexv 1898 for a version with fewer disjoint variable conditions. (Contributed by NM, 19-Apr-2017.) Avoid ax-7 1428. (Revised by Gino Giotto, 25-Aug-2024.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑥𝜑 ↔ ∃𝑦𝜓) | ||
25-Aug-2024 | cbvalvw 1899 | Change bound variable. See cbvalv 1897 for a version with fewer disjoint variable conditions. (Contributed by NM, 9-Apr-2017.) Avoid ax-7 1428. (Revised by Gino Giotto, 25-Aug-2024.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥𝜑 ↔ ∀𝑦𝜓) | ||
25-Aug-2024 | nfal 1556 | If 𝑥 is not free in 𝜑, it is not free in ∀𝑦𝜑. (Contributed by Mario Carneiro, 11-Aug-2016.) Remove dependency on ax-4 1490. (Revised by Gino Giotto, 25-Aug-2024.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ Ⅎ𝑥∀𝑦𝜑 | ||
24-Aug-2024 | gcdcomd 11862 | The gcd operator is commutative, deduction version. (Contributed by SN, 24-Aug-2024.) |
⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑁 ∈ ℤ) ⇒ ⊢ (𝜑 → (𝑀 gcd 𝑁) = (𝑁 gcd 𝑀)) | ||
17-Aug-2024 | fprodcl2lem 11506 | Finite product closure lemma. (Contributed by Scott Fenton, 14-Dec-2017.) (Revised by Jim Kingdon, 17-Aug-2024.) |
⊢ (𝜑 → 𝑆 ⊆ ℂ) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥 · 𝑦) ∈ 𝑆) & ⊢ (𝜑 → 𝐴 ∈ Fin) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐵 ∈ 𝑆) & ⊢ (𝜑 → 𝐴 ≠ ∅) ⇒ ⊢ (𝜑 → ∏𝑘 ∈ 𝐴 𝐵 ∈ 𝑆) | ||
16-Aug-2024 | if0ab 13422 |
Expression of a conditional class as a class abstraction when the False
alternative is the empty class: in that case, the conditional class is
the extension, in the True alternative, of the condition.
Remark: a consequence which could be formalized is the inclusion ⊢ if(𝜑, 𝐴, ∅) ⊆ 𝐴 and therefore, using elpwg 3551, ⊢ (𝐴 ∈ 𝑉 → if(𝜑, 𝐴, ∅) ∈ 𝒫 𝐴), from which fmelpw1o 13423 could be derived, yielding an alternative proof. (Contributed by BJ, 16-Aug-2024.) |
⊢ if(𝜑, 𝐴, ∅) = {𝑥 ∈ 𝐴 ∣ 𝜑} | ||
16-Aug-2024 | fprodunsn 11505 | Multiply in an additional term in a finite product. See also fprodsplitsn 11534 which is the same but with a Ⅎ𝑘𝜑 hypothesis in place of the distinct variable condition between 𝜑 and 𝑘. (Contributed by Jim Kingdon, 16-Aug-2024.) |
⊢ Ⅎ𝑘𝐷 & ⊢ (𝜑 → 𝐴 ∈ Fin) & ⊢ (𝜑 → 𝐵 ∈ 𝑉) & ⊢ (𝜑 → ¬ 𝐵 ∈ 𝐴) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐶 ∈ ℂ) & ⊢ (𝜑 → 𝐷 ∈ ℂ) & ⊢ (𝑘 = 𝐵 → 𝐶 = 𝐷) ⇒ ⊢ (𝜑 → ∏𝑘 ∈ (𝐴 ∪ {𝐵})𝐶 = (∏𝑘 ∈ 𝐴 𝐶 · 𝐷)) | ||
15-Aug-2024 | bj-charfundcALT 13426 | Alternate proof of bj-charfundc 13425. It was expected to be much shorter since it uses bj-charfun 13424 for the main part of the proof and the rest is basic computations, but these turn out to be lengthy, maybe because of the limited library of available lemmas. (Contributed by BJ, 15-Aug-2024.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝜑 → 𝐹 = (𝑥 ∈ 𝑋 ↦ if(𝑥 ∈ 𝐴, 1o, ∅))) & ⊢ (𝜑 → ∀𝑥 ∈ 𝑋 DECID 𝑥 ∈ 𝐴) ⇒ ⊢ (𝜑 → (𝐹:𝑋⟶2o ∧ (∀𝑥 ∈ (𝑋 ∩ 𝐴)(𝐹‘𝑥) = 1o ∧ ∀𝑥 ∈ (𝑋 ∖ 𝐴)(𝐹‘𝑥) = ∅))) | ||
15-Aug-2024 | bj-charfun 13424 | Properties of the characteristic function on the class 𝑋 of the class 𝐴. (Contributed by BJ, 15-Aug-2024.) |
⊢ (𝜑 → 𝐹 = (𝑥 ∈ 𝑋 ↦ if(𝑥 ∈ 𝐴, 1o, ∅))) ⇒ ⊢ (𝜑 → ((𝐹:𝑋⟶𝒫 1o ∧ (𝐹 ↾ ((𝑋 ∩ 𝐴) ∪ (𝑋 ∖ 𝐴))):((𝑋 ∩ 𝐴) ∪ (𝑋 ∖ 𝐴))⟶2o) ∧ (∀𝑥 ∈ (𝑋 ∩ 𝐴)(𝐹‘𝑥) = 1o ∧ ∀𝑥 ∈ (𝑋 ∖ 𝐴)(𝐹‘𝑥) = ∅))) | ||
15-Aug-2024 | fmelpw1o 13423 |
With a formula 𝜑 one can associate an element of
𝒫 1o, which
can therefore be thought of as the set of "truth values" (but
recall that
there are no other genuine truth values than ⊤ and ⊥, by
nndc 837, which translate to 1o and ∅
respectively by iftrue 3510
and iffalse 3513, giving pwtrufal 13611).
As proved in if0ab 13422, the associated element of 𝒫 1o is the extension, in 𝒫 1o, of the formula 𝜑. (Contributed by BJ, 15-Aug-2024.) |
⊢ if(𝜑, 1o, ∅) ∈ 𝒫 1o | ||
15-Aug-2024 | cnstab 8521 | Equality of complex numbers is stable. Stability here means ¬ ¬ 𝐴 = 𝐵 → 𝐴 = 𝐵 as defined at df-stab 817. This theorem for real numbers is Proposition 5.2 of [BauerHanson], p. 27. (Contributed by Jim Kingdon, 1-Aug-2023.) (Proof shortened by BJ, 15-Aug-2024.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → STAB 𝐴 = 𝐵) | ||
15-Aug-2024 | subap0d 8520 | Two numbers apart from each other have difference apart from zero. (Contributed by Jim Kingdon, 12-Aug-2021.) (Proof shortened by BJ, 15-Aug-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℂ) & ⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ (𝜑 → 𝐴 # 𝐵) ⇒ ⊢ (𝜑 → (𝐴 − 𝐵) # 0) | ||
15-Aug-2024 | ifexd 4445 | Existence of a conditional class (deduction form). (Contributed by BJ, 15-Aug-2024.) |
⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → 𝐵 ∈ 𝑊) ⇒ ⊢ (𝜑 → if(𝜑, 𝐴, 𝐵) ∈ V) | ||
15-Aug-2024 | ifelpwun 4444 | Existence of a conditional class, quantitative version (inference form). (Contributed by BJ, 15-Aug-2024.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ if(𝜑, 𝐴, 𝐵) ∈ 𝒫 (𝐴 ∪ 𝐵) | ||
15-Aug-2024 | ifelpwund 4443 | Existence of a conditional class, quantitative version (deduction form). (Contributed by BJ, 15-Aug-2024.) |
⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → 𝐵 ∈ 𝑊) ⇒ ⊢ (𝜑 → if(𝜑, 𝐴, 𝐵) ∈ 𝒫 (𝐴 ∪ 𝐵)) | ||
15-Aug-2024 | ifelpwung 4442 | Existence of a conditional class, quantitative version (closed form). (Contributed by BJ, 15-Aug-2024.) |
⊢ ((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊) → if(𝜑, 𝐴, 𝐵) ∈ 𝒫 (𝐴 ∪ 𝐵)) | ||
15-Aug-2024 | ifidss 3520 | A conditional class whose two alternatives are equal is included in that alternative. With excluded middle, we can prove it is equal to it. (Contributed by BJ, 15-Aug-2024.) |
⊢ if(𝜑, 𝐴, 𝐴) ⊆ 𝐴 | ||
15-Aug-2024 | ifssun 3519 | A conditional class is included in the union of its two alternatives. (Contributed by BJ, 15-Aug-2024.) |
⊢ if(𝜑, 𝐴, 𝐵) ⊆ (𝐴 ∪ 𝐵) | ||
12-Aug-2024 | exmidontriimlem2 7158 | Lemma for exmidontriim 7161. (Contributed by Jim Kingdon, 12-Aug-2024.) |
⊢ (𝜑 → 𝐵 ∈ On) & ⊢ (𝜑 → EXMID) & ⊢ (𝜑 → ∀𝑦 ∈ 𝐵 (𝐴 ∈ 𝑦 ∨ 𝐴 = 𝑦 ∨ 𝑦 ∈ 𝐴)) ⇒ ⊢ (𝜑 → (𝐴 ∈ 𝐵 ∨ ∀𝑦 ∈ 𝐵 𝑦 ∈ 𝐴)) | ||
12-Aug-2024 | exmidontriimlem1 7157 | Lemma for exmidontriim 7161. A variation of r19.30dc 2604. (Contributed by Jim Kingdon, 12-Aug-2024.) |
⊢ ((∀𝑥 ∈ 𝐴 (𝜑 ∨ 𝜓 ∨ 𝜒) ∧ EXMID) → (∃𝑥 ∈ 𝐴 𝜑 ∨ ∃𝑥 ∈ 𝐴 𝜓 ∨ ∀𝑥 ∈ 𝐴 𝜒)) | ||
11-Aug-2024 | nndc 837 |
Double negation of decidability of a formula. Intuitionistic logic
refutes the negation of decidability (but does not prove decidability) of
any formula.
This should not trick the reader into thinking that ¬ ¬ EXMID is provable in intuitionistic logic. Indeed, if we could quantify over formula metavariables, then generalizing nnexmid 836 over 𝜑 would give "⊢ ∀𝜑¬ ¬ DECID 𝜑", but EXMID is "∀𝜑DECID 𝜑", so proving ¬ ¬ EXMID would amount to proving "¬ ¬ ∀𝜑DECID 𝜑", which is not implied by the above theorem. Indeed, the converse of nnal 1629 does not hold. Since our system does not allow quantification over formula metavariables, we can reproduce this argument by representing formulas as subsets of 𝒫 1o, like we do in our definition of EXMID (df-exmid 4157): then, we can prove ∀𝑥 ∈ 𝒫 1o¬ ¬ DECID 𝑥 = 1o but we cannot prove ¬ ¬ ∀𝑥 ∈ 𝒫 1oDECID 𝑥 = 1o because the converse of nnral 2447 does not hold. Actually, ¬ ¬ EXMID is not provable in intuitionistic logic since intuitionistic logic has models satisfying ¬ EXMID and noncontradiction holds (pm3.24 683). (Contributed by BJ, 9-Oct-2019.) Add explanation on non-provability of ¬ ¬ EXMID. (Revised by BJ, 11-Aug-2024.) |
⊢ ¬ ¬ DECID 𝜑 | ||
10-Aug-2024 | exmidontriim 7161 | Excluded middle implies ordinal trichotomy. Lemma 10.4.1 of [HoTT], p. (varies). The proof follows the proof from the HoTT book fairly closely. (Contributed by Jim Kingdon, 10-Aug-2024.) |
⊢ (EXMID → ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥)) | ||
10-Aug-2024 | exmidontriimlem4 7160 | Lemma for exmidontriim 7161. The induction step for the induction on 𝐴. (Contributed by Jim Kingdon, 10-Aug-2024.) |
⊢ (𝜑 → 𝐴 ∈ On) & ⊢ (𝜑 → 𝐵 ∈ On) & ⊢ (𝜑 → EXMID) & ⊢ (𝜑 → ∀𝑧 ∈ 𝐴 ∀𝑦 ∈ On (𝑧 ∈ 𝑦 ∨ 𝑧 = 𝑦 ∨ 𝑦 ∈ 𝑧)) ⇒ ⊢ (𝜑 → (𝐴 ∈ 𝐵 ∨ 𝐴 = 𝐵 ∨ 𝐵 ∈ 𝐴)) | ||
10-Aug-2024 | exmidontriimlem3 7159 | Lemma for exmidontriim 7161. What we get to do based on induction on both 𝐴 and 𝐵. (Contributed by Jim Kingdon, 10-Aug-2024.) |
⊢ (𝜑 → 𝐴 ∈ On) & ⊢ (𝜑 → 𝐵 ∈ On) & ⊢ (𝜑 → EXMID) & ⊢ (𝜑 → ∀𝑧 ∈ 𝐴 ∀𝑦 ∈ On (𝑧 ∈ 𝑦 ∨ 𝑧 = 𝑦 ∨ 𝑦 ∈ 𝑧)) & ⊢ (𝜑 → ∀𝑦 ∈ 𝐵 (𝐴 ∈ 𝑦 ∨ 𝐴 = 𝑦 ∨ 𝑦 ∈ 𝐴)) ⇒ ⊢ (𝜑 → (𝐴 ∈ 𝐵 ∨ 𝐴 = 𝐵 ∨ 𝐵 ∈ 𝐴)) | ||
10-Aug-2024 | nnnninf2 7071 | Canonical embedding of suc ω into ℕ∞. (Contributed by BJ, 10-Aug-2024.) |
⊢ (𝑁 ∈ suc ω → (𝑖 ∈ ω ↦ if(𝑖 ∈ 𝑁, 1o, ∅)) ∈ ℕ∞) | ||
10-Aug-2024 | infnninf 7068 | The point at infinity in ℕ∞ is the constant sequence equal to 1o. Note that with our encoding of functions, that constant function can also be expressed as (ω × {1o}), as fconstmpt 4634 shows. (Contributed by Jim Kingdon, 14-Jul-2022.) Use maps-to notation. (Revised by BJ, 10-Aug-2024.) |
⊢ (𝑖 ∈ ω ↦ 1o) ∈ ℕ∞ | ||
9-Aug-2024 | ss1o0el1o 6858 | Reformulation of ss1o0el1 4159 using 1o instead of {∅}. (Contributed by BJ, 9-Aug-2024.) |
⊢ (𝐴 ⊆ 1o → (∅ ∈ 𝐴 ↔ 𝐴 = 1o)) | ||
9-Aug-2024 | pw1dc0el 6857 | Another equivalent of excluded middle, which is a mere reformulation of the definition. (Contributed by BJ, 9-Aug-2024.) |
⊢ (EXMID ↔ ∀𝑥 ∈ 𝒫 1oDECID ∅ ∈ 𝑥) | ||
9-Aug-2024 | ss1o0el1 4159 | A subclass of {∅} contains the empty set if and only if it equals {∅}. (Contributed by BJ and Jim Kingdon, 9-Aug-2024.) |
⊢ (𝐴 ⊆ {∅} → (∅ ∈ 𝐴 ↔ 𝐴 = {∅})) | ||
8-Aug-2024 | pw1dc1 6859 | If, in the set of truth values (the powerset of 1o), equality to 1o is decidable, then excluded middle holds (and conversely). (Contributed by BJ and Jim Kingdon, 8-Aug-2024.) |
⊢ (EXMID ↔ ∀𝑥 ∈ 𝒫 1oDECID 𝑥 = 1o) | ||
7-Aug-2024 | pw1fin 6856 | Excluded middle is equivalent to the power set of 1o being finite. (Contributed by SN and Jim Kingdon, 7-Aug-2024.) |
⊢ (EXMID ↔ 𝒫 1o ∈ Fin) | ||
7-Aug-2024 | elomssom 4565 | A natural number ordinal is, as a set, included in the set of natural number ordinals. (Contributed by NM, 21-Jun-1998.) Extract this result from the previous proof of elnn 4566. (Revised by BJ, 7-Aug-2024.) |
⊢ (𝐴 ∈ ω → 𝐴 ⊆ ω) | ||
6-Aug-2024 | bj-charfunbi 13428 |
In an ambient set 𝑋, if membership in 𝐴 is
stable, then it is
decidable if and only if 𝐴 has a characteristic function.
This characterization can be applied to singletons when the set 𝑋 has stable equality, which is the case as soon as it has a tight apartness relation. (Contributed by BJ, 6-Aug-2024.) |
⊢ (𝜑 → 𝑋 ∈ 𝑉) & ⊢ (𝜑 → ∀𝑥 ∈ 𝑋 STAB 𝑥 ∈ 𝐴) ⇒ ⊢ (𝜑 → (∀𝑥 ∈ 𝑋 DECID 𝑥 ∈ 𝐴 ↔ ∃𝑓 ∈ (2o ↑𝑚 𝑋)(∀𝑥 ∈ (𝑋 ∩ 𝐴)(𝑓‘𝑥) = 1o ∧ ∀𝑥 ∈ (𝑋 ∖ 𝐴)(𝑓‘𝑥) = ∅))) | ||
6-Aug-2024 | bj-charfunr 13427 |
If a class 𝐴 has a "weak"
characteristic function on a class 𝑋,
then negated membership in 𝐴 is decidable (in other words,
membership in 𝐴 is testable) in 𝑋.
The hypothesis imposes that 𝑋 be a set. As usual, it could be formulated as ⊢ (𝜑 → (𝐹:𝑋⟶ω ∧ ...)) to deal with general classes, but that extra generality would not make the theorem much more useful. The theorem would still hold if the codomain of 𝑓 were any class with testable equality to the point where (𝑋 ∖ 𝐴) is sent. (Contributed by BJ, 6-Aug-2024.) |
⊢ (𝜑 → ∃𝑓 ∈ (ω ↑𝑚 𝑋)(∀𝑥 ∈ (𝑋 ∩ 𝐴)(𝑓‘𝑥) ≠ ∅ ∧ ∀𝑥 ∈ (𝑋 ∖ 𝐴)(𝑓‘𝑥) = ∅)) ⇒ ⊢ (𝜑 → ∀𝑥 ∈ 𝑋 DECID ¬ 𝑥 ∈ 𝐴) | ||
6-Aug-2024 | bj-charfundc 13425 | Properties of the characteristic function on the class 𝑋 of the class 𝐴, provided membership in 𝐴 is decidable in 𝑋. (Contributed by BJ, 6-Aug-2024.) |
⊢ (𝜑 → 𝐹 = (𝑥 ∈ 𝑋 ↦ if(𝑥 ∈ 𝐴, 1o, ∅))) & ⊢ (𝜑 → ∀𝑥 ∈ 𝑋 DECID 𝑥 ∈ 𝐴) ⇒ ⊢ (𝜑 → (𝐹:𝑋⟶2o ∧ (∀𝑥 ∈ (𝑋 ∩ 𝐴)(𝐹‘𝑥) = 1o ∧ ∀𝑥 ∈ (𝑋 ∖ 𝐴)(𝐹‘𝑥) = ∅))) | ||
6-Aug-2024 | prodssdc 11490 | Change the index set to a subset in an upper integer product. (Contributed by Scott Fenton, 11-Dec-2017.) (Revised by Jim Kingdon, 6-Aug-2024.) |
⊢ (𝜑 → 𝐴 ⊆ 𝐵) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐶 ∈ ℂ) & ⊢ (𝜑 → ∃𝑛 ∈ (ℤ≥‘𝑀)∃𝑦(𝑦 # 0 ∧ seq𝑛( · , (𝑘 ∈ (ℤ≥‘𝑀) ↦ if(𝑘 ∈ 𝐵, 𝐶, 1))) ⇝ 𝑦)) & ⊢ (𝜑 → ∀𝑗 ∈ (ℤ≥‘𝑀)DECID 𝑗 ∈ 𝐴) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (𝐵 ∖ 𝐴)) → 𝐶 = 1) & ⊢ (𝜑 → 𝐵 ⊆ (ℤ≥‘𝑀)) & ⊢ (𝜑 → ∀𝑗 ∈ (ℤ≥‘𝑀)DECID 𝑗 ∈ 𝐵) ⇒ ⊢ (𝜑 → ∏𝑘 ∈ 𝐴 𝐶 = ∏𝑘 ∈ 𝐵 𝐶) | ||
5-Aug-2024 | fnmptd 13421 | The maps-to notation defines a function with domain (deduction form). (Contributed by BJ, 5-Aug-2024.) |
⊢ (𝜑 → 𝐹 = (𝑥 ∈ 𝐴 ↦ 𝐵)) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝐵 ∈ 𝑉) ⇒ ⊢ (𝜑 → 𝐹 Fn 𝐴) | ||
5-Aug-2024 | funmptd 13420 |
The maps-to notation defines a function (deduction form).
Note: one should similarly prove a deduction form of funopab4 5208, then prove funmptd 13420 from it, and then prove funmpt 5209 from that: this would reduce global proof length. (Contributed by BJ, 5-Aug-2024.) |
⊢ (𝜑 → 𝐹 = (𝑥 ∈ 𝐴 ↦ 𝐵)) ⇒ ⊢ (𝜑 → Fun 𝐹) | ||
5-Aug-2024 | 2ssom 13419 | The ordinal 2 is included in the set of natural number ordinals. (Contributed by BJ, 5-Aug-2024.) |
⊢ 2o ⊆ ω | ||
5-Aug-2024 | bj-dcfal 13369 | The false truth value is decidable. (Contributed by BJ, 5-Aug-2024.) |
⊢ DECID ⊥ | ||
5-Aug-2024 | bj-dctru 13367 | The true truth value is decidable. (Contributed by BJ, 5-Aug-2024.) |
⊢ DECID ⊤ | ||
5-Aug-2024 | bj-stfal 13359 | The false truth value is stable. (Contributed by BJ, 5-Aug-2024.) |
⊢ STAB ⊥ | ||
5-Aug-2024 | bj-sttru 13357 | The true truth value is stable. (Contributed by BJ, 5-Aug-2024.) |
⊢ STAB ⊤ | ||
5-Aug-2024 | prod1dc 11487 | Any product of one over a valid set is one. (Contributed by Scott Fenton, 7-Dec-2017.) (Revised by Jim Kingdon, 5-Aug-2024.) |
⊢ (((𝑀 ∈ ℤ ∧ 𝐴 ⊆ (ℤ≥‘𝑀) ∧ ∀𝑗 ∈ (ℤ≥‘𝑀)DECID 𝑗 ∈ 𝐴) ∨ 𝐴 ∈ Fin) → ∏𝑘 ∈ 𝐴 1 = 1) | ||
2-Aug-2024 | onntri52 7180 | Double negated ordinal trichotomy. (Contributed by James E. Hanson and Jim Kingdon, 2-Aug-2024.) |
⊢ (¬ ¬ EXMID → ¬ ¬ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥)) | ||
2-Aug-2024 | onntri24 7178 | Double negated ordinal trichotomy. (Contributed by James E. Hanson and Jim Kingdon, 2-Aug-2024.) |
⊢ (¬ ¬ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥) → ∀𝑥 ∈ On ∀𝑦 ∈ On ¬ ¬ (𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥)) | ||
2-Aug-2024 | onntri45 7177 | Double negated ordinal trichotomy. (Contributed by James E. Hanson and Jim Kingdon, 2-Aug-2024.) |
⊢ (∀𝑥 ∈ On ∀𝑦 ∈ On ¬ ¬ (𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥) → ¬ ¬ EXMID) | ||
2-Aug-2024 | onntri51 7176 | Double negated ordinal trichotomy. (Contributed by James E. Hanson and Jim Kingdon, 2-Aug-2024.) |
⊢ (¬ ¬ EXMID → ¬ ¬ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥)) | ||
2-Aug-2024 | onntri13 7174 | Double negated ordinal trichotomy. (Contributed by James E. Hanson and Jim Kingdon, 2-Aug-2024.) |
⊢ (¬ ¬ ∀𝑥 ∈ On ∀𝑦 ∈ On (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥) → ∀𝑥 ∈ On ∀𝑦 ∈ On ¬ ¬ (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥)) | ||
2-Aug-2024 | onntri35 7173 |
Double negated ordinal trichotomy.
There are five equivalent statements: (1) ¬ ¬ ∀𝑥 ∈ On∀𝑦 ∈ On(𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥), (2) ¬ ¬ ∀𝑥 ∈ On∀𝑦 ∈ On(𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥), (3) ∀𝑥 ∈ On∀𝑦 ∈ On¬ ¬ (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥), (4) ∀𝑥 ∈ On∀𝑦 ∈ On¬ ¬ (𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥), and (5) ¬ ¬ EXMID. That these are all equivalent is expressed by (1) implies (3) (onntri13 7174), (3) implies (5) (onntri35 7173), (5) implies (1) (onntri51 7176), (2) implies (4) (onntri24 7178), (4) implies (5) (onntri45 7177), and (5) implies (2) (onntri52 7180). Another way of stating this is that EXMID is equivalent to trichotomy, either the 𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥 or the 𝑥 ⊆ 𝑦 ∨ 𝑦 ⊆ 𝑥 form, as shown in exmidontri 7175 and exmidontri2or 7179, respectively. Thus ¬ ¬ EXMID is equivalent to (1) or (2). In addition, ¬ ¬ EXMID is equivalent to (3) by onntri3or 7181 and (4) by onntri2or 7182. (Contributed by James E. Hanson and Jim Kingdon, 2-Aug-2024.) |
⊢ (∀𝑥 ∈ On ∀𝑦 ∈ On ¬ ¬ (𝑥 ∈ 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 ∈ 𝑥) → ¬ ¬ EXMID) | ||
1-Aug-2024 | nnral 2447 | The double negation of a universal quantification implies the universal quantification of the double negation. Restricted quantifier version of nnal 1629. (Contributed by Jim Kingdon, 1-Aug-2024.) |
⊢ (¬ ¬ ∀𝑥 ∈ 𝐴 𝜑 → ∀𝑥 ∈ 𝐴 ¬ ¬ 𝜑) | ||
31-Jul-2024 | 3nsssucpw1 7172 | Negated excluded middle implies that 3o is not a subset of the successor of the power set of 1o. (Contributed by James E. Hanson and Jim Kingdon, 31-Jul-2024.) |
⊢ (¬ EXMID → ¬ 3o ⊆ suc 𝒫 1o) | ||
31-Jul-2024 | sucpw1nss3 7171 | Negated excluded middle implies that the successor of the power set of 1o is not a subset of 3o. (Contributed by James E. Hanson and Jim Kingdon, 31-Jul-2024.) |
⊢ (¬ EXMID → ¬ suc 𝒫 1o ⊆ 3o) | ||
30-Jul-2024 | 3nelsucpw1 7170 | Three is not an element of the successor of the power set of 1o. (Contributed by James E. Hanson and Jim Kingdon, 30-Jul-2024.) |
⊢ ¬ 3o ∈ suc 𝒫 1o | ||
30-Jul-2024 | sucpw1nel3 7169 | The successor of the power set of 1o is not an element of 3o. (Contributed by James E. Hanson and Jim Kingdon, 30-Jul-2024.) |
⊢ ¬ suc 𝒫 1o ∈ 3o | ||
30-Jul-2024 | sucpw1ne3 7168 | Negated excluded middle implies that the successor of the power set of 1o is not three . (Contributed by James E. Hanson and Jim Kingdon, 30-Jul-2024.) |
⊢ (¬ EXMID → suc 𝒫 1o ≠ 3o) | ||
30-Jul-2024 | pw1nel3 7167 | Negated excluded middle implies that the power set of 1o is not an element of 3o. (Contributed by James E. Hanson and Jim Kingdon, 30-Jul-2024.) |
⊢ (¬ EXMID → ¬ 𝒫 1o ∈ 3o) | ||
30-Jul-2024 | pw1ne3 7166 | The power set of 1o is not three. (Contributed by James E. Hanson and Jim Kingdon, 30-Jul-2024.) |
⊢ 𝒫 1o ≠ 3o | ||
30-Jul-2024 | pw1ne1 7165 | The power set of 1o is not one. (Contributed by Jim Kingdon, 30-Jul-2024.) |
⊢ 𝒫 1o ≠ 1o | ||
30-Jul-2024 | pw1ne0 7164 | The power set of 1o is not zero. (Contributed by Jim Kingdon, 30-Jul-2024.) |
⊢ 𝒫 1o ≠ ∅ | ||
29-Jul-2024 | pw1on 7162 | The power set of 1o is an ordinal. (Contributed by Jim Kingdon, 29-Jul-2024.) |
⊢ 𝒫 1o ∈ On | ||
28-Jul-2024 | exmidpweq 6855 | Excluded middle is equivalent to the power set of 1o being 2o. (Contributed by Jim Kingdon, 28-Jul-2024.) |
⊢ (EXMID ↔ 𝒫 1o = 2o) | ||
27-Jul-2024 | dcapnconstALT 13674 | Decidability of real number apartness implies the existence of a certain non-constant function from real numbers to integers. A proof of dcapnconst 13673 by means of dceqnconst 13672. (Contributed by Jim Kingdon, 27-Jul-2024.) (New usage is discouraged.) (Proof modification is discouraged.) |
⊢ (∀𝑥 ∈ ℝ DECID 𝑥 # 0 → ∃𝑓(𝑓:ℝ⟶ℤ ∧ (𝑓‘0) = 0 ∧ ∀𝑥 ∈ ℝ+ (𝑓‘𝑥) ≠ 0)) | ||
27-Jul-2024 | reap0 13671 | Real number trichotomy is equivalent to decidability of apartness from zero. (Contributed by Jim Kingdon, 27-Jul-2024.) |
⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 < 𝑥) ↔ ∀𝑧 ∈ ℝ DECID 𝑧 # 0) | ||
26-Jul-2024 | nconstwlpolemgt0 13676 | Lemma for nconstwlpo 13678. If one of the terms of series is positive, so is the sum. (Contributed by Jim Kingdon, 26-Jul-2024.) |
⊢ (𝜑 → 𝐺:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐺‘𝑖)) & ⊢ (𝜑 → ∃𝑥 ∈ ℕ (𝐺‘𝑥) = 1) ⇒ ⊢ (𝜑 → 0 < 𝐴) | ||
26-Jul-2024 | nconstwlpolem0 13675 | Lemma for nconstwlpo 13678. If all the terms of the series are zero, so is their sum. (Contributed by Jim Kingdon, 26-Jul-2024.) |
⊢ (𝜑 → 𝐺:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐺‘𝑖)) & ⊢ (𝜑 → ∀𝑥 ∈ ℕ (𝐺‘𝑥) = 0) ⇒ ⊢ (𝜑 → 𝐴 = 0) | ||
24-Jul-2024 | tridceq 13669 | Real trichotomy implies decidability of real number equality. Or in other words, analytic LPO implies analytic WLPO (see trilpo 13656 and redcwlpo 13668). Thus, this is an analytic analogue to lpowlpo 7112. (Contributed by Jim Kingdon, 24-Jul-2024.) |
⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 < 𝑥) → ∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ DECID 𝑥 = 𝑦) | ||
24-Jul-2024 | iswomni0 13664 | Weak omniscience stated in terms of equality with 0. Like iswomninn 13663 but with zero in place of one. (Contributed by Jim Kingdon, 24-Jul-2024.) |
⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ WOmni ↔ ∀𝑓 ∈ ({0, 1} ↑𝑚 𝐴)DECID ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) = 0)) | ||
24-Jul-2024 | lpowlpo 7112 | LPO implies WLPO. Easy corollary of the more general omniwomnimkv 7111. There is an analogue in terms of analytic omniscience principles at tridceq 13669. (Contributed by Jim Kingdon, 24-Jul-2024.) |
⊢ (ω ∈ Omni → ω ∈ WOmni) | ||
23-Jul-2024 | nconstwlpolem 13677 | Lemma for nconstwlpo 13678. (Contributed by Jim Kingdon, 23-Jul-2024.) |
⊢ (𝜑 → 𝐹:ℝ⟶ℤ) & ⊢ (𝜑 → (𝐹‘0) = 0) & ⊢ ((𝜑 ∧ 𝑥 ∈ ℝ+) → (𝐹‘𝑥) ≠ 0) & ⊢ (𝜑 → 𝐺:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐺‘𝑖)) ⇒ ⊢ (𝜑 → (∀𝑦 ∈ ℕ (𝐺‘𝑦) = 0 ∨ ¬ ∀𝑦 ∈ ℕ (𝐺‘𝑦) = 0)) | ||
23-Jul-2024 | dceqnconst 13672 | Decidability of real number equality implies the existence of a certain non-constant function from real numbers to integers. Variation of Exercise 11.6(i) of [HoTT], p. (varies). See redcwlpo 13668 for more discussion of decidability of real number equality. (Contributed by BJ and Jim Kingdon, 24-Jun-2024.) (Revised by Jim Kingdon, 23-Jul-2024.) |
⊢ (∀𝑥 ∈ ℝ DECID 𝑥 = 0 → ∃𝑓(𝑓:ℝ⟶ℤ ∧ (𝑓‘0) = 0 ∧ ∀𝑥 ∈ ℝ+ (𝑓‘𝑥) ≠ 0)) | ||
23-Jul-2024 | redc0 13670 | Two ways to express decidability of real number equality. (Contributed by Jim Kingdon, 23-Jul-2024.) |
⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ DECID 𝑥 = 𝑦 ↔ ∀𝑧 ∈ ℝ DECID 𝑧 = 0) | ||
23-Jul-2024 | canth 5779 | No set 𝐴 is equinumerous to its power set (Cantor's theorem), i.e., no function can map 𝐴 onto its power set. Compare Theorem 6B(b) of [Enderton] p. 132. (Use nex 1480 if you want the form ¬ ∃𝑓𝑓:𝐴–onto→𝒫 𝐴.) (Contributed by NM, 7-Aug-1994.) (Revised by Noah R Kingdon, 23-Jul-2024.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ¬ 𝐹:𝐴–onto→𝒫 𝐴 | ||
22-Jul-2024 | nconstwlpo 13678 | Existence of a certain non-constant function from reals to integers implies ω ∈ WOmni (the Weak Limited Principle of Omniscience or WLPO). Based on Exercise 11.6(ii) of [HoTT], p. (varies). (Contributed by BJ and Jim Kingdon, 22-Jul-2024.) |
⊢ (𝜑 → 𝐹:ℝ⟶ℤ) & ⊢ (𝜑 → (𝐹‘0) = 0) & ⊢ ((𝜑 ∧ 𝑥 ∈ ℝ+) → (𝐹‘𝑥) ≠ 0) ⇒ ⊢ (𝜑 → ω ∈ WOmni) | ||
15-Jul-2024 | fprodseq 11484 | The value of a product over a nonempty finite set. (Contributed by Scott Fenton, 6-Dec-2017.) (Revised by Jim Kingdon, 15-Jul-2024.) |
⊢ (𝑘 = (𝐹‘𝑛) → 𝐵 = 𝐶) & ⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝐹:(1...𝑀)–1-1-onto→𝐴) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐵 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑛 ∈ (1...𝑀)) → (𝐺‘𝑛) = 𝐶) ⇒ ⊢ (𝜑 → ∏𝑘 ∈ 𝐴 𝐵 = (seq1( · , (𝑛 ∈ ℕ ↦ if(𝑛 ≤ 𝑀, (𝐺‘𝑛), 1)))‘𝑀)) | ||
14-Jul-2024 | rexbid2 2462 | Formula-building rule for restricted existential quantifier (deduction form). (Contributed by BJ, 14-Jul-2024.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → ((𝑥 ∈ 𝐴 ∧ 𝜓) ↔ (𝑥 ∈ 𝐵 ∧ 𝜒))) ⇒ ⊢ (𝜑 → (∃𝑥 ∈ 𝐴 𝜓 ↔ ∃𝑥 ∈ 𝐵 𝜒)) | ||
14-Jul-2024 | ralbid2 2461 | Formula-building rule for restricted universal quantifier (deduction form). (Contributed by BJ, 14-Jul-2024.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → ((𝑥 ∈ 𝐴 → 𝜓) ↔ (𝑥 ∈ 𝐵 → 𝜒))) ⇒ ⊢ (𝜑 → (∀𝑥 ∈ 𝐴 𝜓 ↔ ∀𝑥 ∈ 𝐵 𝜒)) | ||
12-Jul-2024 | 2irrexpqap 13337 | There exist real numbers 𝑎 and 𝑏 which are irrational (in the sense of being apart from any rational number) such that (𝑎↑𝑏) is rational. Statement in the Metamath book, section 1.1.5, footnote 27 on page 17, and the "constructive proof" for theorem 1.2 of [Bauer], p. 483. This is a constructive proof because it is based on two explicitly named irrational numbers (√‘2) and (2 logb 9), see sqrt2irrap 12059, 2logb9irrap 13336 and sqrt2cxp2logb9e3 13334. Therefore, this proof is acceptable/usable in intuitionistic logic. (Contributed by Jim Kingdon, 12-Jul-2024.) |
⊢ ∃𝑎 ∈ ℝ ∃𝑏 ∈ ℝ (∀𝑝 ∈ ℚ 𝑎 # 𝑝 ∧ ∀𝑞 ∈ ℚ 𝑏 # 𝑞 ∧ (𝑎↑𝑐𝑏) ∈ ℚ) | ||
12-Jul-2024 | 2logb9irrap 13336 | Example for logbgcd1irrap 13329. The logarithm of nine to base two is irrational (in the sense of being apart from any rational number). (Contributed by Jim Kingdon, 12-Jul-2024.) |
⊢ (𝑄 ∈ ℚ → (2 logb 9) # 𝑄) | ||
11-Jul-2024 | logbgcd1irraplemexp 13327 | Lemma for logbgcd1irrap 13329. Apartness of 𝑋↑𝑁 and 𝐵↑𝑀. (Contributed by Jim Kingdon, 11-Jul-2024.) |
⊢ (𝜑 → 𝑋 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐵 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → (𝑋 gcd 𝐵) = 1) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑁 ∈ ℕ) ⇒ ⊢ (𝜑 → (𝑋↑𝑁) # (𝐵↑𝑀)) | ||
11-Jul-2024 | reapef 13141 | Apartness and the exponential function for reals. (Contributed by Jim Kingdon, 11-Jul-2024.) |
⊢ ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴 # 𝐵 ↔ (exp‘𝐴) # (exp‘𝐵))) | ||
10-Jul-2024 | apcxp2 13300 | Apartness and real exponentiation. (Contributed by Jim Kingdon, 10-Jul-2024.) |
⊢ (((𝐴 ∈ ℝ+ ∧ 𝐴 # 1) ∧ (𝐵 ∈ ℝ ∧ 𝐶 ∈ ℝ)) → (𝐵 # 𝐶 ↔ (𝐴↑𝑐𝐵) # (𝐴↑𝑐𝐶))) | ||
9-Jul-2024 | logbgcd1irraplemap 13328 | Lemma for logbgcd1irrap 13329. The result, with the rational number expressed as numerator and denominator. (Contributed by Jim Kingdon, 9-Jul-2024.) |
⊢ (𝜑 → 𝑋 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → 𝐵 ∈ (ℤ≥‘2)) & ⊢ (𝜑 → (𝑋 gcd 𝐵) = 1) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝑁 ∈ ℕ) ⇒ ⊢ (𝜑 → (𝐵 logb 𝑋) # (𝑀 / 𝑁)) | ||
9-Jul-2024 | apexp1 10596 | Exponentiation and apartness. (Contributed by Jim Kingdon, 9-Jul-2024.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝑁 ∈ ℕ) → ((𝐴↑𝑁) # (𝐵↑𝑁) → 𝐴 # 𝐵)) | ||
5-Jul-2024 | logrpap0 13240 | The logarithm is apart from 0 if its argument is apart from 1. (Contributed by Jim Kingdon, 5-Jul-2024.) |
⊢ ((𝐴 ∈ ℝ+ ∧ 𝐴 # 1) → (log‘𝐴) # 0) | ||
3-Jul-2024 | rplogbval 13304 | Define the value of the logb function, the logarithm generalized to an arbitrary base, when used as infix. Most Metamath statements select variables in order of their use, but to make the order clearer we use "B" for base and "X" for the argument of the logarithm function here. (Contributed by David A. Wheeler, 21-Jan-2017.) (Revised by Jim Kingdon, 3-Jul-2024.) |
⊢ ((𝐵 ∈ ℝ+ ∧ 𝐵 # 1 ∧ 𝑋 ∈ ℝ+) → (𝐵 logb 𝑋) = ((log‘𝑋) / (log‘𝐵))) | ||
3-Jul-2024 | logrpap0d 13241 | Deduction form of logrpap0 13240. (Contributed by Jim Kingdon, 3-Jul-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐴 # 1) ⇒ ⊢ (𝜑 → (log‘𝐴) # 0) | ||
3-Jul-2024 | logrpap0b 13239 | The logarithm is apart from 0 if and only if its argument is apart from 1. (Contributed by Jim Kingdon, 3-Jul-2024.) |
⊢ (𝐴 ∈ ℝ+ → (𝐴 # 1 ↔ (log‘𝐴) # 0)) | ||
28-Jun-2024 | 2o01f 13610 | Mapping zero and one between ω and ℕ0 style integers. (Contributed by Jim Kingdon, 28-Jun-2024.) |
⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) ⇒ ⊢ (𝐺 ↾ 2o):2o⟶{0, 1} | ||
28-Jun-2024 | 012of 13609 | Mapping zero and one between ℕ0 and ω style integers. (Contributed by Jim Kingdon, 28-Jun-2024.) |
⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) ⇒ ⊢ (◡𝐺 ↾ {0, 1}):{0, 1}⟶2o | ||
27-Jun-2024 | iooreen 13648 | An open interval is equinumerous to the real numbers. (Contributed by Jim Kingdon, 27-Jun-2024.) |
⊢ (0(,)1) ≈ ℝ | ||
27-Jun-2024 | iooref1o 13647 | A one-to-one mapping from the real numbers onto the open unit interval. (Contributed by Jim Kingdon, 27-Jun-2024.) |
⊢ 𝐹 = (𝑥 ∈ ℝ ↦ (1 / (1 + (exp‘𝑥)))) ⇒ ⊢ 𝐹:ℝ–1-1-onto→(0(,)1) | ||
25-Jun-2024 | neapmkvlem 13679 | Lemma for neapmkv 13680. The result, with a few hypotheses broken out for convenience. (Contributed by Jim Kingdon, 25-Jun-2024.) |
⊢ (𝜑 → 𝐹:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐹‘𝑖)) & ⊢ ((𝜑 ∧ 𝐴 ≠ 1) → 𝐴 # 1) ⇒ ⊢ (𝜑 → (¬ ∀𝑥 ∈ ℕ (𝐹‘𝑥) = 1 → ∃𝑥 ∈ ℕ (𝐹‘𝑥) = 0)) | ||
25-Jun-2024 | ismkvnn 13666 | The predicate of being Markov stated in terms of set exponentiation. (Contributed by Jim Kingdon, 25-Jun-2024.) |
⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ Markov ↔ ∀𝑓 ∈ ({0, 1} ↑𝑚 𝐴)(¬ ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) = 1 → ∃𝑥 ∈ 𝐴 (𝑓‘𝑥) = 0))) | ||
25-Jun-2024 | ismkvnnlem 13665 | Lemma for ismkvnn 13666. The result, with a hypothesis to give a name to an expression for convenience. (Contributed by Jim Kingdon, 25-Jun-2024.) |
⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) ⇒ ⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ Markov ↔ ∀𝑓 ∈ ({0, 1} ↑𝑚 𝐴)(¬ ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) = 1 → ∃𝑥 ∈ 𝐴 (𝑓‘𝑥) = 0))) | ||
25-Jun-2024 | enmkvlem 7105 | Lemma for enmkv 7106. One direction of the biconditional. (Contributed by Jim Kingdon, 25-Jun-2024.) |
⊢ (𝐴 ≈ 𝐵 → (𝐴 ∈ Markov → 𝐵 ∈ Markov)) | ||
24-Jun-2024 | neapmkv 13680 | If negated equality for real numbers implies apartness, Markov's Principle follows. Exercise 11.10 of [HoTT], p. (varies). (Contributed by Jim Kingdon, 24-Jun-2024.) |
⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 ≠ 𝑦 → 𝑥 # 𝑦) → ω ∈ Markov) | ||
24-Jun-2024 | dcapnconst 13673 |
Decidability of real number apartness implies the existence of a certain
non-constant function from real numbers to integers. Variation of
Exercise 11.6(i) of [HoTT], p. (varies).
See trilpo 13656 for more
discussion of decidability of real number apartness.
This is a weaker form of dceqnconst 13672 and in fact this theorem can be proved using dceqnconst 13672 as shown at dcapnconstALT 13674. (Contributed by BJ and Jim Kingdon, 24-Jun-2024.) |
⊢ (∀𝑥 ∈ ℝ DECID 𝑥 # 0 → ∃𝑓(𝑓:ℝ⟶ℤ ∧ (𝑓‘0) = 0 ∧ ∀𝑥 ∈ ℝ+ (𝑓‘𝑥) ≠ 0)) | ||
24-Jun-2024 | enmkv 7106 | Being Markov is invariant with respect to equinumerosity. For example, this means that we can express the Markov's Principle as either ω ∈ Markov or ℕ0 ∈ Markov. The former is a better match to conventional notation in the sense that df2o3 6378 says that 2o = {∅, 1o} whereas the corresponding relationship does not exist between 2 and {0, 1}. (Contributed by Jim Kingdon, 24-Jun-2024.) |
⊢ (𝐴 ≈ 𝐵 → (𝐴 ∈ Markov ↔ 𝐵 ∈ Markov)) | ||
21-Jun-2024 | redcwlpolemeq1 13667 | Lemma for redcwlpo 13668. A biconditionalized version of trilpolemeq1 13653. (Contributed by Jim Kingdon, 21-Jun-2024.) |
⊢ (𝜑 → 𝐹:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐹‘𝑖)) ⇒ ⊢ (𝜑 → (𝐴 = 1 ↔ ∀𝑥 ∈ ℕ (𝐹‘𝑥) = 1)) | ||
20-Jun-2024 | redcwlpo 13668 |
Decidability of real number equality implies the Weak Limited Principle
of Omniscience (WLPO). We expect that we'd need some form of countable
choice to prove the converse.
Here's the outline of the proof. Given an infinite sequence F of zeroes and ones, we need to show the sequence is all ones or it is not. Construct a real number A whose representation in base two consists of a zero, a decimal point, and then the numbers of the sequence. This real number will equal one if and only if the sequence is all ones (redcwlpolemeq1 13667). Therefore decidability of real number equality would imply decidability of whether the sequence is all ones. Because of this theorem, decidability of real number equality is sometimes called "analytic WLPO". WLPO is known to not be provable in IZF (and most constructive foundations), so this theorem establishes that we will be unable to prove an analogue to qdceq 10150 for real numbers. (Contributed by Jim Kingdon, 20-Jun-2024.) |
⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ DECID 𝑥 = 𝑦 → ω ∈ WOmni) | ||
20-Jun-2024 | iswomninn 13663 | Weak omniscience stated in terms of natural numbers. Similar to iswomnimap 7110 but it will sometimes be more convenient to use 0 and 1 rather than ∅ and 1o. (Contributed by Jim Kingdon, 20-Jun-2024.) |
⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ WOmni ↔ ∀𝑓 ∈ ({0, 1} ↑𝑚 𝐴)DECID ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) = 1)) | ||
20-Jun-2024 | iswomninnlem 13662 | Lemma for iswomnimap 7110. The result, with a hypothesis for convenience. (Contributed by Jim Kingdon, 20-Jun-2024.) |
⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) ⇒ ⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ WOmni ↔ ∀𝑓 ∈ ({0, 1} ↑𝑚 𝐴)DECID ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) = 1)) | ||
20-Jun-2024 | enwomni 7114 | Weak omniscience is invariant with respect to equinumerosity. For example, this means that we can express the Weak Limited Principle of Omniscience as either ω ∈ WOmni or ℕ0 ∈ WOmni. The former is a better match to conventional notation in the sense that df2o3 6378 says that 2o = {∅, 1o} whereas the corresponding relationship does not exist between 2 and {0, 1}. (Contributed by Jim Kingdon, 20-Jun-2024.) |
⊢ (𝐴 ≈ 𝐵 → (𝐴 ∈ WOmni ↔ 𝐵 ∈ WOmni)) | ||
20-Jun-2024 | enwomnilem 7113 | Lemma for enwomni 7114. One direction of the biconditional. (Contributed by Jim Kingdon, 20-Jun-2024.) |
⊢ (𝐴 ≈ 𝐵 → (𝐴 ∈ WOmni → 𝐵 ∈ WOmni)) | ||
19-Jun-2024 | rpabscxpbnd 13301 | Bound on the absolute value of a complex power. (Contributed by Mario Carneiro, 15-Sep-2014.) (Revised by Jim Kingdon, 19-Jun-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ+) & ⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ (𝜑 → 0 < (ℜ‘𝐵)) & ⊢ (𝜑 → 𝑀 ∈ ℝ) & ⊢ (𝜑 → (abs‘𝐴) ≤ 𝑀) ⇒ ⊢ (𝜑 → (abs‘(𝐴↑𝑐𝐵)) ≤ ((𝑀↑𝑐(ℜ‘𝐵)) · (exp‘((abs‘𝐵) · π)))) | ||
16-Jun-2024 | rpcxpsqrt 13284 | The exponential function with exponent 1 / 2 exactly matches the square root function, and thus serves as a suitable generalization to other 𝑛-th roots and irrational roots. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 16-Jun-2024.) |
⊢ (𝐴 ∈ ℝ+ → (𝐴↑𝑐(1 / 2)) = (√‘𝐴)) | ||
13-Jun-2024 | rpcxpadd 13268 | Sum of exponents law for complex exponentiation. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 13-Jun-2024.) |
⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) → (𝐴↑𝑐(𝐵 + 𝐶)) = ((𝐴↑𝑐𝐵) · (𝐴↑𝑐𝐶))) | ||
12-Jun-2024 | cxpap0 13267 | Complex exponentiation is apart from zero. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 12-Jun-2024.) |
⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ) → (𝐴↑𝑐𝐵) # 0) | ||
12-Jun-2024 | rpcncxpcl 13265 | Closure of the complex power function. (Contributed by Jim Kingdon, 12-Jun-2024.) |
⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ) → (𝐴↑𝑐𝐵) ∈ ℂ) | ||
12-Jun-2024 | rpcxp0 13261 | Value of the complex power function when the second argument is zero. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 12-Jun-2024.) |
⊢ (𝐴 ∈ ℝ+ → (𝐴↑𝑐0) = 1) | ||
12-Jun-2024 | cxpexpnn 13259 | Relate the complex power function to the integer power function. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 12-Jun-2024.) |
⊢ ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℤ) → (𝐴↑𝑐𝐵) = (𝐴↑𝐵)) | ||
12-Jun-2024 | cxpexprp 13258 | Relate the complex power function to the integer power function. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 12-Jun-2024.) |
⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℤ) → (𝐴↑𝑐𝐵) = (𝐴↑𝐵)) | ||
12-Jun-2024 | rpcxpef 13257 | Value of the complex power function. (Contributed by Mario Carneiro, 2-Aug-2014.) (Revised by Jim Kingdon, 12-Jun-2024.) |
⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℂ) → (𝐴↑𝑐𝐵) = (exp‘(𝐵 · (log‘𝐴)))) | ||
12-Jun-2024 | df-rpcxp 13222 | Define the power function on complex numbers. Because df-relog 13221 is only defined on positive reals, this definition only allows for a base which is a positive real. (Contributed by Jim Kingdon, 12-Jun-2024.) |
⊢ ↑𝑐 = (𝑥 ∈ ℝ+, 𝑦 ∈ ℂ ↦ (exp‘(𝑦 · (log‘𝑥)))) | ||
10-Jun-2024 | trirec0xor 13658 |
Version of trirec0 13657 with exclusive-or.
The definition of a discrete field is sometimes stated in terms of exclusive-or but as proved here, this is equivalent to inclusive-or because the two disjuncts cannot be simultaneously true. (Contributed by Jim Kingdon, 10-Jun-2024.) |
⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 < 𝑥) ↔ ∀𝑥 ∈ ℝ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ⊻ 𝑥 = 0)) | ||
10-Jun-2024 | trirec0 13657 |
Every real number having a reciprocal or equaling zero is equivalent to
real number trichotomy.
This is the key part of the definition of what is known as a discrete field, so "the real numbers are a discrete field" can be taken as an equivalent way to state real trichotomy (see further discussion at trilpo 13656). (Contributed by Jim Kingdon, 10-Jun-2024.) |
⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 < 𝑥) ↔ ∀𝑥 ∈ ℝ (∃𝑧 ∈ ℝ (𝑥 · 𝑧) = 1 ∨ 𝑥 = 0)) | ||
9-Jun-2024 | omniwomnimkv 7111 | A set is omniscient if and only if it is weakly omniscient and Markov. The case 𝐴 = ω says that LPO ↔ WLPO ∧ MP which is a remark following Definition 2.5 of [Pierik], p. 9. (Contributed by Jim Kingdon, 9-Jun-2024.) |
⊢ (𝐴 ∈ Omni ↔ (𝐴 ∈ WOmni ∧ 𝐴 ∈ Markov)) | ||
9-Jun-2024 | iswomnimap 7110 | The predicate of being weakly omniscient stated in terms of set exponentiation. (Contributed by Jim Kingdon, 9-Jun-2024.) |
⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ WOmni ↔ ∀𝑓 ∈ (2o ↑𝑚 𝐴)DECID ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) = 1o)) | ||
9-Jun-2024 | iswomni 7109 | The predicate of being weakly omniscient. (Contributed by Jim Kingdon, 9-Jun-2024.) |
⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ WOmni ↔ ∀𝑓(𝑓:𝐴⟶2o → DECID ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) = 1o))) | ||
9-Jun-2024 | df-womni 7108 |
A weakly omniscient set is one where we can decide whether a predicate
(here represented by a function 𝑓) holds (is equal to 1o) for
all elements or not. Generalization of definition 2.4 of [Pierik],
p. 9.
In particular, ω ∈ WOmni is known as the Weak Limited Principle of Omniscience (WLPO). The term WLPO is common in the literature; there appears to be no widespread term for what we are calling a weakly omniscient set. (Contributed by Jim Kingdon, 9-Jun-2024.) |
⊢ WOmni = {𝑦 ∣ ∀𝑓(𝑓:𝑦⟶2o → DECID ∀𝑥 ∈ 𝑦 (𝑓‘𝑥) = 1o)} | ||
29-May-2024 | pw1nct 13617 | A condition which ensures that the powerset of a singleton is not countable. The antecedent here can be referred to as the uniformity principle. Based on Mastodon posts by Andrej Bauer and Rahul Chhabra. (Contributed by Jim Kingdon, 29-May-2024.) |
⊢ (∀𝑟(𝑟 ⊆ (𝒫 1o × ω) → (∀𝑝 ∈ 𝒫 1o∃𝑛 ∈ ω 𝑝𝑟𝑛 → ∃𝑚 ∈ ω ∀𝑞 ∈ 𝒫 1o𝑞𝑟𝑚)) → ¬ ∃𝑓 𝑓:ω–onto→(𝒫 1o ⊔ 1o)) | ||
28-May-2024 | sssneq 13616 | Any two elements of a subset of a singleton are equal. (Contributed by Jim Kingdon, 28-May-2024.) |
⊢ (𝐴 ⊆ {𝐵} → ∀𝑦 ∈ 𝐴 ∀𝑧 ∈ 𝐴 𝑦 = 𝑧) | ||
26-May-2024 | elpwi2 4120 | Membership in a power class. (Contributed by Glauco Siliprandi, 3-Mar-2021.) (Proof shortened by Wolf Lammen, 26-May-2024.) |
⊢ 𝐵 ∈ 𝑉 & ⊢ 𝐴 ⊆ 𝐵 ⇒ ⊢ 𝐴 ∈ 𝒫 𝐵 | ||
24-May-2024 | dvmptcjx 13128 | Function-builder for derivative, conjugate rule. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 24-May-2024.) |
⊢ ((𝜑 ∧ 𝑥 ∈ 𝑋) → 𝐴 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝑋) → 𝐵 ∈ 𝑉) & ⊢ (𝜑 → (ℝ D (𝑥 ∈ 𝑋 ↦ 𝐴)) = (𝑥 ∈ 𝑋 ↦ 𝐵)) & ⊢ (𝜑 → 𝑋 ⊆ ℝ) ⇒ ⊢ (𝜑 → (ℝ D (𝑥 ∈ 𝑋 ↦ (∗‘𝐴))) = (𝑥 ∈ 𝑋 ↦ (∗‘𝐵))) | ||
22-May-2024 | efltlemlt 13137 | Lemma for eflt 13138. The converse of efltim 11599 plus the epsilon-delta setup. (Contributed by Jim Kingdon, 22-May-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → (exp‘𝐴) < (exp‘𝐵)) & ⊢ (𝜑 → 𝐷 ∈ ℝ+) & ⊢ (𝜑 → ((abs‘(𝐴 − 𝐵)) < 𝐷 → (abs‘((exp‘𝐴) − (exp‘𝐵))) < ((exp‘𝐵) − (exp‘𝐴)))) ⇒ ⊢ (𝜑 → 𝐴 < 𝐵) | ||
21-May-2024 | eflt 13138 | The exponential function on the reals is strictly increasing. (Contributed by Paul Chapman, 21-Aug-2007.) (Revised by Jim Kingdon, 21-May-2024.) |
⊢ ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → (𝐴 < 𝐵 ↔ (exp‘𝐴) < (exp‘𝐵))) | ||
19-May-2024 | apdifflemr 13660 | Lemma for apdiff 13661. (Contributed by Jim Kingdon, 19-May-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝑆 ∈ ℚ) & ⊢ (𝜑 → (abs‘(𝐴 − -1)) # (abs‘(𝐴 − 1))) & ⊢ ((𝜑 ∧ 𝑆 ≠ 0) → (abs‘(𝐴 − 0)) # (abs‘(𝐴 − (2 · 𝑆)))) ⇒ ⊢ (𝜑 → 𝐴 # 𝑆) | ||
18-May-2024 | apdifflemf 13659 | Lemma for apdiff 13661. Being apart from the point halfway between 𝑄 and 𝑅 suffices for 𝐴 to be a different distance from 𝑄 and from 𝑅. (Contributed by Jim Kingdon, 18-May-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝑄 ∈ ℚ) & ⊢ (𝜑 → 𝑅 ∈ ℚ) & ⊢ (𝜑 → 𝑄 < 𝑅) & ⊢ (𝜑 → ((𝑄 + 𝑅) / 2) # 𝐴) ⇒ ⊢ (𝜑 → (abs‘(𝐴 − 𝑄)) # (abs‘(𝐴 − 𝑅))) | ||
17-May-2024 | apdiff 13661 | The irrationals (reals apart from any rational) are exactly those reals that are a different distance from every rational. (Contributed by Jim Kingdon, 17-May-2024.) |
⊢ (𝐴 ∈ ℝ → (∀𝑞 ∈ ℚ 𝐴 # 𝑞 ↔ ∀𝑞 ∈ ℚ ∀𝑟 ∈ ℚ (𝑞 ≠ 𝑟 → (abs‘(𝐴 − 𝑞)) # (abs‘(𝐴 − 𝑟))))) | ||
15-May-2024 | reeff1oleme 13135 | Lemma for reeff1o 13136. (Contributed by Jim Kingdon, 15-May-2024.) |
⊢ (𝑈 ∈ (0(,)e) → ∃𝑥 ∈ ℝ (exp‘𝑥) = 𝑈) | ||
14-May-2024 | df-relog 13221 | Define the natural logarithm function. Defining the logarithm on complex numbers is similar to square root - there are ways to define it but they tend to make use of excluded middle. Therefore, we merely define logarithms on positive reals. See http://en.wikipedia.org/wiki/Natural_logarithm and https://en.wikipedia.org/wiki/Complex_logarithm. (Contributed by Jim Kingdon, 14-May-2024.) |
⊢ log = ◡(exp ↾ ℝ) | ||
7-May-2024 | ioocosf1o 13217 | The cosine function is a bijection when restricted to its principal domain. (Contributed by Mario Carneiro, 12-May-2014.) (Revised by Jim Kingdon, 7-May-2024.) |
⊢ (cos ↾ (0(,)π)):(0(,)π)–1-1-onto→(-1(,)1) | ||
7-May-2024 | cos0pilt1 13215 | Cosine is between minus one and one on the open interval between zero and π. (Contributed by Jim Kingdon, 7-May-2024.) |
⊢ (𝐴 ∈ (0(,)π) → (cos‘𝐴) ∈ (-1(,)1)) | ||
6-May-2024 | cos11 13216 | Cosine is one-to-one over the closed interval from 0 to π. (Contributed by Paul Chapman, 16-Mar-2008.) (Revised by Jim Kingdon, 6-May-2024.) |
⊢ ((𝐴 ∈ (0[,]π) ∧ 𝐵 ∈ (0[,]π)) → (𝐴 = 𝐵 ↔ (cos‘𝐴) = (cos‘𝐵))) | ||
5-May-2024 | omiunct 12215 | The union of a countably infinite collection of countable sets is countable. Theorem 8.1.28 of [AczelRathjen], p. 78. Compare with ctiunct 12211 which has a stronger hypothesis but does not require countable choice. (Contributed by Jim Kingdon, 5-May-2024.) |
⊢ (𝜑 → CCHOICE) & ⊢ ((𝜑 ∧ 𝑥 ∈ ω) → ∃𝑔 𝑔:ω–onto→(𝐵 ⊔ 1o)) ⇒ ⊢ (𝜑 → ∃ℎ ℎ:ω–onto→(∪ 𝑥 ∈ ω 𝐵 ⊔ 1o)) | ||
5-May-2024 | ctiunctal 12212 | Variation of ctiunct 12211 which allows 𝑥 to be present in 𝜑. (Contributed by Jim Kingdon, 5-May-2024.) |
⊢ (𝜑 → 𝐹:ω–onto→(𝐴 ⊔ 1o)) & ⊢ (𝜑 → ∀𝑥 ∈ 𝐴 𝐺:ω–onto→(𝐵 ⊔ 1o)) ⇒ ⊢ (𝜑 → ∃ℎ ℎ:ω–onto→(∪ 𝑥 ∈ 𝐴 𝐵 ⊔ 1o)) | ||
3-May-2024 | cc4n 7192 | Countable choice with a simpler restriction on how every set in the countable collection needs to be inhabited. That is, compared with cc4 7191, the hypotheses only require an A(n) for each value of 𝑛, not a single set 𝐴 which suffices for every 𝑛 ∈ ω. (Contributed by Mario Carneiro, 7-Apr-2013.) (Revised by Jim Kingdon, 3-May-2024.) |
⊢ (𝜑 → CCHOICE) & ⊢ (𝜑 → ∀𝑛 ∈ 𝑁 {𝑥 ∈ 𝐴 ∣ 𝜓} ∈ 𝑉) & ⊢ (𝜑 → 𝑁 ≈ ω) & ⊢ (𝑥 = (𝑓‘𝑛) → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → ∀𝑛 ∈ 𝑁 ∃𝑥 ∈ 𝐴 𝜓) ⇒ ⊢ (𝜑 → ∃𝑓(𝑓 Fn 𝑁 ∧ ∀𝑛 ∈ 𝑁 𝜒)) | ||
3-May-2024 | cc4f 7190 | Countable choice by showing the existence of a function 𝑓 which can choose a value at each index 𝑛 such that 𝜒 holds. (Contributed by Mario Carneiro, 7-Apr-2013.) (Revised by Jim Kingdon, 3-May-2024.) |
⊢ (𝜑 → CCHOICE) & ⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ Ⅎ𝑛𝐴 & ⊢ (𝜑 → 𝑁 ≈ ω) & ⊢ (𝑥 = (𝑓‘𝑛) → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → ∀𝑛 ∈ 𝑁 ∃𝑥 ∈ 𝐴 𝜓) ⇒ ⊢ (𝜑 → ∃𝑓(𝑓:𝑁⟶𝐴 ∧ ∀𝑛 ∈ 𝑁 𝜒)) | ||
1-May-2024 | cc4 7191 | Countable choice by showing the existence of a function 𝑓 which can choose a value at each index 𝑛 such that 𝜒 holds. (Contributed by Mario Carneiro, 7-Apr-2013.) (Revised by Jim Kingdon, 1-May-2024.) |
⊢ (𝜑 → CCHOICE) & ⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → 𝑁 ≈ ω) & ⊢ (𝑥 = (𝑓‘𝑛) → (𝜓 ↔ 𝜒)) & ⊢ (𝜑 → ∀𝑛 ∈ 𝑁 ∃𝑥 ∈ 𝐴 𝜓) ⇒ ⊢ (𝜑 → ∃𝑓(𝑓:𝑁⟶𝐴 ∧ ∀𝑛 ∈ 𝑁 𝜒)) | ||
29-Apr-2024 | cc3 7189 | Countable choice using a sequence F(n) . (Contributed by Mario Carneiro, 8-Feb-2013.) (Revised by Jim Kingdon, 29-Apr-2024.) |
⊢ (𝜑 → CCHOICE) & ⊢ (𝜑 → ∀𝑛 ∈ 𝑁 𝐹 ∈ V) & ⊢ (𝜑 → ∀𝑛 ∈ 𝑁 ∃𝑤 𝑤 ∈ 𝐹) & ⊢ (𝜑 → 𝑁 ≈ ω) ⇒ ⊢ (𝜑 → ∃𝑓(𝑓 Fn 𝑁 ∧ ∀𝑛 ∈ 𝑁 (𝑓‘𝑛) ∈ 𝐹)) | ||
27-Apr-2024 | cc2 7188 | Countable choice using sequences instead of countable sets. (Contributed by Jim Kingdon, 27-Apr-2024.) |
⊢ (𝜑 → CCHOICE) & ⊢ (𝜑 → 𝐹 Fn ω) & ⊢ (𝜑 → ∀𝑥 ∈ ω ∃𝑤 𝑤 ∈ (𝐹‘𝑥)) ⇒ ⊢ (𝜑 → ∃𝑔(𝑔 Fn ω ∧ ∀𝑛 ∈ ω (𝑔‘𝑛) ∈ (𝐹‘𝑛))) | ||
27-Apr-2024 | cc2lem 7187 | Lemma for cc2 7188. (Contributed by Jim Kingdon, 27-Apr-2024.) |
⊢ (𝜑 → CCHOICE) & ⊢ (𝜑 → 𝐹 Fn ω) & ⊢ (𝜑 → ∀𝑥 ∈ ω ∃𝑤 𝑤 ∈ (𝐹‘𝑥)) & ⊢ 𝐴 = (𝑛 ∈ ω ↦ ({𝑛} × (𝐹‘𝑛))) & ⊢ 𝐺 = (𝑛 ∈ ω ↦ (2nd ‘(𝑓‘(𝐴‘𝑛)))) ⇒ ⊢ (𝜑 → ∃𝑔(𝑔 Fn ω ∧ ∀𝑛 ∈ ω (𝑔‘𝑛) ∈ (𝐹‘𝑛))) | ||
27-Apr-2024 | cc1 7186 | Countable choice in terms of a choice function on a countably infinite set of inhabited sets. (Contributed by Jim Kingdon, 27-Apr-2024.) |
⊢ (CCHOICE → ∀𝑥((𝑥 ≈ ω ∧ ∀𝑧 ∈ 𝑥 ∃𝑤 𝑤 ∈ 𝑧) → ∃𝑓∀𝑧 ∈ 𝑥 (𝑓‘𝑧) ∈ 𝑧)) | ||
19-Apr-2024 | omctfn 12214 | Using countable choice to find a sequence of enumerations for a collection of countable sets. Lemma 8.1.27 of [AczelRathjen], p. 77. (Contributed by Jim Kingdon, 19-Apr-2024.) |
⊢ (𝜑 → CCHOICE) & ⊢ ((𝜑 ∧ 𝑥 ∈ ω) → ∃𝑔 𝑔:ω–onto→(𝐵 ⊔ 1o)) ⇒ ⊢ (𝜑 → ∃𝑓(𝑓 Fn ω ∧ ∀𝑥 ∈ ω (𝑓‘𝑥):ω–onto→(𝐵 ⊔ 1o))) | ||
13-Apr-2024 | prodmodclem2 11478 | Lemma for prodmodc 11479. (Contributed by Scott Fenton, 4-Dec-2017.) (Revised by Jim Kingdon, 13-Apr-2024.) |
⊢ 𝐹 = (𝑘 ∈ ℤ ↦ if(𝑘 ∈ 𝐴, 𝐵, 1)) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐵 ∈ ℂ) & ⊢ 𝐺 = (𝑗 ∈ ℕ ↦ if(𝑗 ≤ (♯‘𝐴), ⦋(𝑓‘𝑗) / 𝑘⦌𝐵, 1)) ⇒ ⊢ ((𝜑 ∧ ∃𝑚 ∈ ℤ ((𝐴 ⊆ (ℤ≥‘𝑚) ∧ ∀𝑗 ∈ (ℤ≥‘𝑚)DECID 𝑗 ∈ 𝐴) ∧ (∃𝑛 ∈ (ℤ≥‘𝑚)∃𝑦(𝑦 # 0 ∧ seq𝑛( · , 𝐹) ⇝ 𝑦) ∧ seq𝑚( · , 𝐹) ⇝ 𝑥))) → (∃𝑚 ∈ ℕ ∃𝑓(𝑓:(1...𝑚)–1-1-onto→𝐴 ∧ 𝑧 = (seq1( · , 𝐺)‘𝑚)) → 𝑥 = 𝑧)) | ||
11-Apr-2024 | prodmodclem2a 11477 | Lemma for prodmodc 11479. (Contributed by Scott Fenton, 4-Dec-2017.) (Revised by Jim Kingdon, 11-Apr-2024.) |
⊢ 𝐹 = (𝑘 ∈ ℤ ↦ if(𝑘 ∈ 𝐴, 𝐵, 1)) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐵 ∈ ℂ) & ⊢ 𝐺 = (𝑗 ∈ ℕ ↦ if(𝑗 ≤ (♯‘𝐴), ⦋(𝑓‘𝑗) / 𝑘⦌𝐵, 1)) & ⊢ 𝐻 = (𝑗 ∈ ℕ ↦ if(𝑗 ≤ (♯‘𝐴), ⦋(𝐾‘𝑗) / 𝑘⦌𝐵, 1)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → DECID 𝑘 ∈ 𝐴) & ⊢ (𝜑 → 𝑁 ∈ ℕ) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝐴 ⊆ (ℤ≥‘𝑀)) & ⊢ (𝜑 → 𝑓:(1...𝑁)–1-1-onto→𝐴) & ⊢ (𝜑 → 𝐾 Isom < , < ((1...(♯‘𝐴)), 𝐴)) ⇒ ⊢ (𝜑 → seq𝑀( · , 𝐹) ⇝ (seq1( · , 𝐺)‘𝑁)) | ||
11-Apr-2024 | prodmodclem3 11476 | Lemma for prodmodc 11479. (Contributed by Scott Fenton, 4-Dec-2017.) (Revised by Jim Kingdon, 11-Apr-2024.) |
⊢ 𝐹 = (𝑘 ∈ ℤ ↦ if(𝑘 ∈ 𝐴, 𝐵, 1)) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐵 ∈ ℂ) & ⊢ 𝐺 = (𝑗 ∈ ℕ ↦ if(𝑗 ≤ (♯‘𝐴), ⦋(𝑓‘𝑗) / 𝑘⦌𝐵, 1)) & ⊢ 𝐻 = (𝑗 ∈ ℕ ↦ if(𝑗 ≤ (♯‘𝐴), ⦋(𝐾‘𝑗) / 𝑘⦌𝐵, 1)) & ⊢ (𝜑 → (𝑀 ∈ ℕ ∧ 𝑁 ∈ ℕ)) & ⊢ (𝜑 → 𝑓:(1...𝑀)–1-1-onto→𝐴) & ⊢ (𝜑 → 𝐾:(1...𝑁)–1-1-onto→𝐴) ⇒ ⊢ (𝜑 → (seq1( · , 𝐺)‘𝑀) = (seq1( · , 𝐻)‘𝑁)) | ||
10-Apr-2024 | jcnd 642 | Deduction joining the consequents of two premises. (Contributed by Glauco Siliprandi, 11-Dec-2019.) (Proof shortened by Wolf Lammen, 10-Apr-2024.) |
⊢ (𝜑 → 𝜓) & ⊢ (𝜑 → ¬ 𝜒) ⇒ ⊢ (𝜑 → ¬ (𝜓 → 𝜒)) | ||
4-Apr-2024 | prodrbdclem 11472 | Lemma for prodrbdc 11475. (Contributed by Scott Fenton, 4-Dec-2017.) (Revised by Jim Kingdon, 4-Apr-2024.) |
⊢ 𝐹 = (𝑘 ∈ ℤ ↦ if(𝑘 ∈ 𝐴, 𝐵, 1)) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → 𝐵 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → DECID 𝑘 ∈ 𝐴) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) ⇒ ⊢ ((𝜑 ∧ 𝐴 ⊆ (ℤ≥‘𝑁)) → (seq𝑀( · , 𝐹) ↾ (ℤ≥‘𝑁)) = seq𝑁( · , 𝐹)) | ||
24-Mar-2024 | prodfdivap 11448 | The quotient of two products. (Contributed by Scott Fenton, 15-Jan-2018.) (Revised by Jim Kingdon, 24-Mar-2024.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑘) ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑘) # 0) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐻‘𝑘) = ((𝐹‘𝑘) / (𝐺‘𝑘))) ⇒ ⊢ (𝜑 → (seq𝑀( · , 𝐻)‘𝑁) = ((seq𝑀( · , 𝐹)‘𝑁) / (seq𝑀( · , 𝐺)‘𝑁))) | ||
24-Mar-2024 | prodfrecap 11447 | The reciprocal of a finite product. (Contributed by Scott Fenton, 15-Jan-2018.) (Revised by Jim Kingdon, 24-Mar-2024.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → (𝐹‘𝑘) # 0) & ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → (𝐺‘𝑘) = (1 / (𝐹‘𝑘))) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑘) ∈ ℂ) ⇒ ⊢ (𝜑 → (seq𝑀( · , 𝐺)‘𝑁) = (1 / (seq𝑀( · , 𝐹)‘𝑁))) | ||
23-Mar-2024 | prodfap0 11446 | The product of finitely many terms apart from zero is apart from zero. (Contributed by Scott Fenton, 14-Jan-2018.) (Revised by Jim Kingdon, 23-Mar-2024.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (𝑀...𝑁)) → (𝐹‘𝑘) # 0) ⇒ ⊢ (𝜑 → (seq𝑀( · , 𝐹)‘𝑁) # 0) | ||
22-Mar-2024 | prod3fmul 11442 | The product of two infinite products. (Contributed by Scott Fenton, 18-Dec-2017.) (Revised by Jim Kingdon, 22-Mar-2024.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐹‘𝑘) ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐺‘𝑘) ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑘 ∈ (ℤ≥‘𝑀)) → (𝐻‘𝑘) = ((𝐹‘𝑘) · (𝐺‘𝑘))) ⇒ ⊢ (𝜑 → (seq𝑀( · , 𝐻)‘𝑁) = ((seq𝑀( · , 𝐹)‘𝑁) · (seq𝑀( · , 𝐺)‘𝑁))) | ||
21-Mar-2024 | df-proddc 11452 | Define the product of a series with an index set of integers 𝐴. This definition takes most of the aspects of df-sumdc 11255 and adapts them for multiplication instead of addition. However, we insist that in the infinite case, there is a nonzero tail of the sequence. This ensures that the convergence criteria match those of infinite sums. (Contributed by Scott Fenton, 4-Dec-2017.) (Revised by Jim Kingdon, 21-Mar-2024.) |
⊢ ∏𝑘 ∈ 𝐴 𝐵 = (℩𝑥(∃𝑚 ∈ ℤ ((𝐴 ⊆ (ℤ≥‘𝑚) ∧ ∀𝑗 ∈ (ℤ≥‘𝑚)DECID 𝑗 ∈ 𝐴) ∧ (∃𝑛 ∈ (ℤ≥‘𝑚)∃𝑦(𝑦 # 0 ∧ seq𝑛( · , (𝑘 ∈ ℤ ↦ if(𝑘 ∈ 𝐴, 𝐵, 1))) ⇝ 𝑦) ∧ seq𝑚( · , (𝑘 ∈ ℤ ↦ if(𝑘 ∈ 𝐴, 𝐵, 1))) ⇝ 𝑥)) ∨ ∃𝑚 ∈ ℕ ∃𝑓(𝑓:(1...𝑚)–1-1-onto→𝐴 ∧ 𝑥 = (seq1( · , (𝑛 ∈ ℕ ↦ if(𝑛 ≤ 𝑚, ⦋(𝑓‘𝑛) / 𝑘⦌𝐵, 1)))‘𝑚)))) | ||
19-Mar-2024 | cos02pilt1 13214 | Cosine is less than one between zero and 2 · π. (Contributed by Jim Kingdon, 19-Mar-2024.) |
⊢ (𝐴 ∈ (0(,)(2 · π)) → (cos‘𝐴) < 1) | ||
19-Mar-2024 | cosq34lt1 13213 | Cosine is less than one in the third and fourth quadrants. (Contributed by Jim Kingdon, 19-Mar-2024.) |
⊢ (𝐴 ∈ (π[,)(2 · π)) → (cos‘𝐴) < 1) | ||
14-Mar-2024 | coseq0q4123 13197 | Location of the zeroes of cosine in (-(π / 2)(,)(3 · (π / 2))). (Contributed by Jim Kingdon, 14-Mar-2024.) |
⊢ (𝐴 ∈ (-(π / 2)(,)(3 · (π / 2))) → ((cos‘𝐴) = 0 ↔ 𝐴 = (π / 2))) | ||
14-Mar-2024 | cosq23lt0 13196 | The cosine of a number in the second and third quadrants is negative. (Contributed by Jim Kingdon, 14-Mar-2024.) |
⊢ (𝐴 ∈ ((π / 2)(,)(3 · (π / 2))) → (cos‘𝐴) < 0) | ||
9-Mar-2024 | pilem3 13146 | Lemma for pi related theorems. (Contributed by Jim Kingdon, 9-Mar-2024.) |
⊢ (π ∈ (2(,)4) ∧ (sin‘π) = 0) | ||
9-Mar-2024 | exmidonfin 7130 | If a finite ordinal is a natural number, excluded middle follows. That excluded middle implies that a finite ordinal is a natural number is proved in the Metamath Proof Explorer. That a natural number is a finite ordinal is shown at nnfi 6818 and nnon 4570. (Contributed by Andrew W Swan and Jim Kingdon, 9-Mar-2024.) |
⊢ (ω = (On ∩ Fin) → EXMID) | ||
9-Mar-2024 | exmidonfinlem 7129 | Lemma for exmidonfin 7130. (Contributed by Andrew W Swan and Jim Kingdon, 9-Mar-2024.) |
⊢ 𝐴 = {{𝑥 ∈ {∅} ∣ 𝜑}, {𝑥 ∈ {∅} ∣ ¬ 𝜑}} ⇒ ⊢ (ω = (On ∩ Fin) → DECID 𝜑) | ||
8-Mar-2024 | sin0pilem2 13145 | Lemma for pi related theorems. (Contributed by Mario Carneiro and Jim Kingdon, 8-Mar-2024.) |
⊢ ∃𝑞 ∈ (2(,)4)((sin‘𝑞) = 0 ∧ ∀𝑥 ∈ (0(,)𝑞)0 < (sin‘𝑥)) | ||
8-Mar-2024 | sin0pilem1 13144 | Lemma for pi related theorems. (Contributed by Mario Carneiro and Jim Kingdon, 8-Mar-2024.) |
⊢ ∃𝑝 ∈ (1(,)2)((cos‘𝑝) = 0 ∧ ∀𝑥 ∈ (𝑝(,)(2 · 𝑝))0 < (sin‘𝑥)) | ||
7-Mar-2024 | cosz12 13143 | Cosine has a zero between 1 and 2. (Contributed by Mario Carneiro and Jim Kingdon, 7-Mar-2024.) |
⊢ ∃𝑝 ∈ (1(,)2)(cos‘𝑝) = 0 | ||
6-Mar-2024 | cos12dec 11668 | Cosine is decreasing from one to two. (Contributed by Mario Carneiro and Jim Kingdon, 6-Mar-2024.) |
⊢ ((𝐴 ∈ (1[,]2) ∧ 𝐵 ∈ (1[,]2) ∧ 𝐴 < 𝐵) → (cos‘𝐵) < (cos‘𝐴)) | ||
25-Feb-2024 | mul2lt0pn 9672 | The product of multiplicands of different signs is negative. (Contributed by Jim Kingdon, 25-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 0) & ⊢ (𝜑 → 0 < 𝐵) ⇒ ⊢ (𝜑 → (𝐵 · 𝐴) < 0) | ||
25-Feb-2024 | mul2lt0np 9671 | The product of multiplicands of different signs is negative. (Contributed by Jim Kingdon, 25-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 0) & ⊢ (𝜑 → 0 < 𝐵) ⇒ ⊢ (𝜑 → (𝐴 · 𝐵) < 0) | ||
25-Feb-2024 | lt0ap0 8524 | A number which is less than zero is apart from zero. (Contributed by Jim Kingdon, 25-Feb-2024.) |
⊢ ((𝐴 ∈ ℝ ∧ 𝐴 < 0) → 𝐴 # 0) | ||
25-Feb-2024 | negap0d 8507 | The negative of a number apart from zero is apart from zero. (Contributed by Jim Kingdon, 25-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℂ) & ⊢ (𝜑 → 𝐴 # 0) ⇒ ⊢ (𝜑 → -𝐴 # 0) | ||
24-Feb-2024 | lt0ap0d 8525 | A real number less than zero is apart from zero. Deduction form. (Contributed by Jim Kingdon, 24-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 0) ⇒ ⊢ (𝜑 → 𝐴 # 0) | ||
20-Feb-2024 | ivthdec 13064 | The intermediate value theorem, decreasing case, for a strictly monotonic function. (Contributed by Jim Kingdon, 20-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐵) < 𝑈 ∧ 𝑈 < (𝐹‘𝐴))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑦) < (𝐹‘𝑥)) ⇒ ⊢ (𝜑 → ∃𝑐 ∈ (𝐴(,)𝐵)(𝐹‘𝑐) = 𝑈) | ||
20-Feb-2024 | ivthinclemex 13062 | Lemma for ivthinc 13063. Existence of a number between the lower cut and the upper cut. (Contributed by Jim Kingdon, 20-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∃!𝑧 ∈ (𝐴(,)𝐵)(∀𝑞 ∈ 𝐿 𝑞 < 𝑧 ∧ ∀𝑟 ∈ 𝑅 𝑧 < 𝑟)) | ||
19-Feb-2024 | ivthinclemuopn 13058 | Lemma for ivthinc 13063. The upper cut is open. (Contributed by Jim Kingdon, 19-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} & ⊢ (𝜑 → 𝑆 ∈ 𝑅) ⇒ ⊢ (𝜑 → ∃𝑞 ∈ 𝑅 𝑞 < 𝑆) | ||
19-Feb-2024 | dedekindicc 13053 | A Dedekind cut identifies a unique real number. Similar to df-inp 7387 except that the Dedekind cut is formed by sets of reals (rather than positive rationals). But in both cases the defining property of a Dedekind cut is that it is inhabited (bounded), rounded, disjoint, and located. (Contributed by Jim Kingdon, 19-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 < 𝐵) ⇒ ⊢ (𝜑 → ∃!𝑥 ∈ (𝐴(,)𝐵)(∀𝑞 ∈ 𝐿 𝑞 < 𝑥 ∧ ∀𝑟 ∈ 𝑈 𝑥 < 𝑟)) | ||
18-Feb-2024 | ivthinclemloc 13061 | Lemma for ivthinc 13063. Locatedness. (Contributed by Jim Kingdon, 18-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑅))) | ||
18-Feb-2024 | ivthinclemdisj 13060 | Lemma for ivthinc 13063. The lower and upper cuts are disjoint. (Contributed by Jim Kingdon, 18-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → (𝐿 ∩ 𝑅) = ∅) | ||
18-Feb-2024 | ivthinclemur 13059 | Lemma for ivthinc 13063. The upper cut is rounded. (Contributed by Jim Kingdon, 18-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑅 ↔ ∃𝑞 ∈ 𝑅 𝑞 < 𝑟)) | ||
18-Feb-2024 | ivthinclemlr 13057 | Lemma for ivthinc 13063. The lower cut is rounded. (Contributed by Jim Kingdon, 18-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) | ||
18-Feb-2024 | ivthinclemum 13055 | Lemma for ivthinc 13063. The upper cut is bounded. (Contributed by Jim Kingdon, 18-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑅) | ||
18-Feb-2024 | ivthinclemlm 13054 | Lemma for ivthinc 13063. The lower cut is bounded. (Contributed by Jim Kingdon, 18-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) | ||
15-Feb-2024 | dedekindicclemeu 13051 | Lemma for dedekindicc 13053. Part of proving uniqueness. (Contributed by Jim Kingdon, 15-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → 𝐶 ∈ (𝐴[,]𝐵)) & ⊢ (𝜑 → (∀𝑞 ∈ 𝐿 𝑞 < 𝐶 ∧ ∀𝑟 ∈ 𝑈 𝐶 < 𝑟)) & ⊢ (𝜑 → 𝐷 ∈ (𝐴[,]𝐵)) & ⊢ (𝜑 → (∀𝑞 ∈ 𝐿 𝑞 < 𝐷 ∧ ∀𝑟 ∈ 𝑈 𝐷 < 𝑟)) & ⊢ (𝜑 → 𝐶 < 𝐷) ⇒ ⊢ (𝜑 → ⊥) | ||
15-Feb-2024 | dedekindicclemlu 13050 | Lemma for dedekindicc 13053. There is a number which separates the lower and upper cuts. (Contributed by Jim Kingdon, 15-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 < 𝐵) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ (𝐴[,]𝐵)(∀𝑞 ∈ 𝐿 𝑞 < 𝑥 ∧ ∀𝑟 ∈ 𝑈 𝑥 < 𝑟)) | ||
15-Feb-2024 | dedekindicclemlub 13049 | Lemma for dedekindicc 13053. The set L has a least upper bound. (Contributed by Jim Kingdon, 15-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 < 𝐵) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ (𝐴[,]𝐵)(∀𝑦 ∈ 𝐿 ¬ 𝑥 < 𝑦 ∧ ∀𝑦 ∈ (𝐴[,]𝐵)(𝑦 < 𝑥 → ∃𝑧 ∈ 𝐿 𝑦 < 𝑧))) | ||
15-Feb-2024 | dedekindicclemloc 13048 | Lemma for dedekindicc 13053. The set L is located. (Contributed by Jim Kingdon, 15-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∀𝑥 ∈ (𝐴[,]𝐵)∀𝑦 ∈ (𝐴[,]𝐵)(𝑥 < 𝑦 → (∃𝑧 ∈ 𝐿 𝑥 < 𝑧 ∨ ∀𝑧 ∈ 𝐿 𝑧 < 𝑦))) | ||
15-Feb-2024 | dedekindicclemub 13047 | Lemma for dedekindicc 13053. The lower cut has an upper bound. (Contributed by Jim Kingdon, 15-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ (𝐴[,]𝐵)∀𝑦 ∈ 𝐿 𝑦 < 𝑥) | ||
15-Feb-2024 | dedekindicclemuub 13046 | Lemma for dedekindicc 13053. Any element of the upper cut is an upper bound for the lower cut. (Contributed by Jim Kingdon, 15-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐶 ∈ 𝑈) ⇒ ⊢ (𝜑 → ∀𝑧 ∈ 𝐿 𝑧 < 𝐶) | ||
14-Feb-2024 | suplociccex 13045 | An inhabited, bounded-above, located set of reals in a closed interval has a supremum. A similar theorem is axsuploc 7951 but that one is for the entire real line rather than a closed interval. (Contributed by Jim Kingdon, 14-Feb-2024.) |
⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐶 ∈ ℝ) & ⊢ (𝜑 → 𝐵 < 𝐶) & ⊢ (𝜑 → 𝐴 ⊆ (𝐵[,]𝐶)) & ⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥 ∈ (𝐵[,]𝐶)∀𝑦 ∈ (𝐵[,]𝐶)(𝑥 < 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 < 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 < 𝑦))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ (𝐵[,]𝐶)(∀𝑦 ∈ 𝐴 ¬ 𝑥 < 𝑦 ∧ ∀𝑦 ∈ (𝐵[,]𝐶)(𝑦 < 𝑥 → ∃𝑧 ∈ 𝐴 𝑦 < 𝑧))) | ||
14-Feb-2024 | suplociccreex 13044 | An inhabited, bounded-above, located set of reals in a closed interval has a supremum. A similar theorem is axsuploc 7951 but that one is for the entire real line rather than a closed interval. (Contributed by Jim Kingdon, 14-Feb-2024.) |
⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐶 ∈ ℝ) & ⊢ (𝜑 → 𝐵 < 𝐶) & ⊢ (𝜑 → 𝐴 ⊆ (𝐵[,]𝐶)) & ⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥 ∈ (𝐵[,]𝐶)∀𝑦 ∈ (𝐵[,]𝐶)(𝑥 < 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 < 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 < 𝑦))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℝ (∀𝑦 ∈ 𝐴 ¬ 𝑥 < 𝑦 ∧ ∀𝑦 ∈ ℝ (𝑦 < 𝑥 → ∃𝑧 ∈ 𝐴 𝑦 < 𝑧))) | ||
6-Feb-2024 | ivthinclemlopn 13056 | Lemma for ivthinc 13063. The lower cut is open. (Contributed by Jim Kingdon, 6-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} & ⊢ (𝜑 → 𝑄 ∈ 𝐿) ⇒ ⊢ (𝜑 → ∃𝑟 ∈ 𝐿 𝑄 < 𝑟) | ||
5-Feb-2024 | ivthinc 13063 | The intermediate value theorem, increasing case, for a strictly monotonic function. Theorem 5.5 of [Bauer], p. 494. This is Metamath 100 proof #79. (Contributed by Jim Kingdon, 5-Feb-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) ⇒ ⊢ (𝜑 → ∃𝑐 ∈ (𝐴(,)𝐵)(𝐹‘𝑐) = 𝑈) | ||
2-Feb-2024 | dedekindeulemuub 13037 | Lemma for dedekindeu 13043. Any element of the upper cut is an upper bound for the lower cut. (Contributed by Jim Kingdon, 2-Feb-2024.) |
⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 ∈ 𝑈) ⇒ ⊢ (𝜑 → ∀𝑧 ∈ 𝐿 𝑧 < 𝐴) | ||
31-Jan-2024 | dedekindeulemeu 13042 | Lemma for dedekindeu 13043. Part of proving uniqueness. (Contributed by Jim Kingdon, 31-Jan-2024.) |
⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → (∀𝑞 ∈ 𝐿 𝑞 < 𝐴 ∧ ∀𝑟 ∈ 𝑈 𝐴 < 𝑟)) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → (∀𝑞 ∈ 𝐿 𝑞 < 𝐵 ∧ ∀𝑟 ∈ 𝑈 𝐵 < 𝑟)) & ⊢ (𝜑 → 𝐴 < 𝐵) ⇒ ⊢ (𝜑 → ⊥) | ||
31-Jan-2024 | dedekindeulemlu 13041 | Lemma for dedekindeu 13043. There is a number which separates the lower and upper cuts. (Contributed by Jim Kingdon, 31-Jan-2024.) |
⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℝ (∀𝑞 ∈ 𝐿 𝑞 < 𝑥 ∧ ∀𝑟 ∈ 𝑈 𝑥 < 𝑟)) | ||
31-Jan-2024 | dedekindeulemlub 13040 | Lemma for dedekindeu 13043. The set L has a least upper bound. (Contributed by Jim Kingdon, 31-Jan-2024.) |
⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℝ (∀𝑦 ∈ 𝐿 ¬ 𝑥 < 𝑦 ∧ ∀𝑦 ∈ ℝ (𝑦 < 𝑥 → ∃𝑧 ∈ 𝐿 𝑦 < 𝑧))) | ||
31-Jan-2024 | dedekindeulemloc 13039 | Lemma for dedekindeu 13043. The set L is located. (Contributed by Jim Kingdon, 31-Jan-2024.) |
⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 → (∃𝑧 ∈ 𝐿 𝑥 < 𝑧 ∨ ∀𝑧 ∈ 𝐿 𝑧 < 𝑦))) | ||
31-Jan-2024 | dedekindeulemub 13038 | Lemma for dedekindeu 13043. The lower cut has an upper bound. (Contributed by Jim Kingdon, 31-Jan-2024.) |
⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℝ ∀𝑦 ∈ 𝐿 𝑦 < 𝑥) | ||
30-Jan-2024 | axsuploc 7951 | An inhabited, bounded-above, located set of reals has a supremum. Axiom for real and complex numbers, derived from ZF set theory. (This restates ax-pre-suploc 7854 with ordering on the extended reals.) (Contributed by Jim Kingdon, 30-Jan-2024.) |
⊢ (((𝐴 ⊆ ℝ ∧ ∃𝑥 𝑥 ∈ 𝐴) ∧ (∃𝑥 ∈ ℝ ∀𝑦 ∈ 𝐴 𝑦 < 𝑥 ∧ ∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 < 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 < 𝑦)))) → ∃𝑥 ∈ ℝ (∀𝑦 ∈ 𝐴 ¬ 𝑥 < 𝑦 ∧ ∀𝑦 ∈ ℝ (𝑦 < 𝑥 → ∃𝑧 ∈ 𝐴 𝑦 < 𝑧))) | ||
24-Jan-2024 | axpre-suploclemres 7822 | Lemma for axpre-suploc 7823. The result. The proof just needs to define 𝐵 as basically the same set as 𝐴 (but expressed as a subset of R rather than a subset of ℝ), and apply suplocsr 7730. (Contributed by Jim Kingdon, 24-Jan-2024.) |
⊢ (𝜑 → 𝐴 ⊆ ℝ) & ⊢ (𝜑 → 𝐶 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ ℝ ∀𝑦 ∈ 𝐴 𝑦 <ℝ 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 <ℝ 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 <ℝ 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 <ℝ 𝑦))) & ⊢ 𝐵 = {𝑤 ∈ R ∣ 〈𝑤, 0R〉 ∈ 𝐴} ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℝ (∀𝑦 ∈ 𝐴 ¬ 𝑥 <ℝ 𝑦 ∧ ∀𝑦 ∈ ℝ (𝑦 <ℝ 𝑥 → ∃𝑧 ∈ 𝐴 𝑦 <ℝ 𝑧))) | ||
23-Jan-2024 | ax-pre-suploc 7854 |
An inhabited, bounded-above, located set of reals has a supremum.
Locatedness here means that given 𝑥 < 𝑦, either there is an element of the set greater than 𝑥, or 𝑦 is an upper bound. Although this and ax-caucvg 7853 are both completeness properties, countable choice would probably be needed to derive this from ax-caucvg 7853. (Contributed by Jim Kingdon, 23-Jan-2024.) |
⊢ (((𝐴 ⊆ ℝ ∧ ∃𝑥 𝑥 ∈ 𝐴) ∧ (∃𝑥 ∈ ℝ ∀𝑦 ∈ 𝐴 𝑦 <ℝ 𝑥 ∧ ∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 <ℝ 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 <ℝ 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 <ℝ 𝑦)))) → ∃𝑥 ∈ ℝ (∀𝑦 ∈ 𝐴 ¬ 𝑥 <ℝ 𝑦 ∧ ∀𝑦 ∈ ℝ (𝑦 <ℝ 𝑥 → ∃𝑧 ∈ 𝐴 𝑦 <ℝ 𝑧))) | ||
23-Jan-2024 | axpre-suploc 7823 |
An inhabited, bounded-above, located set of reals has a supremum.
Locatedness here means that given 𝑥 < 𝑦, either there is an element of the set greater than 𝑥, or 𝑦 is an upper bound. This construction-dependent theorem should not be referenced directly; instead, use ax-pre-suploc 7854. (Contributed by Jim Kingdon, 23-Jan-2024.) (New usage is discouraged.) |
⊢ (((𝐴 ⊆ ℝ ∧ ∃𝑥 𝑥 ∈ 𝐴) ∧ (∃𝑥 ∈ ℝ ∀𝑦 ∈ 𝐴 𝑦 <ℝ 𝑥 ∧ ∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 <ℝ 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 <ℝ 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 <ℝ 𝑦)))) → ∃𝑥 ∈ ℝ (∀𝑦 ∈ 𝐴 ¬ 𝑥 <ℝ 𝑦 ∧ ∀𝑦 ∈ ℝ (𝑦 <ℝ 𝑥 → ∃𝑧 ∈ 𝐴 𝑦 <ℝ 𝑧))) | ||
22-Jan-2024 | suplocsr 7730 | An inhabited, bounded, located set of signed reals has a supremum. (Contributed by Jim Kingdon, 22-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ R ∀𝑦 ∈ 𝐴 𝑦 <R 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ R ∀𝑦 ∈ R (𝑥 <R 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 <R 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 <R 𝑦))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ R (∀𝑦 ∈ 𝐴 ¬ 𝑥 <R 𝑦 ∧ ∀𝑦 ∈ R (𝑦 <R 𝑥 → ∃𝑧 ∈ 𝐴 𝑦 <R 𝑧))) | ||
21-Jan-2024 | bj-el2oss1o 13390 | Shorter proof of el2oss1o 6391 using more axioms. (Contributed by BJ, 21-Jan-2024.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝐴 ∈ 2o → 𝐴 ⊆ 1o) | ||
21-Jan-2024 | ltm1sr 7698 | Adding minus one to a signed real yields a smaller signed real. (Contributed by Jim Kingdon, 21-Jan-2024.) |
⊢ (𝐴 ∈ R → (𝐴 +R -1R) <R 𝐴) | ||
19-Jan-2024 | suplocsrlempr 7728 | Lemma for suplocsr 7730. The set 𝐵 has a least upper bound. (Contributed by Jim Kingdon, 19-Jan-2024.) |
⊢ 𝐵 = {𝑤 ∈ P ∣ (𝐶 +R [〈𝑤, 1P〉] ~R ) ∈ 𝐴} & ⊢ (𝜑 → 𝐴 ⊆ R) & ⊢ (𝜑 → 𝐶 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ R ∀𝑦 ∈ 𝐴 𝑦 <R 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ R ∀𝑦 ∈ R (𝑥 <R 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 <R 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 <R 𝑦))) ⇒ ⊢ (𝜑 → ∃𝑣 ∈ P (∀𝑤 ∈ 𝐵 ¬ 𝑣<P 𝑤 ∧ ∀𝑤 ∈ P (𝑤<P 𝑣 → ∃𝑢 ∈ 𝐵 𝑤<P 𝑢))) | ||
18-Jan-2024 | suplocsrlemb 7727 | Lemma for suplocsr 7730. The set 𝐵 is located. (Contributed by Jim Kingdon, 18-Jan-2024.) |
⊢ 𝐵 = {𝑤 ∈ P ∣ (𝐶 +R [〈𝑤, 1P〉] ~R ) ∈ 𝐴} & ⊢ (𝜑 → 𝐴 ⊆ R) & ⊢ (𝜑 → 𝐶 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ R ∀𝑦 ∈ 𝐴 𝑦 <R 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ R ∀𝑦 ∈ R (𝑥 <R 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 <R 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 <R 𝑦))) ⇒ ⊢ (𝜑 → ∀𝑢 ∈ P ∀𝑣 ∈ P (𝑢<P 𝑣 → (∃𝑞 ∈ 𝐵 𝑢<P 𝑞 ∨ ∀𝑞 ∈ 𝐵 𝑞<P 𝑣))) | ||
16-Jan-2024 | suplocsrlem 7729 | Lemma for suplocsr 7730. The set 𝐴 has a least upper bound. (Contributed by Jim Kingdon, 16-Jan-2024.) |
⊢ 𝐵 = {𝑤 ∈ P ∣ (𝐶 +R [〈𝑤, 1P〉] ~R ) ∈ 𝐴} & ⊢ (𝜑 → 𝐴 ⊆ R) & ⊢ (𝜑 → 𝐶 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ R ∀𝑦 ∈ 𝐴 𝑦 <R 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ R ∀𝑦 ∈ R (𝑥 <R 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 <R 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 <R 𝑦))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ R (∀𝑦 ∈ 𝐴 ¬ 𝑥 <R 𝑦 ∧ ∀𝑦 ∈ R (𝑦 <R 𝑥 → ∃𝑧 ∈ 𝐴 𝑦 <R 𝑧))) | ||
14-Jan-2024 | suplocexprlemlub 7645 | Lemma for suplocexpr 7646. The putative supremum is a least upper bound. (Contributed by Jim Kingdon, 14-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) & ⊢ 𝐵 = 〈∪ (1st “ 𝐴), {𝑢 ∈ Q ∣ ∃𝑤 ∈ ∩ (2nd “ 𝐴)𝑤 <Q 𝑢}〉 ⇒ ⊢ (𝜑 → (𝑦<P 𝐵 → ∃𝑧 ∈ 𝐴 𝑦<P 𝑧)) | ||
14-Jan-2024 | suplocexprlemub 7644 | Lemma for suplocexpr 7646. The putative supremum is an upper bound. (Contributed by Jim Kingdon, 14-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) & ⊢ 𝐵 = 〈∪ (1st “ 𝐴), {𝑢 ∈ Q ∣ ∃𝑤 ∈ ∩ (2nd “ 𝐴)𝑤 <Q 𝑢}〉 ⇒ ⊢ (𝜑 → ∀𝑦 ∈ 𝐴 ¬ 𝐵<P 𝑦) | ||
10-Jan-2024 | nfcsbw 3067 | Bound-variable hypothesis builder for substitution into a class. Version of nfcsb 3068 with a disjoint variable condition. (Contributed by Mario Carneiro, 12-Oct-2016.) (Revised by Gino Giotto, 10-Jan-2024.) |
⊢ Ⅎ𝑥𝐴 & ⊢ Ⅎ𝑥𝐵 ⇒ ⊢ Ⅎ𝑥⦋𝐴 / 𝑦⦌𝐵 | ||
10-Jan-2024 | nfsbcdw 3065 | Version of nfsbcd 2956 with a disjoint variable condition. (Contributed by NM, 23-Nov-2005.) (Revised by Gino Giotto, 10-Jan-2024.) |
⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝐴) & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥[𝐴 / 𝑦]𝜓) | ||
10-Jan-2024 | cbvcsbw 3035 | Version of cbvcsb 3036 with a disjoint variable condition. (Contributed by Gino Giotto, 10-Jan-2024.) |
⊢ Ⅎ𝑦𝐶 & ⊢ Ⅎ𝑥𝐷 & ⊢ (𝑥 = 𝑦 → 𝐶 = 𝐷) ⇒ ⊢ ⦋𝐴 / 𝑥⦌𝐶 = ⦋𝐴 / 𝑦⦌𝐷 | ||
10-Jan-2024 | cbvsbcw 2964 | Version of cbvsbc 2965 with a disjoint variable condition. (Contributed by Gino Giotto, 10-Jan-2024.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ([𝐴 / 𝑥]𝜑 ↔ [𝐴 / 𝑦]𝜓) | ||
10-Jan-2024 | nfralw 2494 | Bound-variable hypothesis builder for restricted quantification. See nfralya 2497 for a version with 𝑦 and 𝐴 distinct instead of 𝑥 and 𝑦. (Contributed by NM, 1-Sep-1999.) (Revised by Gino Giotto, 10-Jan-2024.) |
⊢ Ⅎ𝑥𝐴 & ⊢ Ⅎ𝑥𝜑 ⇒ ⊢ Ⅎ𝑥∀𝑦 ∈ 𝐴 𝜑 | ||
10-Jan-2024 | nfraldw 2489 | Not-free for restricted universal quantification where 𝑥 and 𝑦 are distinct. See nfraldya 2492 for a version with 𝑦 and 𝐴 distinct instead. (Contributed by NM, 15-Feb-2013.) (Revised by Gino Giotto, 10-Jan-2024.) |
⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝐴) & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥∀𝑦 ∈ 𝐴 𝜓) | ||
10-Jan-2024 | nfabdw 2318 | Bound-variable hypothesis builder for a class abstraction. Version of nfabd 2319 with a disjoint variable condition. (Contributed by Mario Carneiro, 8-Oct-2016.) (Revised by Gino Giotto, 10-Jan-2024.) |
⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥{𝑦 ∣ 𝜓}) | ||
10-Jan-2024 | cbv2w 1730 | Rule used to change bound variables, using implicit substitution. Version of cbv2 1729 with a disjoint variable condition. (Contributed by NM, 5-Aug-1993.) (Revised by Gino Giotto, 10-Jan-2024.) |
⊢ Ⅎ𝑥𝜑 & ⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑦𝜓) & ⊢ (𝜑 → Ⅎ𝑥𝜒) & ⊢ (𝜑 → (𝑥 = 𝑦 → (𝜓 ↔ 𝜒))) ⇒ ⊢ (𝜑 → (∀𝑥𝜓 ↔ ∀𝑦𝜒)) | ||
9-Jan-2024 | suplocexprlemloc 7642 | Lemma for suplocexpr 7646. The putative supremum is located. (Contributed by Jim Kingdon, 9-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) & ⊢ 𝐵 = 〈∪ (1st “ 𝐴), {𝑢 ∈ Q ∣ ∃𝑤 ∈ ∩ (2nd “ 𝐴)𝑤 <Q 𝑢}〉 ⇒ ⊢ (𝜑 → ∀𝑞 ∈ Q ∀𝑟 ∈ Q (𝑞 <Q 𝑟 → (𝑞 ∈ ∪ (1st “ 𝐴) ∨ 𝑟 ∈ (2nd ‘𝐵)))) | ||
9-Jan-2024 | suplocexprlemdisj 7641 | Lemma for suplocexpr 7646. The putative supremum is disjoint. (Contributed by Jim Kingdon, 9-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) & ⊢ 𝐵 = 〈∪ (1st “ 𝐴), {𝑢 ∈ Q ∣ ∃𝑤 ∈ ∩ (2nd “ 𝐴)𝑤 <Q 𝑢}〉 ⇒ ⊢ (𝜑 → ∀𝑞 ∈ Q ¬ (𝑞 ∈ ∪ (1st “ 𝐴) ∧ 𝑞 ∈ (2nd ‘𝐵))) | ||
9-Jan-2024 | suplocexprlemru 7640 | Lemma for suplocexpr 7646. The upper cut of the putative supremum is rounded. (Contributed by Jim Kingdon, 9-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) & ⊢ 𝐵 = 〈∪ (1st “ 𝐴), {𝑢 ∈ Q ∣ ∃𝑤 ∈ ∩ (2nd “ 𝐴)𝑤 <Q 𝑢}〉 ⇒ ⊢ (𝜑 → ∀𝑟 ∈ Q (𝑟 ∈ (2nd ‘𝐵) ↔ ∃𝑞 ∈ Q (𝑞 <Q 𝑟 ∧ 𝑞 ∈ (2nd ‘𝐵)))) | ||
9-Jan-2024 | suplocexprlemrl 7638 | Lemma for suplocexpr 7646. The lower cut of the putative supremum is rounded. (Contributed by Jim Kingdon, 9-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) ⇒ ⊢ (𝜑 → ∀𝑞 ∈ Q (𝑞 ∈ ∪ (1st “ 𝐴) ↔ ∃𝑟 ∈ Q (𝑞 <Q 𝑟 ∧ 𝑟 ∈ ∪ (1st “ 𝐴)))) | ||
9-Jan-2024 | suplocexprlem2b 7635 | Lemma for suplocexpr 7646. Expression for the lower cut of the putative supremum. (Contributed by Jim Kingdon, 9-Jan-2024.) |
⊢ 𝐵 = 〈∪ (1st “ 𝐴), {𝑢 ∈ Q ∣ ∃𝑤 ∈ ∩ (2nd “ 𝐴)𝑤 <Q 𝑢}〉 ⇒ ⊢ (𝐴 ⊆ P → (2nd ‘𝐵) = {𝑢 ∈ Q ∣ ∃𝑤 ∈ ∩ (2nd “ 𝐴)𝑤 <Q 𝑢}) | ||
9-Jan-2024 | suplocexprlemell 7634 | Lemma for suplocexpr 7646. Membership in the lower cut of the putative supremum. (Contributed by Jim Kingdon, 9-Jan-2024.) |
⊢ (𝐵 ∈ ∪ (1st “ 𝐴) ↔ ∃𝑥 ∈ 𝐴 𝐵 ∈ (1st ‘𝑥)) | ||
7-Jan-2024 | suplocexpr 7646 | An inhabited, bounded-above, located set of positive reals has a supremum. (Contributed by Jim Kingdon, 7-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ P (∀𝑦 ∈ 𝐴 ¬ 𝑥<P 𝑦 ∧ ∀𝑦 ∈ P (𝑦<P 𝑥 → ∃𝑧 ∈ 𝐴 𝑦<P 𝑧))) | ||
7-Jan-2024 | suplocexprlemex 7643 | Lemma for suplocexpr 7646. The putative supremum is a positive real. (Contributed by Jim Kingdon, 7-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) & ⊢ 𝐵 = 〈∪ (1st “ 𝐴), {𝑢 ∈ Q ∣ ∃𝑤 ∈ ∩ (2nd “ 𝐴)𝑤 <Q 𝑢}〉 ⇒ ⊢ (𝜑 → 𝐵 ∈ P) | ||
7-Jan-2024 | suplocexprlemmu 7639 | Lemma for suplocexpr 7646. The upper cut of the putative supremum is inhabited. (Contributed by Jim Kingdon, 7-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) & ⊢ 𝐵 = 〈∪ (1st “ 𝐴), {𝑢 ∈ Q ∣ ∃𝑤 ∈ ∩ (2nd “ 𝐴)𝑤 <Q 𝑢}〉 ⇒ ⊢ (𝜑 → ∃𝑠 ∈ Q 𝑠 ∈ (2nd ‘𝐵)) | ||
7-Jan-2024 | suplocexprlemml 7637 | Lemma for suplocexpr 7646. The lower cut of the putative supremum is inhabited. (Contributed by Jim Kingdon, 7-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) ⇒ ⊢ (𝜑 → ∃𝑠 ∈ Q 𝑠 ∈ ∪ (1st “ 𝐴)) | ||
7-Jan-2024 | suplocexprlemss 7636 | Lemma for suplocexpr 7646. 𝐴 is a set of positive reals. (Contributed by Jim Kingdon, 7-Jan-2024.) |
⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∃𝑥 ∈ P ∀𝑦 ∈ 𝐴 𝑦<P 𝑥) & ⊢ (𝜑 → ∀𝑥 ∈ P ∀𝑦 ∈ P (𝑥<P 𝑦 → (∃𝑧 ∈ 𝐴 𝑥<P 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧<P 𝑦))) ⇒ ⊢ (𝜑 → 𝐴 ⊆ P) | ||
5-Jan-2024 | dedekindicclemicc 13052 | Lemma for dedekindicc 13053. Same as dedekindicc 13053, except that we merely show 𝑥 to be an element of (𝐴[,]𝐵). Later we will strengthen that to (𝐴(,)𝐵). (Contributed by Jim Kingdon, 5-Jan-2024.) |
⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 < 𝐵) ⇒ ⊢ (𝜑 → ∃!𝑥 ∈ (𝐴[,]𝐵)(∀𝑞 ∈ 𝐿 𝑞 < 𝑥 ∧ ∀𝑟 ∈ 𝑈 𝑥 < 𝑟)) | ||
5-Jan-2024 | dedekindeu 13043 | A Dedekind cut identifies a unique real number. Similar to df-inp 7387 except that the the Dedekind cut is formed by sets of reals (rather than positive rationals). But in both cases the defining property of a Dedekind cut is that it is inhabited (bounded), rounded, disjoint, and located. (Contributed by Jim Kingdon, 5-Jan-2024.) |
⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∃!𝑥 ∈ ℝ (∀𝑞 ∈ 𝐿 𝑞 < 𝑥 ∧ ∀𝑟 ∈ 𝑈 𝑥 < 𝑟)) | ||
31-Dec-2023 | dvmptsubcn 13127 | Function-builder for derivative, subtraction rule. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 31-Dec-2023.) |
⊢ ((𝜑 ∧ 𝑥 ∈ ℂ) → 𝐴 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑥 ∈ ℂ) → 𝐵 ∈ 𝑉) & ⊢ (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐴)) = (𝑥 ∈ ℂ ↦ 𝐵)) & ⊢ ((𝜑 ∧ 𝑥 ∈ ℂ) → 𝐶 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑥 ∈ ℂ) → 𝐷 ∈ 𝑊) & ⊢ (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐶)) = (𝑥 ∈ ℂ ↦ 𝐷)) ⇒ ⊢ (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ (𝐴 − 𝐶))) = (𝑥 ∈ ℂ ↦ (𝐵 − 𝐷))) | ||
31-Dec-2023 | dvmptnegcn 13126 | Function-builder for derivative, product rule for negatives. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 31-Dec-2023.) |
⊢ ((𝜑 ∧ 𝑥 ∈ ℂ) → 𝐴 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑥 ∈ ℂ) → 𝐵 ∈ 𝑉) & ⊢ (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐴)) = (𝑥 ∈ ℂ ↦ 𝐵)) ⇒ ⊢ (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ -𝐴)) = (𝑥 ∈ ℂ ↦ -𝐵)) | ||
31-Dec-2023 | dvmptcmulcn 13125 | Function-builder for derivative, product rule for constant multiplier. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 31-Dec-2023.) |
⊢ ((𝜑 ∧ 𝑥 ∈ ℂ) → 𝐴 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑥 ∈ ℂ) → 𝐵 ∈ 𝑉) & ⊢ (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐴)) = (𝑥 ∈ ℂ ↦ 𝐵)) & ⊢ (𝜑 → 𝐶 ∈ ℂ) ⇒ ⊢ (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ (𝐶 · 𝐴))) = (𝑥 ∈ ℂ ↦ (𝐶 · 𝐵))) | ||
31-Dec-2023 | brm 4015 | If two sets are in a binary relation, the relation is inhabited. (Contributed by Jim Kingdon, 31-Dec-2023.) |
⊢ (𝐴𝑅𝐵 → ∃𝑥 𝑥 ∈ 𝑅) | ||
30-Dec-2023 | dvmptccn 13121 | Function-builder for derivative: derivative of a constant. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 30-Dec-2023.) |
⊢ (𝜑 → 𝐴 ∈ ℂ) ⇒ ⊢ (𝜑 → (ℂ D (𝑥 ∈ ℂ ↦ 𝐴)) = (𝑥 ∈ ℂ ↦ 0)) | ||
30-Dec-2023 | dvmptidcn 13120 | Function-builder for derivative: derivative of the identity. (Contributed by Mario Carneiro, 1-Sep-2014.) (Revised by Jim Kingdon, 30-Dec-2023.) |
⊢ (ℂ D (𝑥 ∈ ℂ ↦ 𝑥)) = (𝑥 ∈ ℂ ↦ 1) | ||
25-Dec-2023 | ctfoex 7063 | A countable class is a set. (Contributed by Jim Kingdon, 25-Dec-2023.) |
⊢ (∃𝑓 𝑓:ω–onto→(𝐴 ⊔ 1o) → 𝐴 ∈ V) | ||
23-Dec-2023 | enct 12204 | Countability is invariant relative to equinumerosity. (Contributed by Jim Kingdon, 23-Dec-2023.) |
⊢ (𝐴 ≈ 𝐵 → (∃𝑓 𝑓:ω–onto→(𝐴 ⊔ 1o) ↔ ∃𝑔 𝑔:ω–onto→(𝐵 ⊔ 1o))) | ||
23-Dec-2023 | enctlem 12203 | Lemma for enct 12204. One direction of the biconditional. (Contributed by Jim Kingdon, 23-Dec-2023.) |
⊢ (𝐴 ≈ 𝐵 → (∃𝑓 𝑓:ω–onto→(𝐴 ⊔ 1o) → ∃𝑔 𝑔:ω–onto→(𝐵 ⊔ 1o))) | ||
23-Dec-2023 | omct 7062 | ω is countable. (Contributed by Jim Kingdon, 23-Dec-2023.) |
⊢ ∃𝑓 𝑓:ω–onto→(ω ⊔ 1o) | ||
21-Dec-2023 | dvcoapbr 13113 | The chain rule for derivatives at a point. The 𝑢 # 𝐶 → (𝐺‘𝑢) # (𝐺‘𝐶) hypothesis constrains what functions work for 𝐺. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Jim Kingdon, 21-Dec-2023.) |
⊢ (𝜑 → 𝐹:𝑋⟶ℂ) & ⊢ (𝜑 → 𝑋 ⊆ 𝑆) & ⊢ (𝜑 → 𝐺:𝑌⟶𝑋) & ⊢ (𝜑 → 𝑌 ⊆ 𝑇) & ⊢ (𝜑 → ∀𝑢 ∈ 𝑌 (𝑢 # 𝐶 → (𝐺‘𝑢) # (𝐺‘𝐶))) & ⊢ (𝜑 → 𝑆 ⊆ ℂ) & ⊢ (𝜑 → 𝑇 ⊆ ℂ) & ⊢ (𝜑 → (𝐺‘𝐶)(𝑆 D 𝐹)𝐾) & ⊢ (𝜑 → 𝐶(𝑇 D 𝐺)𝐿) & ⊢ 𝐽 = (MetOpen‘(abs ∘ − )) ⇒ ⊢ (𝜑 → 𝐶(𝑇 D (𝐹 ∘ 𝐺))(𝐾 · 𝐿)) | ||
19-Dec-2023 | apsscn 8523 | The points apart from a given point are complex numbers. (Contributed by Jim Kingdon, 19-Dec-2023.) |
⊢ {𝑥 ∈ 𝐴 ∣ 𝑥 # 𝐵} ⊆ ℂ | ||
19-Dec-2023 | aprcl 8522 | Reverse closure for apartness. (Contributed by Jim Kingdon, 19-Dec-2023.) |
⊢ (𝐴 # 𝐵 → (𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ)) | ||
18-Dec-2023 | limccoap 13089 | Composition of two limits. This theorem is only usable in the case where 𝑥 # 𝑋 implies R(x) # 𝐶 so it is less general than might appear at first. (Contributed by Mario Carneiro, 29-Dec-2016.) (Revised by Jim Kingdon, 18-Dec-2023.) |
⊢ ((𝜑 ∧ 𝑥 ∈ {𝑤 ∈ 𝐴 ∣ 𝑤 # 𝑋}) → 𝑅 ∈ {𝑤 ∈ 𝐵 ∣ 𝑤 # 𝐶}) & ⊢ ((𝜑 ∧ 𝑦 ∈ {𝑤 ∈ 𝐵 ∣ 𝑤 # 𝐶}) → 𝑆 ∈ ℂ) & ⊢ (𝜑 → 𝐶 ∈ ((𝑥 ∈ {𝑤 ∈ 𝐴 ∣ 𝑤 # 𝑋} ↦ 𝑅) limℂ 𝑋)) & ⊢ (𝜑 → 𝐷 ∈ ((𝑦 ∈ {𝑤 ∈ 𝐵 ∣ 𝑤 # 𝐶} ↦ 𝑆) limℂ 𝐶)) & ⊢ (𝑦 = 𝑅 → 𝑆 = 𝑇) ⇒ ⊢ (𝜑 → 𝐷 ∈ ((𝑥 ∈ {𝑤 ∈ 𝐴 ∣ 𝑤 # 𝑋} ↦ 𝑇) limℂ 𝑋)) | ||
16-Dec-2023 | cnreim 10882 | Complex apartness in terms of real and imaginary parts. See also apreim 8479 which is similar but with different notation. (Contributed by Jim Kingdon, 16-Dec-2023.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → (𝐴 # 𝐵 ↔ ((ℜ‘𝐴) # (ℜ‘𝐵) ∨ (ℑ‘𝐴) # (ℑ‘𝐵)))) | ||
14-Dec-2023 | cnopnap 13036 | The complex numbers apart from a given complex number form an open set. (Contributed by Jim Kingdon, 14-Dec-2023.) |
⊢ (𝐴 ∈ ℂ → {𝑤 ∈ ℂ ∣ 𝑤 # 𝐴} ∈ (MetOpen‘(abs ∘ − ))) | ||
14-Dec-2023 | cnovex 12638 | The class of all continuous functions from a topology to another is a set. (Contributed by Jim Kingdon, 14-Dec-2023.) |
⊢ ((𝐽 ∈ Top ∧ 𝐾 ∈ Top) → (𝐽 Cn 𝐾) ∈ V) | ||
13-Dec-2023 | reopnap 12980 | The real numbers apart from a given real number form an open set. (Contributed by Jim Kingdon, 13-Dec-2023.) |
⊢ (𝐴 ∈ ℝ → {𝑤 ∈ ℝ ∣ 𝑤 # 𝐴} ∈ (topGen‘ran (,))) | ||
12-Dec-2023 | cnopncntop 12979 | The set of complex numbers is open with respect to the standard topology on complex numbers. (Contributed by Glauco Siliprandi, 11-Dec-2019.) (Revised by Jim Kingdon, 12-Dec-2023.) |
⊢ ℂ ∈ (MetOpen‘(abs ∘ − )) | ||
12-Dec-2023 | unicntopcntop 12978 | The underlying set of the standard topology on the complex numbers is the set of complex numbers. (Contributed by Glauco Siliprandi, 11-Dec-2019.) (Revised by Jim Kingdon, 12-Dec-2023.) |
⊢ ℂ = ∪ (MetOpen‘(abs ∘ − )) | ||
4-Dec-2023 | bj-pm2.18st 13365 | Clavius law for stable formulas. See pm2.18dc 841. (Contributed by BJ, 4-Dec-2023.) |
⊢ (STAB 𝜑 → ((¬ 𝜑 → 𝜑) → 𝜑)) | ||
4-Dec-2023 | bj-nnclavius 13355 | Clavius law with doubly negated consequent. (Contributed by BJ, 4-Dec-2023.) |
⊢ ((¬ 𝜑 → 𝜑) → ¬ ¬ 𝜑) | ||
2-Dec-2023 | dvmulxx 13110 | The product rule for derivatives at a point. For the (more general) relation version, see dvmulxxbr 13108. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Jim Kingdon, 2-Dec-2023.) |
⊢ (𝜑 → 𝐹:𝑋⟶ℂ) & ⊢ (𝜑 → 𝑋 ⊆ 𝑆) & ⊢ (𝜑 → 𝐺:𝑋⟶ℂ) & ⊢ (𝜑 → 𝑆 ∈ {ℝ, ℂ}) & ⊢ (𝜑 → 𝐶 ∈ dom (𝑆 D 𝐹)) & ⊢ (𝜑 → 𝐶 ∈ dom (𝑆 D 𝐺)) ⇒ ⊢ (𝜑 → ((𝑆 D (𝐹 ∘𝑓 · 𝐺))‘𝐶) = ((((𝑆 D 𝐹)‘𝐶) · (𝐺‘𝐶)) + (((𝑆 D 𝐺)‘𝐶) · (𝐹‘𝐶)))) | ||
1-Dec-2023 | dvmulxxbr 13108 | The product rule for derivatives at a point. For the (simpler but more limited) function version, see dvmulxx 13110. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Jim Kingdon, 1-Dec-2023.) |
⊢ (𝜑 → 𝐹:𝑋⟶ℂ) & ⊢ (𝜑 → 𝑋 ⊆ 𝑆) & ⊢ (𝜑 → 𝐺:𝑋⟶ℂ) & ⊢ (𝜑 → 𝑆 ⊆ ℂ) & ⊢ (𝜑 → 𝐶(𝑆 D 𝐹)𝐾) & ⊢ (𝜑 → 𝐶(𝑆 D 𝐺)𝐿) & ⊢ 𝐽 = (MetOpen‘(abs ∘ − )) ⇒ ⊢ (𝜑 → 𝐶(𝑆 D (𝐹 ∘𝑓 · 𝐺))((𝐾 · (𝐺‘𝐶)) + (𝐿 · (𝐹‘𝐶)))) | ||
29-Nov-2023 | subctctexmid 13615 | If every subcountable set is countable and Markov's principle holds, excluded middle follows. Proposition 2.6 of [BauerSwan], p. 14:4. The proof is taken from that paper. (Contributed by Jim Kingdon, 29-Nov-2023.) |
⊢ (𝜑 → ∀𝑥(∃𝑠(𝑠 ⊆ ω ∧ ∃𝑓 𝑓:𝑠–onto→𝑥) → ∃𝑔 𝑔:ω–onto→(𝑥 ⊔ 1o))) & ⊢ (𝜑 → ω ∈ Markov) ⇒ ⊢ (𝜑 → EXMID) | ||
29-Nov-2023 | ismkvnex 7099 | The predicate of being Markov stated in terms of double negation and comparison with 1o. (Contributed by Jim Kingdon, 29-Nov-2023.) |
⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ Markov ↔ ∀𝑓 ∈ (2o ↑𝑚 𝐴)(¬ ¬ ∃𝑥 ∈ 𝐴 (𝑓‘𝑥) = 1o → ∃𝑥 ∈ 𝐴 (𝑓‘𝑥) = 1o))) | ||
28-Nov-2023 | exmid1stab 13614 | If any proposition is stable, excluded middle follows. We are thinking of 𝑥 as a proposition and 𝑥 = {∅} as "x is true". (Contributed by Jim Kingdon, 28-Nov-2023.) |
⊢ ((𝜑 ∧ 𝑥 ⊆ {∅}) → STAB 𝑥 = {∅}) ⇒ ⊢ (𝜑 → EXMID) | ||
28-Nov-2023 | ccfunen 7185 | Existence of a choice function for a countably infinite set. (Contributed by Jim Kingdon, 28-Nov-2023.) |
⊢ (𝜑 → CCHOICE) & ⊢ (𝜑 → 𝐴 ≈ ω) & ⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∃𝑤 𝑤 ∈ 𝑥) ⇒ ⊢ (𝜑 → ∃𝑓(𝑓 Fn 𝐴 ∧ ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) ∈ 𝑥)) | ||
27-Nov-2023 | df-cc 7184 | The expression CCHOICE will be used as a readable shorthand for any form of countable choice, analogous to df-ac 7142 for full choice. (Contributed by Jim Kingdon, 27-Nov-2023.) |
⊢ (CCHOICE ↔ ∀𝑥(dom 𝑥 ≈ ω → ∃𝑓(𝑓 ⊆ 𝑥 ∧ 𝑓 Fn dom 𝑥))) | ||
26-Nov-2023 | offeq 6046 | Convert an identity of the operation to the analogous identity on the function operation. (Contributed by Jim Kingdon, 26-Nov-2023.) |
⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑇)) → (𝑥𝑅𝑦) ∈ 𝑈) & ⊢ (𝜑 → 𝐹:𝐴⟶𝑆) & ⊢ (𝜑 → 𝐺:𝐵⟶𝑇) & ⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → 𝐵 ∈ 𝑊) & ⊢ (𝐴 ∩ 𝐵) = 𝐶 & ⊢ (𝜑 → 𝐻:𝐶⟶𝑈) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → (𝐹‘𝑥) = 𝐷) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐵) → (𝐺‘𝑥) = 𝐸) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐶) → (𝐷𝑅𝐸) = (𝐻‘𝑥)) ⇒ ⊢ (𝜑 → (𝐹 ∘𝑓 𝑅𝐺) = 𝐻) | ||
25-Nov-2023 | dvaddxx 13109 | The sum rule for derivatives at a point. For the (more general) relation version, see dvaddxxbr 13107. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Jim Kingdon, 25-Nov-2023.) |
⊢ (𝜑 → 𝐹:𝑋⟶ℂ) & ⊢ (𝜑 → 𝑋 ⊆ 𝑆) & ⊢ (𝜑 → 𝐺:𝑋⟶ℂ) & ⊢ (𝜑 → 𝑆 ∈ {ℝ, ℂ}) & ⊢ (𝜑 → 𝐶 ∈ dom (𝑆 D 𝐹)) & ⊢ (𝜑 → 𝐶 ∈ dom (𝑆 D 𝐺)) ⇒ ⊢ (𝜑 → ((𝑆 D (𝐹 ∘𝑓 + 𝐺))‘𝐶) = (((𝑆 D 𝐹)‘𝐶) + ((𝑆 D 𝐺)‘𝐶))) | ||
25-Nov-2023 | dvaddxxbr 13107 | The sum rule for derivatives at a point. That is, if the derivative of 𝐹 at 𝐶 is 𝐾 and the derivative of 𝐺 at 𝐶 is 𝐿, then the derivative of the pointwise sum of those two functions at 𝐶 is 𝐾 + 𝐿. (Contributed by Mario Carneiro, 9-Aug-2014.) (Revised by Jim Kingdon, 25-Nov-2023.) |
⊢ (𝜑 → 𝐹:𝑋⟶ℂ) & ⊢ (𝜑 → 𝑋 ⊆ 𝑆) & ⊢ (𝜑 → 𝐺:𝑋⟶ℂ) & ⊢ (𝜑 → 𝑆 ⊆ ℂ) & ⊢ (𝜑 → 𝐶(𝑆 D 𝐹)𝐾) & ⊢ (𝜑 → 𝐶(𝑆 D 𝐺)𝐿) & ⊢ 𝐽 = (MetOpen‘(abs ∘ − )) ⇒ ⊢ (𝜑 → 𝐶(𝑆 D (𝐹 ∘𝑓 + 𝐺))(𝐾 + 𝐿)) | ||
25-Nov-2023 | dcnn 834 | Decidability of the negation of a proposition is equivalent to decidability of its double negation. See also dcn 828. The relation between dcn 828 and dcnn 834 is analogous to that between notnot 619 and notnotnot 624 (and directly stems from it). Using the notion of "testable proposition" (proposition whose negation is decidable), dcnn 834 means that a proposition is testable if and only if its negation is testable, and dcn 828 means that decidability implies testability. (Contributed by David A. Wheeler, 6-Dec-2018.) (Proof shortened by BJ, 25-Nov-2023.) |
⊢ (DECID ¬ 𝜑 ↔ DECID ¬ ¬ 𝜑) | ||
24-Nov-2023 | bj-dcst 13376 | Stability of a proposition is decidable if and only if that proposition is stable. (Contributed by BJ, 24-Nov-2023.) |
⊢ (DECID STAB 𝜑 ↔ STAB 𝜑) | ||
24-Nov-2023 | bj-nnbidc 13371 | If a formula is not refutable, then it is decidable if and only if it is provable. See also comment of bj-nnbist 13360. (Contributed by BJ, 24-Nov-2023.) |
⊢ (¬ ¬ 𝜑 → (DECID 𝜑 ↔ 𝜑)) | ||
24-Nov-2023 | bj-dcstab 13370 | A decidable formula is stable. (Contributed by BJ, 24-Nov-2023.) (Proof modification is discouraged.) |
⊢ (DECID 𝜑 → STAB 𝜑) | ||
24-Nov-2023 | bj-fadc 13368 | A refutable formula is decidable. (Contributed by BJ, 24-Nov-2023.) |
⊢ (¬ 𝜑 → DECID 𝜑) | ||
24-Nov-2023 | bj-trdc 13366 | A provable formula is decidable. (Contributed by BJ, 24-Nov-2023.) |
⊢ (𝜑 → DECID 𝜑) | ||
24-Nov-2023 | bj-stal 13364 | The universal quantification of stable formula is stable. See bj-stim 13361 for implication, stabnot 819 for negation, and bj-stan 13362 for conjunction. (Contributed by BJ, 24-Nov-2023.) |
⊢ (∀𝑥STAB 𝜑 → STAB ∀𝑥𝜑) | ||
24-Nov-2023 | bj-stand 13363 | The conjunction of two stable formulas is stable. Deduction form of bj-stan 13362. Its proof is shorter, so one could prove it first and then bj-stan 13362 from it, the usual way. (Contributed by BJ, 24-Nov-2023.) (Proof modification is discouraged.) |
⊢ (𝜑 → STAB 𝜓) & ⊢ (𝜑 → STAB 𝜒) ⇒ ⊢ (𝜑 → STAB (𝜓 ∧ 𝜒)) | ||
24-Nov-2023 | bj-stan 13362 | The conjunction of two stable formulas is stable. See bj-stim 13361 for implication, stabnot 819 for negation, and bj-stal 13364 for universal quantification. (Contributed by BJ, 24-Nov-2023.) |
⊢ ((STAB 𝜑 ∧ STAB 𝜓) → STAB (𝜑 ∧ 𝜓)) | ||
24-Nov-2023 | bj-stim 13361 | A conjunction with a stable consequent is stable. See stabnot 819 for negation and bj-stan 13362 for conjunction. (Contributed by BJ, 24-Nov-2023.) |
⊢ (STAB 𝜓 → STAB (𝜑 → 𝜓)) | ||
24-Nov-2023 | bj-nnbist 13360 | If a formula is not refutable, then it is stable if and only if it is provable. By double-negation translation, if 𝜑 is a classical tautology, then ¬ ¬ 𝜑 is an intuitionistic tautology. Therefore, if 𝜑 is a classical tautology, then 𝜑 is intuitionistically equivalent to its stability (and to its decidability, see bj-nnbidc 13371). (Contributed by BJ, 24-Nov-2023.) |
⊢ (¬ ¬ 𝜑 → (STAB 𝜑 ↔ 𝜑)) | ||
24-Nov-2023 | bj-fast 13358 | A refutable formula is stable. (Contributed by BJ, 24-Nov-2023.) |
⊢ (¬ 𝜑 → STAB 𝜑) | ||
24-Nov-2023 | bj-trst 13356 | A provable formula is stable. (Contributed by BJ, 24-Nov-2023.) |
⊢ (𝜑 → STAB 𝜑) | ||
24-Nov-2023 | bj-nnan 13354 | The double negation of a conjunction implies the conjunction of the double negations. (Contributed by BJ, 24-Nov-2023.) |
⊢ (¬ ¬ (𝜑 ∧ 𝜓) → (¬ ¬ 𝜑 ∧ ¬ ¬ 𝜓)) | ||
24-Nov-2023 | bj-nnim 13353 | The double negation of an implication implies the implication with the consequent doubly negated. (Contributed by BJ, 24-Nov-2023.) |
⊢ (¬ ¬ (𝜑 → 𝜓) → (𝜑 → ¬ ¬ 𝜓)) | ||
24-Nov-2023 | bj-nnsn 13351 | As far as implying a negated formula is concerned, a formula is equivalent to its double negation. (Contributed by BJ, 24-Nov-2023.) |
⊢ ((𝜑 → ¬ 𝜓) ↔ (¬ ¬ 𝜑 → ¬ 𝜓)) | ||
24-Nov-2023 | nnal 1629 | The double negation of a universal quantification implies the universal quantification of the double negation. (Contributed by BJ, 24-Nov-2023.) |
⊢ (¬ ¬ ∀𝑥𝜑 → ∀𝑥 ¬ ¬ 𝜑) | ||
22-Nov-2023 | ofvalg 6042 | Evaluate a function operation at a point. (Contributed by Mario Carneiro, 20-Jul-2014.) (Revised by Jim Kingdon, 22-Nov-2023.) |
⊢ (𝜑 → 𝐹 Fn 𝐴) & ⊢ (𝜑 → 𝐺 Fn 𝐵) & ⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → 𝐵 ∈ 𝑊) & ⊢ (𝐴 ∩ 𝐵) = 𝑆 & ⊢ ((𝜑 ∧ 𝑋 ∈ 𝐴) → (𝐹‘𝑋) = 𝐶) & ⊢ ((𝜑 ∧ 𝑋 ∈ 𝐵) → (𝐺‘𝑋) = 𝐷) & ⊢ ((𝜑 ∧ 𝑋 ∈ 𝑆) → (𝐶𝑅𝐷) ∈ 𝑈) ⇒ ⊢ ((𝜑 ∧ 𝑋 ∈ 𝑆) → ((𝐹 ∘𝑓 𝑅𝐺)‘𝑋) = (𝐶𝑅𝐷)) | ||
21-Nov-2023 | exmidac 7145 | The axiom of choice implies excluded middle. See acexmid 5824 for more discussion of this theorem and a way of stating it without using CHOICE or EXMID. (Contributed by Jim Kingdon, 21-Nov-2023.) |
⊢ (CHOICE → EXMID) | ||
21-Nov-2023 | exmidaclem 7144 | Lemma for exmidac 7145. The result, with a few hypotheses to break out commonly used expressions. (Contributed by Jim Kingdon, 21-Nov-2023.) |
⊢ 𝐴 = {𝑥 ∈ {∅, {∅}} ∣ (𝑥 = ∅ ∨ 𝑦 = {∅})} & ⊢ 𝐵 = {𝑥 ∈ {∅, {∅}} ∣ (𝑥 = {∅} ∨ 𝑦 = {∅})} & ⊢ 𝐶 = {𝐴, 𝐵} ⇒ ⊢ (CHOICE → EXMID) | ||
21-Nov-2023 | exmid1dc 4162 | A convenience theorem for proving that something implies EXMID. Think of this as an alternative to using a proposition, as in proofs like undifexmid 4155 or ordtriexmid 4481. In this context 𝑥 = {∅} can be thought of as "x is true". (Contributed by Jim Kingdon, 21-Nov-2023.) |
⊢ ((𝜑 ∧ 𝑥 ⊆ {∅}) → DECID 𝑥 = {∅}) ⇒ ⊢ (𝜑 → EXMID) | ||
20-Nov-2023 | acfun 7143 | A convenient form of choice. The goal here is to state choice as the existence of a choice function on a set of inhabited sets, while making full use of our notation around functions and function values. (Contributed by Jim Kingdon, 20-Nov-2023.) |
⊢ (𝜑 → CHOICE) & ⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∃𝑤 𝑤 ∈ 𝑥) ⇒ ⊢ (𝜑 → ∃𝑓(𝑓 Fn 𝐴 ∧ ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) ∈ 𝑥)) | ||
18-Nov-2023 | condc 839 |
Contraposition of a decidable proposition.
This theorem swaps or "transposes" the order of the consequents when negation is removed. An informal example is that the statement "if there are no clouds in the sky, it is not raining" implies the statement "if it is raining, there are clouds in the sky." This theorem (without the decidability condition, of course) is called Transp or "the principle of transposition" in Principia Mathematica (Theorem *2.17 of [WhiteheadRussell] p. 103) and is Axiom A3 of [Margaris] p. 49. We will also use the term "contraposition" for this principle, although the reader is advised that in the field of philosophical logic, "contraposition" has a different technical meaning. (Contributed by Jim Kingdon, 13-Mar-2018.) (Proof shortened by BJ, 18-Nov-2023.) |
⊢ (DECID 𝜑 → ((¬ 𝜑 → ¬ 𝜓) → (𝜓 → 𝜑))) | ||
18-Nov-2023 | const 838 | Contraposition of a stable proposition. See comment of condc 839. (Contributed by BJ, 18-Nov-2023.) |
⊢ (STAB 𝜑 → ((¬ 𝜑 → ¬ 𝜓) → (𝜓 → 𝜑))) | ||
18-Nov-2023 | stdcn 833 | A formula is stable if and only if the decidability of its negation implies its decidability. Note that the right-hand side of this biconditional is the converse of dcn 828. (Contributed by BJ, 18-Nov-2023.) |
⊢ (STAB 𝜑 ↔ (DECID ¬ 𝜑 → DECID 𝜑)) | ||
17-Nov-2023 | cnplimclemr 13080 | Lemma for cnplimccntop 13081. The reverse direction. (Contributed by Mario Carneiro and Jim Kingdon, 17-Nov-2023.) |
⊢ 𝐾 = (MetOpen‘(abs ∘ − )) & ⊢ 𝐽 = (𝐾 ↾t 𝐴) & ⊢ (𝜑 → 𝐴 ⊆ ℂ) & ⊢ (𝜑 → 𝐹:𝐴⟶ℂ) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → (𝐹‘𝐵) ∈ (𝐹 limℂ 𝐵)) ⇒ ⊢ (𝜑 → 𝐹 ∈ ((𝐽 CnP 𝐾)‘𝐵)) | ||
17-Nov-2023 | cnplimclemle 13079 | Lemma for cnplimccntop 13081. Satisfying the epsilon condition for continuity. (Contributed by Mario Carneiro and Jim Kingdon, 17-Nov-2023.) |
⊢ 𝐾 = (MetOpen‘(abs ∘ − )) & ⊢ 𝐽 = (𝐾 ↾t 𝐴) & ⊢ (𝜑 → 𝐴 ⊆ ℂ) & ⊢ (𝜑 → 𝐹:𝐴⟶ℂ) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → (𝐹‘𝐵) ∈ (𝐹 limℂ 𝐵)) & ⊢ (𝜑 → 𝐸 ∈ ℝ+) & ⊢ (𝜑 → 𝐷 ∈ ℝ+) & ⊢ (𝜑 → 𝑍 ∈ 𝐴) & ⊢ ((𝜑 ∧ 𝑍 # 𝐵 ∧ (abs‘(𝑍 − 𝐵)) < 𝐷) → (abs‘((𝐹‘𝑍) − (𝐹‘𝐵))) < (𝐸 / 2)) & ⊢ (𝜑 → (abs‘(𝑍 − 𝐵)) < 𝐷) ⇒ ⊢ (𝜑 → (abs‘((𝐹‘𝑍) − (𝐹‘𝐵))) < 𝐸) | ||
14-Nov-2023 | limccnp2cntop 13088 | The image of a convergent sequence under a continuous map is convergent to the image of the original point. Binary operation version. (Contributed by Mario Carneiro, 28-Dec-2016.) (Revised by Jim Kingdon, 14-Nov-2023.) |
⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑅 ∈ 𝑋) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑆 ∈ 𝑌) & ⊢ (𝜑 → 𝑋 ⊆ ℂ) & ⊢ (𝜑 → 𝑌 ⊆ ℂ) & ⊢ 𝐾 = (MetOpen‘(abs ∘ − )) & ⊢ 𝐽 = ((𝐾 ×t 𝐾) ↾t (𝑋 × 𝑌)) & ⊢ (𝜑 → 𝐶 ∈ ((𝑥 ∈ 𝐴 ↦ 𝑅) limℂ 𝐵)) & ⊢ (𝜑 → 𝐷 ∈ ((𝑥 ∈ 𝐴 ↦ 𝑆) limℂ 𝐵)) & ⊢ (𝜑 → 𝐻 ∈ ((𝐽 CnP 𝐾)‘〈𝐶, 𝐷〉)) ⇒ ⊢ (𝜑 → (𝐶𝐻𝐷) ∈ ((𝑥 ∈ 𝐴 ↦ (𝑅𝐻𝑆)) limℂ 𝐵)) | ||
10-Nov-2023 | rpmaxcl 11127 | The maximum of two positive real numbers is a positive real number. (Contributed by Jim Kingdon, 10-Nov-2023.) |
⊢ ((𝐴 ∈ ℝ+ ∧ 𝐵 ∈ ℝ+) → sup({𝐴, 𝐵}, ℝ, < ) ∈ ℝ+) | ||
9-Nov-2023 | limccnp2lem 13087 | Lemma for limccnp2cntop 13088. This is most of the result, expressed in epsilon-delta form, with a large number of hypotheses so that lengthy expressions do not need to be repeated. (Contributed by Jim Kingdon, 9-Nov-2023.) |
⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑅 ∈ 𝑋) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑆 ∈ 𝑌) & ⊢ (𝜑 → 𝑋 ⊆ ℂ) & ⊢ (𝜑 → 𝑌 ⊆ ℂ) & ⊢ 𝐾 = (MetOpen‘(abs ∘ − )) & ⊢ 𝐽 = ((𝐾 ×t 𝐾) ↾t (𝑋 × 𝑌)) & ⊢ (𝜑 → 𝐶 ∈ ((𝑥 ∈ 𝐴 ↦ 𝑅) limℂ 𝐵)) & ⊢ (𝜑 → 𝐷 ∈ ((𝑥 ∈ 𝐴 ↦ 𝑆) limℂ 𝐵)) & ⊢ (𝜑 → 𝐻 ∈ ((𝐽 CnP 𝐾)‘〈𝐶, 𝐷〉)) & ⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → 𝐸 ∈ ℝ+) & ⊢ (𝜑 → 𝐿 ∈ ℝ+) & ⊢ (𝜑 → ∀𝑟 ∈ 𝑋 ∀𝑠 ∈ 𝑌 (((𝐶((abs ∘ − ) ↾ (𝑋 × 𝑋))𝑟) < 𝐿 ∧ (𝐷((abs ∘ − ) ↾ (𝑌 × 𝑌))𝑠) < 𝐿) → ((𝐶𝐻𝐷)(abs ∘ − )(𝑟𝐻𝑠)) < 𝐸)) & ⊢ (𝜑 → 𝐹 ∈ ℝ+) & ⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ((𝑥 # 𝐵 ∧ (abs‘(𝑥 − 𝐵)) < 𝐹) → (abs‘(𝑅 − 𝐶)) < 𝐿)) & ⊢ (𝜑 → 𝐺 ∈ ℝ+) & ⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ((𝑥 # 𝐵 ∧ (abs‘(𝑥 − 𝐵)) < 𝐺) → (abs‘(𝑆 − 𝐷)) < 𝐿)) ⇒ ⊢ (𝜑 → ∃𝑑 ∈ ℝ+ ∀𝑥 ∈ 𝐴 ((𝑥 # 𝐵 ∧ (abs‘(𝑥 − 𝐵)) < 𝑑) → (abs‘((𝑅𝐻𝑆) − (𝐶𝐻𝐷))) < 𝐸)) | ||
4-Nov-2023 | ellimc3apf 13071 | Write the epsilon-delta definition of a limit. (Contributed by Mario Carneiro, 28-Dec-2016.) (Revised by Jim Kingdon, 4-Nov-2023.) |
⊢ (𝜑 → 𝐹:𝐴⟶ℂ) & ⊢ (𝜑 → 𝐴 ⊆ ℂ) & ⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ Ⅎ𝑧𝐹 ⇒ ⊢ (𝜑 → (𝐶 ∈ (𝐹 limℂ 𝐵) ↔ (𝐶 ∈ ℂ ∧ ∀𝑥 ∈ ℝ+ ∃𝑦 ∈ ℝ+ ∀𝑧 ∈ 𝐴 ((𝑧 # 𝐵 ∧ (abs‘(𝑧 − 𝐵)) < 𝑦) → (abs‘((𝐹‘𝑧) − 𝐶)) < 𝑥)))) | ||
3-Nov-2023 | limcmpted 13074 | Express the limit operator for a function defined by a mapping, via epsilon-delta. (Contributed by Jim Kingdon, 3-Nov-2023.) |
⊢ (𝜑 → 𝐴 ⊆ ℂ) & ⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑧 ∈ 𝐴) → 𝐷 ∈ ℂ) ⇒ ⊢ (𝜑 → (𝐶 ∈ ((𝑧 ∈ 𝐴 ↦ 𝐷) limℂ 𝐵) ↔ (𝐶 ∈ ℂ ∧ ∀𝑥 ∈ ℝ+ ∃𝑦 ∈ ℝ+ ∀𝑧 ∈ 𝐴 ((𝑧 # 𝐵 ∧ (abs‘(𝑧 − 𝐵)) < 𝑦) → (abs‘(𝐷 − 𝐶)) < 𝑥)))) | ||
1-Nov-2023 | unct 12213 | The union of two countable sets is countable. Corollary 8.1.20 of [AczelRathjen], p. 75. (Contributed by Jim Kingdon, 1-Nov-2023.) |
⊢ ((∃𝑓 𝑓:ω–onto→(𝐴 ⊔ 1o) ∧ ∃𝑔 𝑔:ω–onto→(𝐵 ⊔ 1o)) → ∃ℎ ℎ:ω–onto→((𝐴 ∪ 𝐵) ⊔ 1o)) | ||
31-Oct-2023 | ctiunct 12211 |
A sequence of enumerations gives an enumeration of the union. We refer
to "sequence of enumerations" rather than "countably many
countable
sets" because the hypothesis provides more than countability for
each
𝐵(𝑥): it refers to 𝐵(𝑥) together with the 𝐺(𝑥)
which enumerates it. Theorem 8.1.19 of [AczelRathjen], p. 74.
For "countably many countable sets" the key hypothesis would be (𝜑 ∧ 𝑥 ∈ 𝐴) → ∃𝑔𝑔:ω–onto→(𝐵 ⊔ 1o). This is almost omiunct 12215 (which uses countable choice) although that is for a countably infinite collection not any countable collection. Compare with the case of two sets instead of countably many, as seen at unct 12213, which says that the union of two countable sets is countable . The proof proceeds by mapping a natural number to a pair of natural numbers (by xpomen 12166) and using the first number to map to an element 𝑥 of 𝐴 and the second number to map to an element of B(x) . In this way we are able to map to every element of ∪ 𝑥 ∈ 𝐴𝐵. Although it would be possible to work directly with countability expressed as 𝐹:ω–onto→(𝐴 ⊔ 1o), we instead use functions from subsets of the natural numbers via ctssdccl 7056 and ctssdc 7058. (Contributed by Jim Kingdon, 31-Oct-2023.) |
⊢ (𝜑 → 𝐹:ω–onto→(𝐴 ⊔ 1o)) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝐺:ω–onto→(𝐵 ⊔ 1o)) ⇒ ⊢ (𝜑 → ∃ℎ ℎ:ω–onto→(∪ 𝑥 ∈ 𝐴 𝐵 ⊔ 1o)) | ||
30-Oct-2023 | ctssdccl 7056 | A mapping from a decidable subset of the natural numbers onto a countable set. This is similar to one direction of ctssdc 7058 but expressed in terms of classes rather than ∃. (Contributed by Jim Kingdon, 30-Oct-2023.) |
⊢ (𝜑 → 𝐹:ω–onto→(𝐴 ⊔ 1o)) & ⊢ 𝑆 = {𝑥 ∈ ω ∣ (𝐹‘𝑥) ∈ (inl “ 𝐴)} & ⊢ 𝐺 = (◡inl ∘ 𝐹) ⇒ ⊢ (𝜑 → (𝑆 ⊆ ω ∧ 𝐺:𝑆–onto→𝐴 ∧ ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑆)) | ||
28-Oct-2023 | ctiunctlemfo 12210 | Lemma for ctiunct 12211. (Contributed by Jim Kingdon, 28-Oct-2023.) |
⊢ (𝜑 → 𝑆 ⊆ ω) & ⊢ (𝜑 → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑆) & ⊢ (𝜑 → 𝐹:𝑆–onto→𝐴) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑇 ⊆ ω) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑇) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝐺:𝑇–onto→𝐵) & ⊢ (𝜑 → 𝐽:ω–1-1-onto→(ω × ω)) & ⊢ 𝑈 = {𝑧 ∈ ω ∣ ((1st ‘(𝐽‘𝑧)) ∈ 𝑆 ∧ (2nd ‘(𝐽‘𝑧)) ∈ ⦋(𝐹‘(1st ‘(𝐽‘𝑧))) / 𝑥⦌𝑇)} & ⊢ 𝐻 = (𝑛 ∈ 𝑈 ↦ (⦋(𝐹‘(1st ‘(𝐽‘𝑛))) / 𝑥⦌𝐺‘(2nd ‘(𝐽‘𝑛)))) & ⊢ Ⅎ𝑥𝐻 & ⊢ Ⅎ𝑥𝑈 ⇒ ⊢ (𝜑 → 𝐻:𝑈–onto→∪ 𝑥 ∈ 𝐴 𝐵) | ||
28-Oct-2023 | ctiunctlemf 12209 | Lemma for ctiunct 12211. (Contributed by Jim Kingdon, 28-Oct-2023.) |
⊢ (𝜑 → 𝑆 ⊆ ω) & ⊢ (𝜑 → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑆) & ⊢ (𝜑 → 𝐹:𝑆–onto→𝐴) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑇 ⊆ ω) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑇) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝐺:𝑇–onto→𝐵) & ⊢ (𝜑 → 𝐽:ω–1-1-onto→(ω × ω)) & ⊢ 𝑈 = {𝑧 ∈ ω ∣ ((1st ‘(𝐽‘𝑧)) ∈ 𝑆 ∧ (2nd ‘(𝐽‘𝑧)) ∈ ⦋(𝐹‘(1st ‘(𝐽‘𝑧))) / 𝑥⦌𝑇)} & ⊢ 𝐻 = (𝑛 ∈ 𝑈 ↦ (⦋(𝐹‘(1st ‘(𝐽‘𝑛))) / 𝑥⦌𝐺‘(2nd ‘(𝐽‘𝑛)))) ⇒ ⊢ (𝜑 → 𝐻:𝑈⟶∪ 𝑥 ∈ 𝐴 𝐵) | ||
28-Oct-2023 | ctiunctlemudc 12208 | Lemma for ctiunct 12211. (Contributed by Jim Kingdon, 28-Oct-2023.) |
⊢ (𝜑 → 𝑆 ⊆ ω) & ⊢ (𝜑 → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑆) & ⊢ (𝜑 → 𝐹:𝑆–onto→𝐴) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑇 ⊆ ω) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑇) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝐺:𝑇–onto→𝐵) & ⊢ (𝜑 → 𝐽:ω–1-1-onto→(ω × ω)) & ⊢ 𝑈 = {𝑧 ∈ ω ∣ ((1st ‘(𝐽‘𝑧)) ∈ 𝑆 ∧ (2nd ‘(𝐽‘𝑧)) ∈ ⦋(𝐹‘(1st ‘(𝐽‘𝑧))) / 𝑥⦌𝑇)} ⇒ ⊢ (𝜑 → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑈) | ||
28-Oct-2023 | ctiunctlemuom 12207 | Lemma for ctiunct 12211. (Contributed by Jim Kingdon, 28-Oct-2023.) |
⊢ (𝜑 → 𝑆 ⊆ ω) & ⊢ (𝜑 → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑆) & ⊢ (𝜑 → 𝐹:𝑆–onto→𝐴) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑇 ⊆ ω) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑇) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝐺:𝑇–onto→𝐵) & ⊢ (𝜑 → 𝐽:ω–1-1-onto→(ω × ω)) & ⊢ 𝑈 = {𝑧 ∈ ω ∣ ((1st ‘(𝐽‘𝑧)) ∈ 𝑆 ∧ (2nd ‘(𝐽‘𝑧)) ∈ ⦋(𝐹‘(1st ‘(𝐽‘𝑧))) / 𝑥⦌𝑇)} ⇒ ⊢ (𝜑 → 𝑈 ⊆ ω) | ||
28-Oct-2023 | ctiunctlemu2nd 12206 | Lemma for ctiunct 12211. (Contributed by Jim Kingdon, 28-Oct-2023.) |
⊢ (𝜑 → 𝑆 ⊆ ω) & ⊢ (𝜑 → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑆) & ⊢ (𝜑 → 𝐹:𝑆–onto→𝐴) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑇 ⊆ ω) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑇) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝐺:𝑇–onto→𝐵) & ⊢ (𝜑 → 𝐽:ω–1-1-onto→(ω × ω)) & ⊢ 𝑈 = {𝑧 ∈ ω ∣ ((1st ‘(𝐽‘𝑧)) ∈ 𝑆 ∧ (2nd ‘(𝐽‘𝑧)) ∈ ⦋(𝐹‘(1st ‘(𝐽‘𝑧))) / 𝑥⦌𝑇)} & ⊢ (𝜑 → 𝑁 ∈ 𝑈) ⇒ ⊢ (𝜑 → (2nd ‘(𝐽‘𝑁)) ∈ ⦋(𝐹‘(1st ‘(𝐽‘𝑁))) / 𝑥⦌𝑇) | ||
28-Oct-2023 | ctiunctlemu1st 12205 | Lemma for ctiunct 12211. (Contributed by Jim Kingdon, 28-Oct-2023.) |
⊢ (𝜑 → 𝑆 ⊆ ω) & ⊢ (𝜑 → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑆) & ⊢ (𝜑 → 𝐹:𝑆–onto→𝐴) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑇 ⊆ ω) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑇) & ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝐺:𝑇–onto→𝐵) & ⊢ (𝜑 → 𝐽:ω–1-1-onto→(ω × ω)) & ⊢ 𝑈 = {𝑧 ∈ ω ∣ ((1st ‘(𝐽‘𝑧)) ∈ 𝑆 ∧ (2nd ‘(𝐽‘𝑧)) ∈ ⦋(𝐹‘(1st ‘(𝐽‘𝑧))) / 𝑥⦌𝑇)} & ⊢ (𝜑 → 𝑁 ∈ 𝑈) ⇒ ⊢ (𝜑 → (1st ‘(𝐽‘𝑁)) ∈ 𝑆) | ||
28-Oct-2023 | pm2.521gdc 854 | A general instance of Theorem *2.521 of [WhiteheadRussell] p. 107, under a decidability condition. (Contributed by BJ, 28-Oct-2023.) |
⊢ (DECID 𝜑 → (¬ (𝜑 → 𝜓) → (𝜒 → 𝜑))) | ||
28-Oct-2023 | stdcndc 831 | A formula is decidable if and only if its negation is decidable and it is stable (that is, it is testable and stable). (Contributed by David A. Wheeler, 13-Aug-2018.) (Proof shortened by BJ, 28-Oct-2023.) |
⊢ ((STAB 𝜑 ∧ DECID ¬ 𝜑) ↔ DECID 𝜑) | ||
28-Oct-2023 | conax1k 644 | Weakening of conax1 643. General instance of pm2.51 645 and of pm2.52 646. (Contributed by BJ, 28-Oct-2023.) |
⊢ (¬ (𝜑 → 𝜓) → (𝜒 → ¬ 𝜓)) | ||
28-Oct-2023 | conax1 643 | Contrapositive of ax-1 6. (Contributed by BJ, 28-Oct-2023.) |
⊢ (¬ (𝜑 → 𝜓) → ¬ 𝜓) | ||
25-Oct-2023 | divcnap 12997 | Complex number division is a continuous function, when the second argument is apart from zero. (Contributed by Mario Carneiro, 12-Aug-2014.) (Revised by Jim Kingdon, 25-Oct-2023.) |
⊢ 𝐽 = (MetOpen‘(abs ∘ − )) & ⊢ 𝐾 = (𝐽 ↾t {𝑥 ∈ ℂ ∣ 𝑥 # 0}) ⇒ ⊢ (𝑦 ∈ ℂ, 𝑧 ∈ {𝑥 ∈ ℂ ∣ 𝑥 # 0} ↦ (𝑦 / 𝑧)) ∈ ((𝐽 ×t 𝐾) Cn 𝐽) | ||
23-Oct-2023 | cnm 7753 | A complex number is an inhabited set. Note: do not use this after the real number axioms are developed, since it is a construction-dependent property. (Contributed by Jim Kingdon, 23-Oct-2023.) (New usage is discouraged.) |
⊢ (𝐴 ∈ ℂ → ∃𝑥 𝑥 ∈ 𝐴) | ||
23-Oct-2023 | oprssdmm 6120 | Domain of closure of an operation. (Contributed by Jim Kingdon, 23-Oct-2023.) |
⊢ ((𝜑 ∧ 𝑢 ∈ 𝑆) → ∃𝑣 𝑣 ∈ 𝑢) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑆 ∧ 𝑦 ∈ 𝑆)) → (𝑥𝐹𝑦) ∈ 𝑆) & ⊢ (𝜑 → Rel 𝐹) ⇒ ⊢ (𝜑 → (𝑆 × 𝑆) ⊆ dom 𝐹) | ||
22-Oct-2023 | addcncntoplem 12993 | Lemma for addcncntop 12994, subcncntop 12995, and mulcncntop 12996. (Contributed by Mario Carneiro, 5-May-2014.) (Revised by Jim Kingdon, 22-Oct-2023.) |
⊢ 𝐽 = (MetOpen‘(abs ∘ − )) & ⊢ + :(ℂ × ℂ)⟶ℂ & ⊢ ((𝑎 ∈ ℝ+ ∧ 𝑏 ∈ ℂ ∧ 𝑐 ∈ ℂ) → ∃𝑦 ∈ ℝ+ ∃𝑧 ∈ ℝ+ ∀𝑢 ∈ ℂ ∀𝑣 ∈ ℂ (((abs‘(𝑢 − 𝑏)) < 𝑦 ∧ (abs‘(𝑣 − 𝑐)) < 𝑧) → (abs‘((𝑢 + 𝑣) − (𝑏 + 𝑐))) < 𝑎)) ⇒ ⊢ + ∈ ((𝐽 ×t 𝐽) Cn 𝐽) | ||
22-Oct-2023 | txmetcnp 12960 | Continuity of a binary operation on metric spaces. (Contributed by Mario Carneiro, 2-Sep-2015.) (Revised by Jim Kingdon, 22-Oct-2023.) |
⊢ 𝐽 = (MetOpen‘𝐶) & ⊢ 𝐾 = (MetOpen‘𝐷) & ⊢ 𝐿 = (MetOpen‘𝐸) ⇒ ⊢ (((𝐶 ∈ (∞Met‘𝑋) ∧ 𝐷 ∈ (∞Met‘𝑌) ∧ 𝐸 ∈ (∞Met‘𝑍)) ∧ (𝐴 ∈ 𝑋 ∧ 𝐵 ∈ 𝑌)) → (𝐹 ∈ (((𝐽 ×t 𝐾) CnP 𝐿)‘〈𝐴, 𝐵〉) ↔ (𝐹:(𝑋 × 𝑌)⟶𝑍 ∧ ∀𝑧 ∈ ℝ+ ∃𝑤 ∈ ℝ+ ∀𝑢 ∈ 𝑋 ∀𝑣 ∈ 𝑌 (((𝐴𝐶𝑢) < 𝑤 ∧ (𝐵𝐷𝑣) < 𝑤) → ((𝐴𝐹𝐵)𝐸(𝑢𝐹𝑣)) < 𝑧)))) | ||
22-Oct-2023 | xmetxpbl 12950 | The maximum metric (Chebyshev distance) on the product of two sets, expressed in terms of balls centered on a point 𝐶 with radius 𝑅. (Contributed by Jim Kingdon, 22-Oct-2023.) |
⊢ 𝑃 = (𝑢 ∈ (𝑋 × 𝑌), 𝑣 ∈ (𝑋 × 𝑌) ↦ sup({((1st ‘𝑢)𝑀(1st ‘𝑣)), ((2nd ‘𝑢)𝑁(2nd ‘𝑣))}, ℝ*, < )) & ⊢ (𝜑 → 𝑀 ∈ (∞Met‘𝑋)) & ⊢ (𝜑 → 𝑁 ∈ (∞Met‘𝑌)) & ⊢ (𝜑 → 𝑅 ∈ ℝ*) & ⊢ (𝜑 → 𝐶 ∈ (𝑋 × 𝑌)) ⇒ ⊢ (𝜑 → (𝐶(ball‘𝑃)𝑅) = (((1st ‘𝐶)(ball‘𝑀)𝑅) × ((2nd ‘𝐶)(ball‘𝑁)𝑅))) | ||
15-Oct-2023 | xmettxlem 12951 | Lemma for xmettx 12952. (Contributed by Jim Kingdon, 15-Oct-2023.) |
⊢ 𝑃 = (𝑢 ∈ (𝑋 × 𝑌), 𝑣 ∈ (𝑋 × 𝑌) ↦ sup({((1st ‘𝑢)𝑀(1st ‘𝑣)), ((2nd ‘𝑢)𝑁(2nd ‘𝑣))}, ℝ*, < )) & ⊢ (𝜑 → 𝑀 ∈ (∞Met‘𝑋)) & ⊢ (𝜑 → 𝑁 ∈ (∞Met‘𝑌)) & ⊢ 𝐽 = (MetOpen‘𝑀) & ⊢ 𝐾 = (MetOpen‘𝑁) & ⊢ 𝐿 = (MetOpen‘𝑃) ⇒ ⊢ (𝜑 → 𝐿 ⊆ (𝐽 ×t 𝐾)) | ||
11-Oct-2023 | xmettx 12952 | The maximum metric (Chebyshev distance) on the product of two sets, expressed as a binary topological product. (Contributed by Jim Kingdon, 11-Oct-2023.) |
⊢ 𝑃 = (𝑢 ∈ (𝑋 × 𝑌), 𝑣 ∈ (𝑋 × 𝑌) ↦ sup({((1st ‘𝑢)𝑀(1st ‘𝑣)), ((2nd ‘𝑢)𝑁(2nd ‘𝑣))}, ℝ*, < )) & ⊢ (𝜑 → 𝑀 ∈ (∞Met‘𝑋)) & ⊢ (𝜑 → 𝑁 ∈ (∞Met‘𝑌)) & ⊢ 𝐽 = (MetOpen‘𝑀) & ⊢ 𝐾 = (MetOpen‘𝑁) & ⊢ 𝐿 = (MetOpen‘𝑃) ⇒ ⊢ (𝜑 → 𝐿 = (𝐽 ×t 𝐾)) | ||
11-Oct-2023 | xmetxp 12949 | The maximum metric (Chebyshev distance) on the product of two sets. (Contributed by Jim Kingdon, 11-Oct-2023.) |
⊢ 𝑃 = (𝑢 ∈ (𝑋 × 𝑌), 𝑣 ∈ (𝑋 × 𝑌) ↦ sup({((1st ‘𝑢)𝑀(1st ‘𝑣)), ((2nd ‘𝑢)𝑁(2nd ‘𝑣))}, ℝ*, < )) & ⊢ (𝜑 → 𝑀 ∈ (∞Met‘𝑋)) & ⊢ (𝜑 → 𝑁 ∈ (∞Met‘𝑌)) ⇒ ⊢ (𝜑 → 𝑃 ∈ (∞Met‘(𝑋 × 𝑌))) | ||
12-Sep-2023 | pwntru 4161 | A slight strengthening of pwtrufal 13611. (Contributed by Mario Carneiro and Jim Kingdon, 12-Sep-2023.) |
⊢ ((𝐴 ⊆ {∅} ∧ 𝐴 ≠ {∅}) → 𝐴 = ∅) | ||
11-Sep-2023 | pwtrufal 13611 | A subset of the singleton {∅} cannot be anything other than ∅ or {∅}. Removing the double negation would change the meaning, as seen at exmid01 4160. If we view a subset of a singleton as a truth value (as seen in theorems like exmidexmid 4158), then this theorem states there are no truth values other than true and false, as described in section 1.1 of [Bauer], p. 481. (Contributed by Mario Carneiro and Jim Kingdon, 11-Sep-2023.) |
⊢ (𝐴 ⊆ {∅} → ¬ ¬ (𝐴 = ∅ ∨ 𝐴 = {∅})) | ||
9-Sep-2023 | mathbox 13350 |
(This theorem is a dummy placeholder for these guidelines. The label
of this theorem, "mathbox", is hard-coded into the Metamath
program to
identify the start of the mathbox section for web page generation.)
A "mathbox" is a user-contributed section that is maintained by its contributor independently from the main part of iset.mm. For contributors: By making a contribution, you agree to release it into the public domain, according to the statement at the beginning of iset.mm. Guidelines: Mathboxes in iset.mm follow the same practices as in set.mm, so refer to the mathbox guidelines there for more details. (Contributed by NM, 20-Feb-2007.) (Revised by the Metamath team, 9-Sep-2023.) (New usage is discouraged.) |
⊢ 𝜑 ⇒ ⊢ 𝜑 | ||
6-Sep-2023 | djuexb 6989 | The disjoint union of two classes is a set iff both classes are sets. (Contributed by Jim Kingdon, 6-Sep-2023.) |
⊢ ((𝐴 ∈ V ∧ 𝐵 ∈ V) ↔ (𝐴 ⊔ 𝐵) ∈ V) | ||
3-Sep-2023 | pwf1oexmid 13613 | An exercise related to 𝑁 copies of a singleton and the power set of a singleton (where the latter can also be thought of as representing truth values). Posed as an exercise by Martin Escardo online. (Contributed by Jim Kingdon, 3-Sep-2023.) |
⊢ 𝑇 = ∪ 𝑥 ∈ 𝑁 ({𝑥} × 1o) ⇒ ⊢ ((𝑁 ∈ ω ∧ 𝐺:𝑇–1-1→𝒫 1o) → (ran 𝐺 = 𝒫 1o ↔ (𝑁 = 2o ∧ EXMID))) | ||
3-Sep-2023 | pwle2 13612 | An exercise related to 𝑁 copies of a singleton and the power set of a singleton (where the latter can also be thought of as representing truth values). Posed as an exercise by Martin Escardo online. (Contributed by Jim Kingdon, 3-Sep-2023.) |
⊢ 𝑇 = ∪ 𝑥 ∈ 𝑁 ({𝑥} × 1o) ⇒ ⊢ ((𝑁 ∈ ω ∧ 𝐺:𝑇–1-1→𝒫 1o) → 𝑁 ⊆ 2o) | ||
30-Aug-2023 | isomninn 13644 | Omniscience stated in terms of natural numbers. Similar to isomnimap 7081 but it will sometimes be more convenient to use 0 and 1 rather than ∅ and 1o. (Contributed by Jim Kingdon, 30-Aug-2023.) |
⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ Omni ↔ ∀𝑓 ∈ ({0, 1} ↑𝑚 𝐴)(∃𝑥 ∈ 𝐴 (𝑓‘𝑥) = 0 ∨ ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) = 1))) | ||
30-Aug-2023 | isomninnlem 13643 | Lemma for isomninn 13644. The result, with a hypothesis to provide a convenient notation. (Contributed by Jim Kingdon, 30-Aug-2023.) |
⊢ 𝐺 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) ⇒ ⊢ (𝐴 ∈ 𝑉 → (𝐴 ∈ Omni ↔ ∀𝑓 ∈ ({0, 1} ↑𝑚 𝐴)(∃𝑥 ∈ 𝐴 (𝑓‘𝑥) = 0 ∨ ∀𝑥 ∈ 𝐴 (𝑓‘𝑥) = 1))) | ||
28-Aug-2023 | trilpolemisumle 13651 | Lemma for trilpo 13656. An upper bound for the sum of the digits beyond a certain point. (Contributed by Jim Kingdon, 28-Aug-2023.) |
⊢ (𝜑 → 𝐹:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐹‘𝑖)) & ⊢ 𝑍 = (ℤ≥‘𝑀) & ⊢ (𝜑 → 𝑀 ∈ ℕ) ⇒ ⊢ (𝜑 → Σ𝑖 ∈ 𝑍 ((1 / (2↑𝑖)) · (𝐹‘𝑖)) ≤ Σ𝑖 ∈ 𝑍 (1 / (2↑𝑖))) | ||
25-Aug-2023 | cvgcmp2n 13646 | A comparison test for convergence of a real infinite series. (Contributed by Jim Kingdon, 25-Aug-2023.) |
⊢ ((𝜑 ∧ 𝑘 ∈ ℕ) → (𝐺‘𝑘) ∈ ℝ) & ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ) → 0 ≤ (𝐺‘𝑘)) & ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ) → (𝐺‘𝑘) ≤ (1 / (2↑𝑘))) ⇒ ⊢ (𝜑 → seq1( + , 𝐺) ∈ dom ⇝ ) | ||
25-Aug-2023 | cvgcmp2nlemabs 13645 | Lemma for cvgcmp2n 13646. The partial sums get closer to each other as we go further out. The proof proceeds by rewriting (seq1( + , 𝐺)‘𝑁) as the sum of (seq1( + , 𝐺)‘𝑀) and a term which gets smaller as 𝑀 gets large. (Contributed by Jim Kingdon, 25-Aug-2023.) |
⊢ ((𝜑 ∧ 𝑘 ∈ ℕ) → (𝐺‘𝑘) ∈ ℝ) & ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ) → 0 ≤ (𝐺‘𝑘)) & ⊢ ((𝜑 ∧ 𝑘 ∈ ℕ) → (𝐺‘𝑘) ≤ (1 / (2↑𝑘))) & ⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) ⇒ ⊢ (𝜑 → (abs‘((seq1( + , 𝐺)‘𝑁) − (seq1( + , 𝐺)‘𝑀))) < (2 / 𝑀)) | ||
24-Aug-2023 | trilpolemclim 13649 | Lemma for trilpo 13656. Convergence of the series. (Contributed by Jim Kingdon, 24-Aug-2023.) |
⊢ (𝜑 → 𝐹:ℕ⟶{0, 1}) & ⊢ 𝐺 = (𝑛 ∈ ℕ ↦ ((1 / (2↑𝑛)) · (𝐹‘𝑛))) ⇒ ⊢ (𝜑 → seq1( + , 𝐺) ∈ dom ⇝ ) | ||
23-Aug-2023 | trilpo 13656 |
Real number trichotomy implies the Limited Principle of Omniscience
(LPO). We expect that we'd need some form of countable choice to prove
the converse.
Here's the outline of the proof. Given an infinite sequence F of zeroes and ones, we need to show the sequence contains a zero or it is all ones. Construct a real number A whose representation in base two consists of a zero, a decimal point, and then the numbers of the sequence. Compare it with one using trichotomy. The three cases from trichotomy are trilpolemlt1 13654 (which means the sequence contains a zero), trilpolemeq1 13653 (which means the sequence is all ones), and trilpolemgt1 13652 (which is not possible). Equivalent ways to state real number trichotomy (sometimes called "analytic LPO") include decidability of real number apartness (see triap 13642) or that the real numbers are a discrete field (see trirec0 13657). LPO is known to not be provable in IZF (and most constructive foundations), so this theorem establishes that we will be unable to prove an analogue to qtri3or 10146 for real numbers. (Contributed by Jim Kingdon, 23-Aug-2023.) |
⊢ (∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 ∨ 𝑥 = 𝑦 ∨ 𝑦 < 𝑥) → ω ∈ Omni) | ||
23-Aug-2023 | trilpolemres 13655 | Lemma for trilpo 13656. The result. (Contributed by Jim Kingdon, 23-Aug-2023.) |
⊢ (𝜑 → 𝐹:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐹‘𝑖)) & ⊢ (𝜑 → (𝐴 < 1 ∨ 𝐴 = 1 ∨ 1 < 𝐴)) ⇒ ⊢ (𝜑 → (∃𝑥 ∈ ℕ (𝐹‘𝑥) = 0 ∨ ∀𝑥 ∈ ℕ (𝐹‘𝑥) = 1)) | ||
23-Aug-2023 | trilpolemlt1 13654 | Lemma for trilpo 13656. The 𝐴 < 1 case. We can use the distance between 𝐴 and one (that is, 1 − 𝐴) to find a position in the sequence 𝑛 where terms after that point will not add up to as much as 1 − 𝐴. By finomni 7084 we know the terms up to 𝑛 either contain a zero or are all one. But if they are all one that contradicts the way we constructed 𝑛, so we know that the sequence contains a zero. (Contributed by Jim Kingdon, 23-Aug-2023.) |
⊢ (𝜑 → 𝐹:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐹‘𝑖)) & ⊢ (𝜑 → 𝐴 < 1) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℕ (𝐹‘𝑥) = 0) | ||
23-Aug-2023 | trilpolemeq1 13653 | Lemma for trilpo 13656. The 𝐴 = 1 case. This is proved by noting that if any (𝐹‘𝑥) is zero, then the infinite sum 𝐴 is less than one based on the term which is zero. We are using the fact that the 𝐹 sequence is decidable (in the sense that each element is either zero or one). (Contributed by Jim Kingdon, 23-Aug-2023.) |
⊢ (𝜑 → 𝐹:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐹‘𝑖)) & ⊢ (𝜑 → 𝐴 = 1) ⇒ ⊢ (𝜑 → ∀𝑥 ∈ ℕ (𝐹‘𝑥) = 1) | ||
23-Aug-2023 | trilpolemgt1 13652 | Lemma for trilpo 13656. The 1 < 𝐴 case. (Contributed by Jim Kingdon, 23-Aug-2023.) |
⊢ (𝜑 → 𝐹:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐹‘𝑖)) ⇒ ⊢ (𝜑 → ¬ 1 < 𝐴) | ||
23-Aug-2023 | trilpolemcl 13650 | Lemma for trilpo 13656. The sum exists. (Contributed by Jim Kingdon, 23-Aug-2023.) |
⊢ (𝜑 → 𝐹:ℕ⟶{0, 1}) & ⊢ 𝐴 = Σ𝑖 ∈ ℕ ((1 / (2↑𝑖)) · (𝐹‘𝑖)) ⇒ ⊢ (𝜑 → 𝐴 ∈ ℝ) | ||
23-Aug-2023 | triap 13642 | Two ways of stating real number trichotomy. (Contributed by Jim Kingdon, 23-Aug-2023.) |
⊢ ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → ((𝐴 < 𝐵 ∨ 𝐴 = 𝐵 ∨ 𝐵 < 𝐴) ↔ DECID 𝐴 # 𝐵)) | ||
19-Aug-2023 | djuenun 7148 | Disjoint union is equinumerous to union for disjoint sets. (Contributed by Mario Carneiro, 29-Apr-2015.) (Revised by Jim Kingdon, 19-Aug-2023.) |
⊢ ((𝐴 ≈ 𝐵 ∧ 𝐶 ≈ 𝐷 ∧ (𝐵 ∩ 𝐷) = ∅) → (𝐴 ⊔ 𝐶) ≈ (𝐵 ∪ 𝐷)) | ||
16-Aug-2023 | ctssdclemr 7057 | Lemma for ctssdc 7058. Showing that our usual definition of countable implies the alternate one. (Contributed by Jim Kingdon, 16-Aug-2023.) |
⊢ (∃𝑓 𝑓:ω–onto→(𝐴 ⊔ 1o) → ∃𝑠(𝑠 ⊆ ω ∧ ∃𝑓 𝑓:𝑠–onto→𝐴 ∧ ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑠)) | ||
16-Aug-2023 | ctssdclemn0 7055 | Lemma for ctssdc 7058. The ¬ ∅ ∈ 𝑆 case. (Contributed by Jim Kingdon, 16-Aug-2023.) |
⊢ (𝜑 → 𝑆 ⊆ ω) & ⊢ (𝜑 → ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑆) & ⊢ (𝜑 → 𝐹:𝑆–onto→𝐴) & ⊢ (𝜑 → ¬ ∅ ∈ 𝑆) ⇒ ⊢ (𝜑 → ∃𝑔 𝑔:ω–onto→(𝐴 ⊔ 1o)) | ||
15-Aug-2023 | ctssexmid 7094 | The decidability condition in ctssdc 7058 is needed. More specifically, ctssdc 7058 minus that condition, plus the Limited Principle of Omniscience (LPO), implies excluded middle. (Contributed by Jim Kingdon, 15-Aug-2023.) |
⊢ ((𝑦 ⊆ ω ∧ ∃𝑓 𝑓:𝑦–onto→𝑥) → ∃𝑓 𝑓:ω–onto→(𝑥 ⊔ 1o)) & ⊢ ω ∈ Omni ⇒ ⊢ (𝜑 ∨ ¬ 𝜑) | ||
15-Aug-2023 | ctssdc 7058 | A set is countable iff there is a surjection from a decidable subset of the natural numbers onto it. The decidability condition is needed as shown at ctssexmid 7094. (Contributed by Jim Kingdon, 15-Aug-2023.) |
⊢ (∃𝑠(𝑠 ⊆ ω ∧ ∃𝑓 𝑓:𝑠–onto→𝐴 ∧ ∀𝑛 ∈ ω DECID 𝑛 ∈ 𝑠) ↔ ∃𝑓 𝑓:ω–onto→(𝐴 ⊔ 1o)) | ||
13-Aug-2023 | ltntri 8004 | Negative trichotomy property for real numbers. It is well known that we cannot prove real number trichotomy, 𝐴 < 𝐵 ∨ 𝐴 = 𝐵 ∨ 𝐵 < 𝐴. Does that mean there is a pair of real numbers where none of those hold (that is, where we can refute each of those three relationships)? Actually, no, as shown here. This is another example of distinguishing between being unable to prove something, or being able to refute it. (Contributed by Jim Kingdon, 13-Aug-2023.) |
⊢ ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ) → ¬ (¬ 𝐴 < 𝐵 ∧ ¬ 𝐴 = 𝐵 ∧ ¬ 𝐵 < 𝐴)) | ||
11-Aug-2023 | qnnen 12202 | The rational numbers are countably infinite. Corollary 8.1.23 of [AczelRathjen], p. 75. This is Metamath 100 proof #3. (Contributed by Jim Kingdon, 11-Aug-2023.) |
⊢ ℚ ≈ ℕ | ||
10-Aug-2023 | ctinfomlemom 12198 | Lemma for ctinfom 12199. Converting between ω and ℕ0. (Contributed by Jim Kingdon, 10-Aug-2023.) |
⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐺 = (𝐹 ∘ ◡𝑁) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ¬ (𝐹‘𝑘) ∈ (𝐹 “ 𝑛)) ⇒ ⊢ (𝜑 → (𝐺:ℕ0–onto→𝐴 ∧ ∀𝑚 ∈ ℕ0 ∃𝑗 ∈ ℕ0 ∀𝑖 ∈ (0...𝑚)(𝐺‘𝑗) ≠ (𝐺‘𝑖))) | ||
9-Aug-2023 | difinfsnlem 7044 | Lemma for difinfsn 7045. The case where we need to swap 𝐵 and (inr‘∅) in building the mapping 𝐺. (Contributed by Jim Kingdon, 9-Aug-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → 𝐹:(ω ⊔ 1o)–1-1→𝐴) & ⊢ (𝜑 → (𝐹‘(inr‘∅)) ≠ 𝐵) & ⊢ 𝐺 = (𝑛 ∈ ω ↦ if((𝐹‘(inl‘𝑛)) = 𝐵, (𝐹‘(inr‘∅)), (𝐹‘(inl‘𝑛)))) ⇒ ⊢ (𝜑 → 𝐺:ω–1-1→(𝐴 ∖ {𝐵})) | ||
8-Aug-2023 | difinfinf 7046 | An infinite set minus a finite subset is infinite. We require that the set has decidable equality. (Contributed by Jim Kingdon, 8-Aug-2023.) |
⊢ (((∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦 ∧ ω ≼ 𝐴) ∧ (𝐵 ⊆ 𝐴 ∧ 𝐵 ∈ Fin)) → ω ≼ (𝐴 ∖ 𝐵)) | ||
8-Aug-2023 | difinfsn 7045 | An infinite set minus one element is infinite. We require that the set has decidable equality. (Contributed by Jim Kingdon, 8-Aug-2023.) |
⊢ ((∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦 ∧ ω ≼ 𝐴 ∧ 𝐵 ∈ 𝐴) → ω ≼ (𝐴 ∖ {𝐵})) | ||
7-Aug-2023 | ctinf 12201 | A set is countably infinite if and only if it has decidable equality, is countable, and is infinite. (Contributed by Jim Kingdon, 7-Aug-2023.) |
⊢ (𝐴 ≈ ℕ ↔ (∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦 ∧ ∃𝑓 𝑓:ω–onto→𝐴 ∧ ω ≼ 𝐴)) | ||
7-Aug-2023 | inffinp1 12200 | An infinite set contains an element not contained in a given finite subset. (Contributed by Jim Kingdon, 7-Aug-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → ω ≼ 𝐴) & ⊢ (𝜑 → 𝐵 ⊆ 𝐴) & ⊢ (𝜑 → 𝐵 ∈ Fin) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ 𝐴 ¬ 𝑥 ∈ 𝐵) | ||
7-Aug-2023 | ctinfom 12199 | A condition for a set being countably infinite. Restates ennnfone 12196 in terms of ω and function image. Like ennnfone 12196 the condition can be summarized as 𝐴 being countable, infinite, and having decidable equality. (Contributed by Jim Kingdon, 7-Aug-2023.) |
⊢ (𝐴 ≈ ℕ ↔ (∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦 ∧ ∃𝑓(𝑓:ω–onto→𝐴 ∧ ∀𝑛 ∈ ω ∃𝑘 ∈ ω ¬ (𝑓‘𝑘) ∈ (𝑓 “ 𝑛)))) | ||
6-Aug-2023 | rerestcntop 12992 | The subspace topology induced by a subset of the reals. (Contributed by Mario Carneiro, 13-Aug-2014.) (Revised by Jim Kingdon, 6-Aug-2023.) |
⊢ 𝐽 = (MetOpen‘(abs ∘ − )) & ⊢ 𝑅 = (topGen‘ran (,)) ⇒ ⊢ (𝐴 ⊆ ℝ → (𝐽 ↾t 𝐴) = (𝑅 ↾t 𝐴)) | ||
6-Aug-2023 | tgioo2cntop 12991 | The standard topology on the reals is a subspace of the complex metric topology. (Contributed by Mario Carneiro, 13-Aug-2014.) (Revised by Jim Kingdon, 6-Aug-2023.) |
⊢ 𝐽 = (MetOpen‘(abs ∘ − )) ⇒ ⊢ (topGen‘ran (,)) = (𝐽 ↾t ℝ) | ||
4-Aug-2023 | nninffeq 13634 | Equality of two functions on ℕ∞ which agree at every integer and at the point at infinity. From an online post by Martin Escardo. Remark: the last two hypotheses can be grouped into one, ⊢ (𝜑 → ∀𝑛 ∈ suc ω...). (Contributed by Jim Kingdon, 4-Aug-2023.) |
⊢ (𝜑 → 𝐹:ℕ∞⟶ℕ0) & ⊢ (𝜑 → 𝐺:ℕ∞⟶ℕ0) & ⊢ (𝜑 → (𝐹‘(𝑥 ∈ ω ↦ 1o)) = (𝐺‘(𝑥 ∈ ω ↦ 1o))) & ⊢ (𝜑 → ∀𝑛 ∈ ω (𝐹‘(𝑖 ∈ ω ↦ if(𝑖 ∈ 𝑛, 1o, ∅))) = (𝐺‘(𝑖 ∈ ω ↦ if(𝑖 ∈ 𝑛, 1o, ∅)))) ⇒ ⊢ (𝜑 → 𝐹 = 𝐺) | ||
3-Aug-2023 | txvalex 12696 | Existence of the binary topological product. If 𝑅 and 𝑆 are known to be topologies, see txtop 12702. (Contributed by Jim Kingdon, 3-Aug-2023.) |
⊢ ((𝑅 ∈ 𝑉 ∧ 𝑆 ∈ 𝑊) → (𝑅 ×t 𝑆) ∈ V) | ||
3-Aug-2023 | dvdsgcdidd 11882 | The greatest common divisor of a positive integer and another integer it divides is itself. (Contributed by Rohan Ridenour, 3-Aug-2023.) |
⊢ (𝜑 → 𝑀 ∈ ℕ) & ⊢ (𝜑 → 𝑁 ∈ ℤ) & ⊢ (𝜑 → 𝑀 ∥ 𝑁) ⇒ ⊢ (𝜑 → (𝑀 gcd 𝑁) = 𝑀) | ||
3-Aug-2023 | gcdmultipled 11881 | The greatest common divisor of a nonnegative integer 𝑀 and a multiple of it is 𝑀 itself. (Contributed by Rohan Ridenour, 3-Aug-2023.) |
⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑁 ∈ ℤ) ⇒ ⊢ (𝜑 → (𝑀 gcd (𝑁 · 𝑀)) = 𝑀) | ||
3-Aug-2023 | phpeqd 6878 | Corollary of the Pigeonhole Principle using equality. Strengthening of phpm 6811 expressed without negation. (Contributed by Rohan Ridenour, 3-Aug-2023.) |
⊢ (𝜑 → 𝐴 ∈ Fin) & ⊢ (𝜑 → 𝐵 ⊆ 𝐴) & ⊢ (𝜑 → 𝐴 ≈ 𝐵) ⇒ ⊢ (𝜑 → 𝐴 = 𝐵) | ||
3-Aug-2023 | enpr2d 6763 | A pair with distinct elements is equinumerous to ordinal two. (Contributed by Rohan Ridenour, 3-Aug-2023.) |
⊢ (𝜑 → 𝐴 ∈ 𝐶) & ⊢ (𝜑 → 𝐵 ∈ 𝐷) & ⊢ (𝜑 → ¬ 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → {𝐴, 𝐵} ≈ 2o) | ||
3-Aug-2023 | elrnmpt2d 4842 | Elementhood in the range of a function in maps-to notation, deduction form. (Contributed by Rohan Ridenour, 3-Aug-2023.) |
⊢ 𝐹 = (𝑥 ∈ 𝐴 ↦ 𝐵) & ⊢ (𝜑 → 𝐶 ∈ ran 𝐹) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ 𝐴 𝐶 = 𝐵) | ||
3-Aug-2023 | elrnmptdv 4841 | Elementhood in the range of a function in maps-to notation, deduction form. (Contributed by Rohan Ridenour, 3-Aug-2023.) |
⊢ 𝐹 = (𝑥 ∈ 𝐴 ↦ 𝐵) & ⊢ (𝜑 → 𝐶 ∈ 𝐴) & ⊢ (𝜑 → 𝐷 ∈ 𝑉) & ⊢ ((𝜑 ∧ 𝑥 = 𝐶) → 𝐷 = 𝐵) ⇒ ⊢ (𝜑 → 𝐷 ∈ ran 𝐹) | ||
3-Aug-2023 | rspcime 2823 | Prove a restricted existential. (Contributed by Rohan Ridenour, 3-Aug-2023.) |
⊢ ((𝜑 ∧ 𝑥 = 𝐴) → 𝜓) & ⊢ (𝜑 → 𝐴 ∈ 𝐵) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ 𝐵 𝜓) | ||
3-Aug-2023 | neqcomd 2162 | Commute an inequality. (Contributed by Rohan Ridenour, 3-Aug-2023.) |
⊢ (𝜑 → ¬ 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → ¬ 𝐵 = 𝐴) | ||
2-Aug-2023 | dvid 13104 | Derivative of the identity function. (Contributed by Mario Carneiro, 8-Aug-2014.) (Revised by Jim Kingdon, 2-Aug-2023.) |
⊢ (ℂ D ( I ↾ ℂ)) = (ℂ × {1}) | ||
2-Aug-2023 | dvconst 13103 | Derivative of a constant function. (Contributed by Mario Carneiro, 8-Aug-2014.) (Revised by Jim Kingdon, 2-Aug-2023.) |
⊢ (𝐴 ∈ ℂ → (ℂ D (ℂ × {𝐴})) = (ℂ × {0})) | ||
2-Aug-2023 | dvidlemap 13102 | Lemma for dvid 13104 and dvconst 13103. (Contributed by Mario Carneiro, 8-Aug-2014.) (Revised by Jim Kingdon, 2-Aug-2023.) |
⊢ (𝜑 → 𝐹:ℂ⟶ℂ) & ⊢ ((𝜑 ∧ (𝑥 ∈ ℂ ∧ 𝑧 ∈ ℂ ∧ 𝑧 # 𝑥)) → (((𝐹‘𝑧) − (𝐹‘𝑥)) / (𝑧 − 𝑥)) = 𝐵) & ⊢ 𝐵 ∈ ℂ ⇒ ⊢ (𝜑 → (ℂ D 𝐹) = (ℂ × {𝐵})) | ||
2-Aug-2023 | diveqap1bd 8709 | If two complex numbers are equal, their quotient is one. One-way deduction form of diveqap1 8579. Converse of diveqap1d 8672. (Contributed by David Moews, 28-Feb-2017.) (Revised by Jim Kingdon, 2-Aug-2023.) |
⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ (𝜑 → 𝐵 # 0) & ⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐴 / 𝐵) = 1) | ||
31-Jul-2023 | mul0inf 11144 | Equality of a product with zero. A bit of a curiosity, in the sense that theorems like abs00ap 10966 and mulap0bd 8532 may better express the ideas behind it. (Contributed by Jim Kingdon, 31-Jul-2023.) |
⊢ ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → ((𝐴 · 𝐵) = 0 ↔ inf({(abs‘𝐴), (abs‘𝐵)}, ℝ, < ) = 0)) | ||
31-Jul-2023 | mul0eqap 8545 | If two numbers are apart from each other and their product is zero, one of them must be zero. (Contributed by Jim Kingdon, 31-Jul-2023.) |
⊢ (𝜑 → 𝐴 ∈ ℂ) & ⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ (𝜑 → 𝐴 # 𝐵) & ⊢ (𝜑 → (𝐴 · 𝐵) = 0) ⇒ ⊢ (𝜑 → (𝐴 = 0 ∨ 𝐵 = 0)) | ||
31-Jul-2023 | apcon4bid 8500 | Contrapositive law deduction for apartness. (Contributed by Jim Kingdon, 31-Jul-2023.) |
⊢ (𝜑 → 𝐴 ∈ ℂ) & ⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ (𝜑 → 𝐶 ∈ ℂ) & ⊢ (𝜑 → 𝐷 ∈ ℂ) & ⊢ (𝜑 → (𝐴 # 𝐵 ↔ 𝐶 # 𝐷)) ⇒ ⊢ (𝜑 → (𝐴 = 𝐵 ↔ 𝐶 = 𝐷)) | ||
30-Jul-2023 | uzennn 10339 | An upper integer set is equinumerous to the set of natural numbers. (Contributed by Jim Kingdon, 30-Jul-2023.) |
⊢ (𝑀 ∈ ℤ → (ℤ≥‘𝑀) ≈ ℕ) | ||
30-Jul-2023 | djuen 7147 | Disjoint unions of equinumerous sets are equinumerous. (Contributed by Jim Kingdon, 30-Jul-2023.) |
⊢ ((𝐴 ≈ 𝐵 ∧ 𝐶 ≈ 𝐷) → (𝐴 ⊔ 𝐶) ≈ (𝐵 ⊔ 𝐷)) | ||
30-Jul-2023 | endjudisj 7146 | Equinumerosity of a disjoint union and a union of two disjoint sets. (Contributed by Jim Kingdon, 30-Jul-2023.) |
⊢ ((𝐴 ∈ 𝑉 ∧ 𝐵 ∈ 𝑊 ∧ (𝐴 ∩ 𝐵) = ∅) → (𝐴 ⊔ 𝐵) ≈ (𝐴 ∪ 𝐵)) | ||
30-Jul-2023 | eninr 7043 | Equinumerosity of a set and its image under right injection. (Contributed by Jim Kingdon, 30-Jul-2023.) |
⊢ (𝐴 ∈ 𝑉 → (inr “ 𝐴) ≈ 𝐴) | ||
30-Jul-2023 | eninl 7042 | Equinumerosity of a set and its image under left injection. (Contributed by Jim Kingdon, 30-Jul-2023.) |
⊢ (𝐴 ∈ 𝑉 → (inl “ 𝐴) ≈ 𝐴) | ||
29-Jul-2023 | exmidunben 12197 | If any unbounded set of positive integers is equinumerous to ℕ, then the Limited Principle of Omniscience (LPO) implies excluded middle. (Contributed by Jim Kingdon, 29-Jul-2023.) |
⊢ ((∀𝑥((𝑥 ⊆ ℕ ∧ ∀𝑚 ∈ ℕ ∃𝑛 ∈ 𝑥 𝑚 < 𝑛) → 𝑥 ≈ ℕ) ∧ ω ∈ Omni) → EXMID) | ||
29-Jul-2023 | exmidsssnc 4165 | Excluded middle in terms of subsets of a singleton. This is similar to exmid01 4160 but lets you choose any set as the element of the singleton rather than just ∅. It is similar to exmidsssn 4164 but for a particular set 𝐵 rather than all sets. (Contributed by Jim Kingdon, 29-Jul-2023.) |
⊢ (𝐵 ∈ 𝑉 → (EXMID ↔ ∀𝑥(𝑥 ⊆ {𝐵} → (𝑥 = ∅ ∨ 𝑥 = {𝐵})))) | ||
28-Jul-2023 | dvfcnpm 13101 | The derivative is a function. (Contributed by Mario Carneiro, 9-Feb-2015.) (Revised by Jim Kingdon, 28-Jul-2023.) |
⊢ (𝐹 ∈ (ℂ ↑pm ℂ) → (ℂ D 𝐹):dom (ℂ D 𝐹)⟶ℂ) | ||
28-Jul-2023 | dvfpm 13100 | The derivative is a function. (Contributed by Mario Carneiro, 8-Aug-2014.) (Revised by Jim Kingdon, 28-Jul-2023.) |
⊢ (𝐹 ∈ (ℂ ↑pm ℝ) → (ℝ D 𝐹):dom (ℝ D 𝐹)⟶ℂ) | ||
23-Jul-2023 | ennnfonelemhdmp1 12180 | Lemma for ennnfone 12196. Domain at a successor where we need to add an element to the sequence. (Contributed by Jim Kingdon, 23-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ (𝜑 → 𝑃 ∈ ℕ0) & ⊢ (𝜑 → ¬ (𝐹‘(◡𝑁‘𝑃)) ∈ (𝐹 “ (◡𝑁‘𝑃))) ⇒ ⊢ (𝜑 → dom (𝐻‘(𝑃 + 1)) = suc dom (𝐻‘𝑃)) | ||
23-Jul-2023 | ennnfonelemp1 12177 | Lemma for ennnfone 12196. Value of 𝐻 at a successor. (Contributed by Jim Kingdon, 23-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ (𝜑 → 𝑃 ∈ ℕ0) ⇒ ⊢ (𝜑 → (𝐻‘(𝑃 + 1)) = if((𝐹‘(◡𝑁‘𝑃)) ∈ (𝐹 “ (◡𝑁‘𝑃)), (𝐻‘𝑃), ((𝐻‘𝑃) ∪ {〈dom (𝐻‘𝑃), (𝐹‘(◡𝑁‘𝑃))〉}))) | ||
22-Jul-2023 | nntr2 6451 | Transitive law for natural numbers. (Contributed by Jim Kingdon, 22-Jul-2023.) |
⊢ ((𝐴 ∈ ω ∧ 𝐶 ∈ ω) → ((𝐴 ⊆ 𝐵 ∧ 𝐵 ∈ 𝐶) → 𝐴 ∈ 𝐶)) | ||
22-Jul-2023 | nnsssuc 6450 | A natural number is a subset of another natural number if and only if it belongs to its successor. (Contributed by Jim Kingdon, 22-Jul-2023.) |
⊢ ((𝐴 ∈ ω ∧ 𝐵 ∈ ω) → (𝐴 ⊆ 𝐵 ↔ 𝐴 ∈ suc 𝐵)) | ||
21-Jul-2023 | ennnfoneleminc 12182 | Lemma for ennnfone 12196. We only add elements to 𝐻 as the index increases. (Contributed by Jim Kingdon, 21-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ (𝜑 → 𝑃 ∈ ℕ0) & ⊢ (𝜑 → 𝑄 ∈ ℕ0) & ⊢ (𝜑 → 𝑃 ≤ 𝑄) ⇒ ⊢ (𝜑 → (𝐻‘𝑃) ⊆ (𝐻‘𝑄)) | ||
20-Jul-2023 | ennnfonelemg 12174 | Lemma for ennnfone 12196. Closure for 𝐺. (Contributed by Jim Kingdon, 20-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) ⇒ ⊢ ((𝜑 ∧ (𝑓 ∈ {𝑔 ∈ (𝐴 ↑pm ω) ∣ dom 𝑔 ∈ ω} ∧ 𝑗 ∈ ω)) → (𝑓𝐺𝑗) ∈ {𝑔 ∈ (𝐴 ↑pm ω) ∣ dom 𝑔 ∈ ω}) | ||
20-Jul-2023 | ennnfonelemjn 12173 | Lemma for ennnfone 12196. Non-initial state for 𝐽. (Contributed by Jim Kingdon, 20-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) ⇒ ⊢ ((𝜑 ∧ 𝑓 ∈ (ℤ≥‘(0 + 1))) → (𝐽‘𝑓) ∈ ω) | ||
20-Jul-2023 | ennnfonelemj0 12172 | Lemma for ennnfone 12196. Initial state for 𝐽. (Contributed by Jim Kingdon, 20-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) ⇒ ⊢ (𝜑 → (𝐽‘0) ∈ {𝑔 ∈ (𝐴 ↑pm ω) ∣ dom 𝑔 ∈ ω}) | ||
20-Jul-2023 | seqp1cd 10369 | Value of the sequence builder function at a successor. A version of seq3p1 10365 which provides two classes 𝐷 and 𝐶 for the terms and the value being accumulated, respectively. (Contributed by Jim Kingdon, 20-Jul-2023.) |
⊢ (𝜑 → 𝑁 ∈ (ℤ≥‘𝑀)) & ⊢ (𝜑 → (𝐹‘𝑀) ∈ 𝐶) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝐶 ∧ 𝑦 ∈ 𝐷)) → (𝑥 + 𝑦) ∈ 𝐶) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘(𝑀 + 1))) → (𝐹‘𝑥) ∈ 𝐷) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹)‘(𝑁 + 1)) = ((seq𝑀( + , 𝐹)‘𝑁) + (𝐹‘(𝑁 + 1)))) | ||
20-Jul-2023 | seqovcd 10366 | A closure law for the recursive sequence builder. This is a lemma for theorems such as seqf2 10367 and seq1cd 10368 and is unlikely to be needed once such theorems are proved. (Contributed by Jim Kingdon, 20-Jul-2023.) |
⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘(𝑀 + 1))) → (𝐹‘𝑥) ∈ 𝐷) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝐶 ∧ 𝑦 ∈ 𝐷)) → (𝑥 + 𝑦) ∈ 𝐶) ⇒ ⊢ ((𝜑 ∧ (𝑥 ∈ (ℤ≥‘𝑀) ∧ 𝑦 ∈ 𝐶)) → (𝑥(𝑧 ∈ (ℤ≥‘𝑀), 𝑤 ∈ 𝐶 ↦ (𝑤 + (𝐹‘(𝑧 + 1))))𝑦) ∈ 𝐶) | ||
19-Jul-2023 | ennnfonelemhom 12186 | Lemma for ennnfone 12196. The sequences in 𝐻 increase in length without bound if you go out far enough. (Contributed by Jim Kingdon, 19-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ (𝜑 → 𝑀 ∈ ω) ⇒ ⊢ (𝜑 → ∃𝑖 ∈ ℕ0 𝑀 ∈ dom (𝐻‘𝑖)) | ||
19-Jul-2023 | ennnfonelemex 12185 | Lemma for ennnfone 12196. Extending the sequence (𝐻‘𝑃) to include an additional element. (Contributed by Jim Kingdon, 19-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ (𝜑 → 𝑃 ∈ ℕ0) ⇒ ⊢ (𝜑 → ∃𝑖 ∈ ℕ0 dom (𝐻‘𝑃) ∈ dom (𝐻‘𝑖)) | ||
19-Jul-2023 | ennnfonelemkh 12183 | Lemma for ennnfone 12196. Because we add zero or one entries for each new index, the length of each sequence is no greater than its index. (Contributed by Jim Kingdon, 19-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ (𝜑 → 𝑃 ∈ ℕ0) ⇒ ⊢ (𝜑 → dom (𝐻‘𝑃) ⊆ (◡𝑁‘𝑃)) | ||
19-Jul-2023 | ennnfonelemom 12179 | Lemma for ennnfone 12196. 𝐻 yields finite sequences. (Contributed by Jim Kingdon, 19-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ (𝜑 → 𝑃 ∈ ℕ0) ⇒ ⊢ (𝜑 → dom (𝐻‘𝑃) ∈ ω) | ||
19-Jul-2023 | ennnfonelem1 12178 | Lemma for ennnfone 12196. Second value. (Contributed by Jim Kingdon, 19-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) ⇒ ⊢ (𝜑 → (𝐻‘1) = {〈∅, (𝐹‘∅)〉}) | ||
19-Jul-2023 | seq1cd 10368 | Initial value of the recursive sequence builder. A version of seq3-1 10363 which provides two classes 𝐷 and 𝐶 for the terms and the value being accumulated, respectively. (Contributed by Jim Kingdon, 19-Jul-2023.) |
⊢ (𝜑 → (𝐹‘𝑀) ∈ 𝐶) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝐶 ∧ 𝑦 ∈ 𝐷)) → (𝑥 + 𝑦) ∈ 𝐶) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ ((𝜑 ∧ 𝑥 ∈ (ℤ≥‘(𝑀 + 1))) → (𝐹‘𝑥) ∈ 𝐷) ⇒ ⊢ (𝜑 → (seq𝑀( + , 𝐹)‘𝑀) = (𝐹‘𝑀)) | ||
17-Jul-2023 | ennnfonelemhf1o 12184 | Lemma for ennnfone 12196. Each of the functions in 𝐻 is one to one and onto an image of 𝐹. (Contributed by Jim Kingdon, 17-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ (𝜑 → 𝑃 ∈ ℕ0) ⇒ ⊢ (𝜑 → (𝐻‘𝑃):dom (𝐻‘𝑃)–1-1-onto→(𝐹 “ (◡𝑁‘𝑃))) | ||
16-Jul-2023 | ennnfonelemen 12192 | Lemma for ennnfone 12196. The result. (Contributed by Jim Kingdon, 16-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ 𝐿 = ∪ 𝑖 ∈ ℕ0 (𝐻‘𝑖) ⇒ ⊢ (𝜑 → 𝐴 ≈ ℕ) | ||
16-Jul-2023 | ennnfonelemdm 12191 | Lemma for ennnfone 12196. The function 𝐿 is defined everywhere. (Contributed by Jim Kingdon, 16-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ 𝐿 = ∪ 𝑖 ∈ ℕ0 (𝐻‘𝑖) ⇒ ⊢ (𝜑 → dom 𝐿 = ω) | ||
16-Jul-2023 | ennnfonelemrn 12190 | Lemma for ennnfone 12196. 𝐿 is onto 𝐴. (Contributed by Jim Kingdon, 16-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ 𝐿 = ∪ 𝑖 ∈ ℕ0 (𝐻‘𝑖) ⇒ ⊢ (𝜑 → ran 𝐿 = 𝐴) | ||
16-Jul-2023 | ennnfonelemf1 12189 | Lemma for ennnfone 12196. 𝐿 is one-to-one. (Contributed by Jim Kingdon, 16-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, (𝐹‘𝑦)〉}))) & ⊢ 𝑁 = frec((𝑥 ∈ ℤ ↦ (𝑥 + 1)), 0) & ⊢ 𝐽 = (𝑥 ∈ ℕ0 ↦ if(𝑥 = 0, ∅, (◡𝑁‘(𝑥 − 1)))) & ⊢ 𝐻 = seq0(𝐺, 𝐽) & ⊢ 𝐿 = ∪ 𝑖 ∈ ℕ0 (𝐻‘𝑖) ⇒ ⊢ (𝜑 → 𝐿:dom 𝐿–1-1→𝐴) | ||
16-Jul-2023 | ennnfonelemfun 12188 | Lemma for ennnfone 12196. 𝐿 is a function. (Contributed by Jim Kingdon, 16-Jul-2023.) |
⊢ (𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 DECID 𝑥 = 𝑦) & ⊢ (𝜑 → 𝐹:ω–onto→𝐴) & ⊢ (𝜑 → ∀𝑛 ∈ ω ∃𝑘 ∈ ω ∀𝑗 ∈ suc 𝑛(𝐹‘𝑘) ≠ (𝐹‘𝑗)) & ⊢ 𝐺 = (𝑥 ∈ (𝐴 ↑pm ω), 𝑦 ∈ ω ↦ if((𝐹‘𝑦) ∈ (𝐹 “ 𝑦), 𝑥, (𝑥 ∪ {〈dom 𝑥, ( |