| Intuitionistic Logic Explorer Theorem List (p. 149 of 158)  | < Previous Next > | |
| Bad symbols? Try the
 GIF version.  | 
||
| 
 Mirrors > Metamath Home Page > ILE Home Page > Theorem List Contents > Recent Proofs This page: Page List  | 
||
| Type | Label | Description | 
|---|---|---|
| Statement | ||
| Theorem | divcnap 14801* | 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 𝐽) | ||
| Theorem | mpomulcn 14802* | Complex number multiplication is a continuous function. (Contributed by GG, 16-Mar-2025.) | 
| ⊢ 𝐽 = (TopOpen‘ℂfld) ⇒ ⊢ (𝑥 ∈ ℂ, 𝑦 ∈ ℂ ↦ (𝑥 · 𝑦)) ∈ ((𝐽 ×t 𝐽) Cn 𝐽) | ||
| Theorem | fsumcncntop 14803* | A finite sum of functions to complex numbers from a common topological space is continuous. The class expression for 𝐵 normally contains free variables 𝑘 and 𝑥 to index it. (Contributed by NM, 8-Aug-2007.) (Revised by Mario Carneiro, 23-Aug-2014.) | 
| ⊢ 𝐾 = (MetOpen‘(abs ∘ − )) & ⊢ (𝜑 → 𝐽 ∈ (TopOn‘𝑋)) & ⊢ (𝜑 → 𝐴 ∈ Fin) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → (𝑥 ∈ 𝑋 ↦ 𝐵) ∈ (𝐽 Cn 𝐾)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ Σ𝑘 ∈ 𝐴 𝐵) ∈ (𝐽 Cn 𝐾)) | ||
| Theorem | fsumcn 14804* | A finite sum of functions to complex numbers from a common topological space is continuous. The class expression for 𝐵 normally contains free variables 𝑘 and 𝑥 to index it. (Contributed by NM, 8-Aug-2007.) (Revised by Mario Carneiro, 23-Aug-2014.) | 
| ⊢ 𝐾 = (TopOpen‘ℂfld) & ⊢ (𝜑 → 𝐽 ∈ (TopOn‘𝑋)) & ⊢ (𝜑 → 𝐴 ∈ Fin) & ⊢ ((𝜑 ∧ 𝑘 ∈ 𝐴) → (𝑥 ∈ 𝑋 ↦ 𝐵) ∈ (𝐽 Cn 𝐾)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ Σ𝑘 ∈ 𝐴 𝐵) ∈ (𝐽 Cn 𝐾)) | ||
| Theorem | expcn 14805* | The power function on complex numbers, for fixed exponent 𝑁, is continuous. (Contributed by Mario Carneiro, 5-May-2014.) (Revised by Mario Carneiro, 23-Aug-2014.) Avoid ax-mulf 8002. (Revised by GG, 16-Mar-2025.) | 
| ⊢ 𝐽 = (TopOpen‘ℂfld) ⇒ ⊢ (𝑁 ∈ ℕ0 → (𝑥 ∈ ℂ ↦ (𝑥↑𝑁)) ∈ (𝐽 Cn 𝐽)) | ||
| Syntax | ccncf 14806 | Extend class notation to include the operation which returns a class of continuous complex functions. | 
| class –cn→ | ||
| Definition | df-cncf 14807* | Define the operation whose value is a class of continuous complex functions. (Contributed by Paul Chapman, 11-Oct-2007.) | 
| ⊢ –cn→ = (𝑎 ∈ 𝒫 ℂ, 𝑏 ∈ 𝒫 ℂ ↦ {𝑓 ∈ (𝑏 ↑𝑚 𝑎) ∣ ∀𝑥 ∈ 𝑎 ∀𝑒 ∈ ℝ+ ∃𝑑 ∈ ℝ+ ∀𝑦 ∈ 𝑎 ((abs‘(𝑥 − 𝑦)) < 𝑑 → (abs‘((𝑓‘𝑥) − (𝑓‘𝑦))) < 𝑒)}) | ||
| Theorem | cncfval 14808* | The value of the continuous complex function operation is the set of continuous functions from 𝐴 to 𝐵. (Contributed by Paul Chapman, 11-Oct-2007.) (Revised by Mario Carneiro, 9-Nov-2013.) | 
| ⊢ ((𝐴 ⊆ ℂ ∧ 𝐵 ⊆ ℂ) → (𝐴–cn→𝐵) = {𝑓 ∈ (𝐵 ↑𝑚 𝐴) ∣ ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ ℝ+ ∃𝑧 ∈ ℝ+ ∀𝑤 ∈ 𝐴 ((abs‘(𝑥 − 𝑤)) < 𝑧 → (abs‘((𝑓‘𝑥) − (𝑓‘𝑤))) < 𝑦)}) | ||
| Theorem | elcncf 14809* | Membership in the set of continuous complex functions from 𝐴 to 𝐵. (Contributed by Paul Chapman, 11-Oct-2007.) (Revised by Mario Carneiro, 9-Nov-2013.) | 
| ⊢ ((𝐴 ⊆ ℂ ∧ 𝐵 ⊆ ℂ) → (𝐹 ∈ (𝐴–cn→𝐵) ↔ (𝐹:𝐴⟶𝐵 ∧ ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ ℝ+ ∃𝑧 ∈ ℝ+ ∀𝑤 ∈ 𝐴 ((abs‘(𝑥 − 𝑤)) < 𝑧 → (abs‘((𝐹‘𝑥) − (𝐹‘𝑤))) < 𝑦)))) | ||
| Theorem | elcncf2 14810* | Version of elcncf 14809 with arguments commuted. (Contributed by Mario Carneiro, 28-Apr-2014.) | 
| ⊢ ((𝐴 ⊆ ℂ ∧ 𝐵 ⊆ ℂ) → (𝐹 ∈ (𝐴–cn→𝐵) ↔ (𝐹:𝐴⟶𝐵 ∧ ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ ℝ+ ∃𝑧 ∈ ℝ+ ∀𝑤 ∈ 𝐴 ((abs‘(𝑤 − 𝑥)) < 𝑧 → (abs‘((𝐹‘𝑤) − (𝐹‘𝑥))) < 𝑦)))) | ||
| Theorem | cncfrss 14811 | Reverse closure of the continuous function predicate. (Contributed by Mario Carneiro, 25-Aug-2014.) | 
| ⊢ (𝐹 ∈ (𝐴–cn→𝐵) → 𝐴 ⊆ ℂ) | ||
| Theorem | cncfrss2 14812 | Reverse closure of the continuous function predicate. (Contributed by Mario Carneiro, 25-Aug-2014.) | 
| ⊢ (𝐹 ∈ (𝐴–cn→𝐵) → 𝐵 ⊆ ℂ) | ||
| Theorem | cncff 14813 | A continuous complex function's domain and codomain. (Contributed by Paul Chapman, 17-Jan-2008.) (Revised by Mario Carneiro, 25-Aug-2014.) | 
| ⊢ (𝐹 ∈ (𝐴–cn→𝐵) → 𝐹:𝐴⟶𝐵) | ||
| Theorem | cncfi 14814* | Defining property of a continuous function. (Contributed by Mario Carneiro, 30-Apr-2014.) (Revised by Mario Carneiro, 25-Aug-2014.) | 
| ⊢ ((𝐹 ∈ (𝐴–cn→𝐵) ∧ 𝐶 ∈ 𝐴 ∧ 𝑅 ∈ ℝ+) → ∃𝑧 ∈ ℝ+ ∀𝑤 ∈ 𝐴 ((abs‘(𝑤 − 𝐶)) < 𝑧 → (abs‘((𝐹‘𝑤) − (𝐹‘𝐶))) < 𝑅)) | ||
| Theorem | elcncf1di 14815* | Membership in the set of continuous complex functions from 𝐴 to 𝐵. (Contributed by Paul Chapman, 26-Nov-2007.) | 
| ⊢ (𝜑 → 𝐹:𝐴⟶𝐵) & ⊢ (𝜑 → ((𝑥 ∈ 𝐴 ∧ 𝑦 ∈ ℝ+) → 𝑍 ∈ ℝ+)) & ⊢ (𝜑 → (((𝑥 ∈ 𝐴 ∧ 𝑤 ∈ 𝐴) ∧ 𝑦 ∈ ℝ+) → ((abs‘(𝑥 − 𝑤)) < 𝑍 → (abs‘((𝐹‘𝑥) − (𝐹‘𝑤))) < 𝑦))) ⇒ ⊢ (𝜑 → ((𝐴 ⊆ ℂ ∧ 𝐵 ⊆ ℂ) → 𝐹 ∈ (𝐴–cn→𝐵))) | ||
| Theorem | elcncf1ii 14816* | Membership in the set of continuous complex functions from 𝐴 to 𝐵. (Contributed by Paul Chapman, 26-Nov-2007.) | 
| ⊢ 𝐹:𝐴⟶𝐵 & ⊢ ((𝑥 ∈ 𝐴 ∧ 𝑦 ∈ ℝ+) → 𝑍 ∈ ℝ+) & ⊢ (((𝑥 ∈ 𝐴 ∧ 𝑤 ∈ 𝐴) ∧ 𝑦 ∈ ℝ+) → ((abs‘(𝑥 − 𝑤)) < 𝑍 → (abs‘((𝐹‘𝑥) − (𝐹‘𝑤))) < 𝑦)) ⇒ ⊢ ((𝐴 ⊆ ℂ ∧ 𝐵 ⊆ ℂ) → 𝐹 ∈ (𝐴–cn→𝐵)) | ||
| Theorem | rescncf 14817 | A continuous complex function restricted to a subset is continuous. (Contributed by Paul Chapman, 18-Oct-2007.) (Revised by Mario Carneiro, 25-Aug-2014.) | 
| ⊢ (𝐶 ⊆ 𝐴 → (𝐹 ∈ (𝐴–cn→𝐵) → (𝐹 ↾ 𝐶) ∈ (𝐶–cn→𝐵))) | ||
| Theorem | cncfcdm 14818 | Change the codomain of a continuous complex function. (Contributed by Paul Chapman, 18-Oct-2007.) (Revised by Mario Carneiro, 1-May-2015.) | 
| ⊢ ((𝐶 ⊆ ℂ ∧ 𝐹 ∈ (𝐴–cn→𝐵)) → (𝐹 ∈ (𝐴–cn→𝐶) ↔ 𝐹:𝐴⟶𝐶)) | ||
| Theorem | cncfss 14819 | The set of continuous functions is expanded when the codomain is expanded. (Contributed by Mario Carneiro, 30-Aug-2014.) | 
| ⊢ ((𝐵 ⊆ 𝐶 ∧ 𝐶 ⊆ ℂ) → (𝐴–cn→𝐵) ⊆ (𝐴–cn→𝐶)) | ||
| Theorem | climcncf 14820 | Image of a limit under a continuous map. (Contributed by Mario Carneiro, 7-Apr-2015.) | 
| ⊢ 𝑍 = (ℤ≥‘𝑀) & ⊢ (𝜑 → 𝑀 ∈ ℤ) & ⊢ (𝜑 → 𝐹 ∈ (𝐴–cn→𝐵)) & ⊢ (𝜑 → 𝐺:𝑍⟶𝐴) & ⊢ (𝜑 → 𝐺 ⇝ 𝐷) & ⊢ (𝜑 → 𝐷 ∈ 𝐴) ⇒ ⊢ (𝜑 → (𝐹 ∘ 𝐺) ⇝ (𝐹‘𝐷)) | ||
| Theorem | abscncf 14821 | Absolute value is continuous. (Contributed by Paul Chapman, 21-Oct-2007.) (Revised by Mario Carneiro, 28-Apr-2014.) | 
| ⊢ abs ∈ (ℂ–cn→ℝ) | ||
| Theorem | recncf 14822 | Real part is continuous. (Contributed by Paul Chapman, 21-Oct-2007.) (Revised by Mario Carneiro, 28-Apr-2014.) | 
| ⊢ ℜ ∈ (ℂ–cn→ℝ) | ||
| Theorem | imcncf 14823 | Imaginary part is continuous. (Contributed by Paul Chapman, 21-Oct-2007.) (Revised by Mario Carneiro, 28-Apr-2014.) | 
| ⊢ ℑ ∈ (ℂ–cn→ℝ) | ||
| Theorem | cjcncf 14824 | Complex conjugate is continuous. (Contributed by Paul Chapman, 21-Oct-2007.) (Revised by Mario Carneiro, 28-Apr-2014.) | 
| ⊢ ∗ ∈ (ℂ–cn→ℂ) | ||
| Theorem | mulc1cncf 14825* | Multiplication by a constant is continuous. (Contributed by Paul Chapman, 28-Nov-2007.) (Revised by Mario Carneiro, 30-Apr-2014.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℂ ↦ (𝐴 · 𝑥)) ⇒ ⊢ (𝐴 ∈ ℂ → 𝐹 ∈ (ℂ–cn→ℂ)) | ||
| Theorem | divccncfap 14826* | Division by a constant is continuous. (Contributed by Paul Chapman, 28-Nov-2007.) (Revised by Jim Kingdon, 9-Jan-2023.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℂ ↦ (𝑥 / 𝐴)) ⇒ ⊢ ((𝐴 ∈ ℂ ∧ 𝐴 # 0) → 𝐹 ∈ (ℂ–cn→ℂ)) | ||
| Theorem | cncfco 14827 | The composition of two continuous maps on complex numbers is also continuous. (Contributed by Jeff Madsen, 2-Sep-2009.) (Revised by Mario Carneiro, 25-Aug-2014.) | 
| ⊢ (𝜑 → 𝐹 ∈ (𝐴–cn→𝐵)) & ⊢ (𝜑 → 𝐺 ∈ (𝐵–cn→𝐶)) ⇒ ⊢ (𝜑 → (𝐺 ∘ 𝐹) ∈ (𝐴–cn→𝐶)) | ||
| Theorem | cncfmet 14828 | Relate complex function continuity to metric space continuity. (Contributed by Paul Chapman, 26-Nov-2007.) (Revised by Mario Carneiro, 7-Sep-2015.) | 
| ⊢ 𝐶 = ((abs ∘ − ) ↾ (𝐴 × 𝐴)) & ⊢ 𝐷 = ((abs ∘ − ) ↾ (𝐵 × 𝐵)) & ⊢ 𝐽 = (MetOpen‘𝐶) & ⊢ 𝐾 = (MetOpen‘𝐷) ⇒ ⊢ ((𝐴 ⊆ ℂ ∧ 𝐵 ⊆ ℂ) → (𝐴–cn→𝐵) = (𝐽 Cn 𝐾)) | ||
| Theorem | cncfcncntop 14829 | Relate complex function continuity to topological continuity. (Contributed by Mario Carneiro, 17-Feb-2015.) | 
| ⊢ 𝐽 = (MetOpen‘(abs ∘ − )) & ⊢ 𝐾 = (𝐽 ↾t 𝐴) & ⊢ 𝐿 = (𝐽 ↾t 𝐵) ⇒ ⊢ ((𝐴 ⊆ ℂ ∧ 𝐵 ⊆ ℂ) → (𝐴–cn→𝐵) = (𝐾 Cn 𝐿)) | ||
| Theorem | cncfcn1cntop 14830 | Relate complex function continuity to topological continuity. (Contributed by Paul Chapman, 28-Nov-2007.) (Revised by Mario Carneiro, 7-Sep-2015.) (Revised by Jim Kingdon, 16-Jun-2023.) | 
| ⊢ 𝐽 = (MetOpen‘(abs ∘ − )) ⇒ ⊢ (ℂ–cn→ℂ) = (𝐽 Cn 𝐽) | ||
| Theorem | cncfcn1 14831 | Relate complex function continuity to topological continuity. (Contributed by Paul Chapman, 28-Nov-2007.) (Revised by Mario Carneiro, 7-Sep-2015.) | 
| ⊢ 𝐽 = (TopOpen‘ℂfld) ⇒ ⊢ (ℂ–cn→ℂ) = (𝐽 Cn 𝐽) | ||
| Theorem | cncfmptc 14832* | A constant function is a continuous function on ℂ. (Contributed by Jeff Madsen, 2-Sep-2009.) (Revised by Mario Carneiro, 7-Sep-2015.) | 
| ⊢ ((𝐴 ∈ 𝑇 ∧ 𝑆 ⊆ ℂ ∧ 𝑇 ⊆ ℂ) → (𝑥 ∈ 𝑆 ↦ 𝐴) ∈ (𝑆–cn→𝑇)) | ||
| Theorem | cncfmptid 14833* | The identity function is a continuous function on ℂ. (Contributed by Jeff Madsen, 11-Jun-2010.) (Revised by Mario Carneiro, 17-May-2016.) | 
| ⊢ ((𝑆 ⊆ 𝑇 ∧ 𝑇 ⊆ ℂ) → (𝑥 ∈ 𝑆 ↦ 𝑥) ∈ (𝑆–cn→𝑇)) | ||
| Theorem | cncfmpt1f 14834* | Composition of continuous functions. –cn→ analogue of cnmpt11f 14520. (Contributed by Mario Carneiro, 3-Sep-2014.) | 
| ⊢ (𝜑 → 𝐹 ∈ (ℂ–cn→ℂ)) & ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐴) ∈ (𝑋–cn→ℂ)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ (𝐹‘𝐴)) ∈ (𝑋–cn→ℂ)) | ||
| Theorem | cncfmpt2fcntop 14835* | Composition of continuous functions. –cn→ analogue of cnmpt12f 14522. (Contributed by Mario Carneiro, 3-Sep-2014.) | 
| ⊢ 𝐽 = (MetOpen‘(abs ∘ − )) & ⊢ (𝜑 → 𝐹 ∈ ((𝐽 ×t 𝐽) Cn 𝐽)) & ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐴) ∈ (𝑋–cn→ℂ)) & ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐵) ∈ (𝑋–cn→ℂ)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ (𝐴𝐹𝐵)) ∈ (𝑋–cn→ℂ)) | ||
| Theorem | addccncf 14836* | Adding a constant is a continuous function. (Contributed by Jeff Madsen, 2-Sep-2009.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℂ ↦ (𝑥 + 𝐴)) ⇒ ⊢ (𝐴 ∈ ℂ → 𝐹 ∈ (ℂ–cn→ℂ)) | ||
| Theorem | idcncf 14837 | The identity function is a continuous function on ℂ. (Contributed by Jeff Madsen, 11-Jun-2010.) (Moved into main set.mm as cncfmptid 14833 and may be deleted by mathbox owner, JM. --MC 12-Sep-2015.) (Revised by Mario Carneiro, 12-Sep-2015.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℂ ↦ 𝑥) ⇒ ⊢ 𝐹 ∈ (ℂ–cn→ℂ) | ||
| Theorem | sub1cncf 14838* | Subtracting a constant is a continuous function. (Contributed by Jeff Madsen, 2-Sep-2009.) (Proof shortened by Mario Carneiro, 12-Sep-2015.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℂ ↦ (𝑥 − 𝐴)) ⇒ ⊢ (𝐴 ∈ ℂ → 𝐹 ∈ (ℂ–cn→ℂ)) | ||
| Theorem | sub2cncf 14839* | Subtraction from a constant is a continuous function. (Contributed by Jeff Madsen, 2-Sep-2009.) (Proof shortened by Mario Carneiro, 12-Sep-2015.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℂ ↦ (𝐴 − 𝑥)) ⇒ ⊢ (𝐴 ∈ ℂ → 𝐹 ∈ (ℂ–cn→ℂ)) | ||
| Theorem | cdivcncfap 14840* | Division with a constant numerator is continuous. (Contributed by Mario Carneiro, 28-Dec-2016.) (Revised by Jim Kingdon, 26-May-2023.) | 
| ⊢ 𝐹 = (𝑥 ∈ {𝑦 ∈ ℂ ∣ 𝑦 # 0} ↦ (𝐴 / 𝑥)) ⇒ ⊢ (𝐴 ∈ ℂ → 𝐹 ∈ ({𝑦 ∈ ℂ ∣ 𝑦 # 0}–cn→ℂ)) | ||
| Theorem | negcncf 14841* | The negative function is continuous. (Contributed by Mario Carneiro, 30-Dec-2016.) | 
| ⊢ 𝐹 = (𝑥 ∈ 𝐴 ↦ -𝑥) ⇒ ⊢ (𝐴 ⊆ ℂ → 𝐹 ∈ (𝐴–cn→ℂ)) | ||
| Theorem | negfcncf 14842* | The negative of a continuous complex function is continuous. (Contributed by Paul Chapman, 21-Jan-2008.) (Revised by Mario Carneiro, 25-Aug-2014.) | 
| ⊢ 𝐺 = (𝑥 ∈ 𝐴 ↦ -(𝐹‘𝑥)) ⇒ ⊢ (𝐹 ∈ (𝐴–cn→ℂ) → 𝐺 ∈ (𝐴–cn→ℂ)) | ||
| Theorem | mulcncflem 14843* | Lemma for mulcncf 14844. (Contributed by Jim Kingdon, 29-May-2023.) | 
| ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐴) ∈ (𝑋–cn→ℂ)) & ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐵) ∈ (𝑋–cn→ℂ)) & ⊢ (𝜑 → 𝑉 ∈ 𝑋) & ⊢ (𝜑 → 𝐸 ∈ ℝ+) & ⊢ (𝜑 → 𝐹 ∈ ℝ+) & ⊢ (𝜑 → 𝐺 ∈ ℝ+) & ⊢ (𝜑 → 𝑆 ∈ ℝ+) & ⊢ (𝜑 → 𝑇 ∈ ℝ+) & ⊢ (𝜑 → ∀𝑢 ∈ 𝑋 ((abs‘(𝑢 − 𝑉)) < 𝑆 → (abs‘(((𝑥 ∈ 𝑋 ↦ 𝐴)‘𝑢) − ((𝑥 ∈ 𝑋 ↦ 𝐴)‘𝑉))) < 𝐹)) & ⊢ (𝜑 → ∀𝑢 ∈ 𝑋 ((abs‘(𝑢 − 𝑉)) < 𝑇 → (abs‘(((𝑥 ∈ 𝑋 ↦ 𝐵)‘𝑢) − ((𝑥 ∈ 𝑋 ↦ 𝐵)‘𝑉))) < 𝐺)) & ⊢ (𝜑 → ∀𝑢 ∈ 𝑋 (((abs‘(⦋𝑢 / 𝑥⦌𝐴 − ⦋𝑉 / 𝑥⦌𝐴)) < 𝐹 ∧ (abs‘(⦋𝑢 / 𝑥⦌𝐵 − ⦋𝑉 / 𝑥⦌𝐵)) < 𝐺) → (abs‘((⦋𝑢 / 𝑥⦌𝐴 · ⦋𝑢 / 𝑥⦌𝐵) − (⦋𝑉 / 𝑥⦌𝐴 · ⦋𝑉 / 𝑥⦌𝐵))) < 𝐸)) ⇒ ⊢ (𝜑 → ∃𝑑 ∈ ℝ+ ∀𝑢 ∈ 𝑋 ((abs‘(𝑢 − 𝑉)) < 𝑑 → (abs‘(((𝑥 ∈ 𝑋 ↦ (𝐴 · 𝐵))‘𝑢) − ((𝑥 ∈ 𝑋 ↦ (𝐴 · 𝐵))‘𝑉))) < 𝐸)) | ||
| Theorem | mulcncf 14844* | The multiplication of two continuous complex functions is continuous. (Contributed by Glauco Siliprandi, 29-Jun-2017.) | 
| ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐴) ∈ (𝑋–cn→ℂ)) & ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐵) ∈ (𝑋–cn→ℂ)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ (𝐴 · 𝐵)) ∈ (𝑋–cn→ℂ)) | ||
| Theorem | expcncf 14845* | The power function on complex numbers, for fixed exponent N, is continuous. (Contributed by Glauco Siliprandi, 29-Jun-2017.) | 
| ⊢ (𝑁 ∈ ℕ0 → (𝑥 ∈ ℂ ↦ (𝑥↑𝑁)) ∈ (ℂ–cn→ℂ)) | ||
| Theorem | cnrehmeocntop 14846* | The canonical bijection from (ℝ × ℝ) to ℂ described in cnref1o 9725 is in fact a homeomorphism of the usual topologies on these sets. (It is also an isometry, if (ℝ × ℝ) is metrized with the l<SUP>2</SUP> norm.) (Contributed by Mario Carneiro, 25-Aug-2014.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℝ, 𝑦 ∈ ℝ ↦ (𝑥 + (i · 𝑦))) & ⊢ 𝐽 = (topGen‘ran (,)) & ⊢ 𝐾 = (MetOpen‘(abs ∘ − )) ⇒ ⊢ 𝐹 ∈ ((𝐽 ×t 𝐽)Homeo𝐾) | ||
| Theorem | cnopnap 14847* | The complex numbers apart from a given complex number form an open set. (Contributed by Jim Kingdon, 14-Dec-2023.) | 
| ⊢ (𝐴 ∈ ℂ → {𝑤 ∈ ℂ ∣ 𝑤 # 𝐴} ∈ (MetOpen‘(abs ∘ − ))) | ||
| Theorem | addcncf 14848* | The addition of two continuous complex functions is continuous. (Contributed by Glauco Siliprandi, 11-Dec-2019.) | 
| ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐴) ∈ (𝑋–cn→ℂ)) & ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐵) ∈ (𝑋–cn→ℂ)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ (𝐴 + 𝐵)) ∈ (𝑋–cn→ℂ)) | ||
| Theorem | subcncf 14849* | The subtraction of two continuous complex functions is continuous. (Contributed by Glauco Siliprandi, 11-Dec-2019.) | 
| ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐴) ∈ (𝑋–cn→ℂ)) & ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐵) ∈ (𝑋–cn→ℂ)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ (𝐴 − 𝐵)) ∈ (𝑋–cn→ℂ)) | ||
| Theorem | divcncfap 14850* | The quotient of two continuous complex functions is continuous. (Contributed by Glauco Siliprandi, 11-Dec-2019.) | 
| ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐴) ∈ (𝑋–cn→ℂ)) & ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐵) ∈ (𝑋–cn→{𝑦 ∈ ℂ ∣ 𝑦 # 0})) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ (𝐴 / 𝐵)) ∈ (𝑋–cn→ℂ)) | ||
| Theorem | maxcncf 14851* | The maximum of two continuous real functions is continuous. (Contributed by Jim Kingdon, 18-Jul-2025.) | 
| ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐴) ∈ (𝑋–cn→ℝ)) & ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐵) ∈ (𝑋–cn→ℝ)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ sup({𝐴, 𝐵}, ℝ, < )) ∈ (𝑋–cn→ℝ)) | ||
| Theorem | mincncf 14852* | The minimum of two continuous real functions is continuous. (Contributed by Jim Kingdon, 19-Jul-2025.) | 
| ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐴) ∈ (𝑋–cn→ℝ)) & ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ 𝐵) ∈ (𝑋–cn→ℝ)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ inf({𝐴, 𝐵}, ℝ, < )) ∈ (𝑋–cn→ℝ)) | ||
| Theorem | dedekindeulemuub 14853* | Lemma for dedekindeu 14859. Any element of the upper cut is an upper bound for the lower cut. (Contributed by Jim Kingdon, 2-Feb-2024.) | 
| ⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 ∈ 𝑈) ⇒ ⊢ (𝜑 → ∀𝑧 ∈ 𝐿 𝑧 < 𝐴) | ||
| Theorem | dedekindeulemub 14854* | Lemma for dedekindeu 14859. The lower cut has an upper bound. (Contributed by Jim Kingdon, 31-Jan-2024.) | 
| ⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℝ ∀𝑦 ∈ 𝐿 𝑦 < 𝑥) | ||
| Theorem | dedekindeulemloc 14855* | Lemma for dedekindeu 14859. The set L is located. (Contributed by Jim Kingdon, 31-Jan-2024.) | 
| ⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 < 𝑦 → (∃𝑧 ∈ 𝐿 𝑥 < 𝑧 ∨ ∀𝑧 ∈ 𝐿 𝑧 < 𝑦))) | ||
| Theorem | dedekindeulemlub 14856* | Lemma for dedekindeu 14859. The set L has a least upper bound. (Contributed by Jim Kingdon, 31-Jan-2024.) | 
| ⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℝ (∀𝑦 ∈ 𝐿 ¬ 𝑥 < 𝑦 ∧ ∀𝑦 ∈ ℝ (𝑦 < 𝑥 → ∃𝑧 ∈ 𝐿 𝑦 < 𝑧))) | ||
| Theorem | dedekindeulemlu 14857* | Lemma for dedekindeu 14859. There is a number which separates the lower and upper cuts. (Contributed by Jim Kingdon, 31-Jan-2024.) | 
| ⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℝ (∀𝑞 ∈ 𝐿 𝑞 < 𝑥 ∧ ∀𝑟 ∈ 𝑈 𝑥 < 𝑟)) | ||
| Theorem | dedekindeulemeu 14858* | Lemma for dedekindeu 14859. Part of proving uniqueness. (Contributed by Jim Kingdon, 31-Jan-2024.) | 
| ⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → (∀𝑞 ∈ 𝐿 𝑞 < 𝐴 ∧ ∀𝑟 ∈ 𝑈 𝐴 < 𝑟)) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → (∀𝑞 ∈ 𝐿 𝑞 < 𝐵 ∧ ∀𝑟 ∈ 𝑈 𝐵 < 𝑟)) & ⊢ (𝜑 → 𝐴 < 𝐵) ⇒ ⊢ (𝜑 → ⊥) | ||
| Theorem | dedekindeu 14859* | A Dedekind cut identifies a unique real number. Similar to df-inp 7533 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.) | 
| ⊢ (𝜑 → 𝐿 ⊆ ℝ) & ⊢ (𝜑 → 𝑈 ⊆ ℝ) & ⊢ (𝜑 → ∃𝑞 ∈ ℝ 𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ ℝ 𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ (𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ ℝ (𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ ℝ ∀𝑟 ∈ ℝ (𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∃!𝑥 ∈ ℝ (∀𝑞 ∈ 𝐿 𝑞 < 𝑥 ∧ ∀𝑟 ∈ 𝑈 𝑥 < 𝑟)) | ||
| Theorem | suplociccreex 14860* | An inhabited, bounded-above, located set of reals in a closed interval has a supremum. A similar theorem is axsuploc 8099 but that one is for the entire real line rather than a closed interval. (Contributed by Jim Kingdon, 14-Feb-2024.) | 
| ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐶 ∈ ℝ) & ⊢ (𝜑 → 𝐵 < 𝐶) & ⊢ (𝜑 → 𝐴 ⊆ (𝐵[,]𝐶)) & ⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥 ∈ (𝐵[,]𝐶)∀𝑦 ∈ (𝐵[,]𝐶)(𝑥 < 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 < 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 < 𝑦))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ ℝ (∀𝑦 ∈ 𝐴 ¬ 𝑥 < 𝑦 ∧ ∀𝑦 ∈ ℝ (𝑦 < 𝑥 → ∃𝑧 ∈ 𝐴 𝑦 < 𝑧))) | ||
| Theorem | suplociccex 14861* | An inhabited, bounded-above, located set of reals in a closed interval has a supremum. A similar theorem is axsuploc 8099 but that one is for the entire real line rather than a closed interval. (Contributed by Jim Kingdon, 14-Feb-2024.) | 
| ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐶 ∈ ℝ) & ⊢ (𝜑 → 𝐵 < 𝐶) & ⊢ (𝜑 → 𝐴 ⊆ (𝐵[,]𝐶)) & ⊢ (𝜑 → ∃𝑥 𝑥 ∈ 𝐴) & ⊢ (𝜑 → ∀𝑥 ∈ (𝐵[,]𝐶)∀𝑦 ∈ (𝐵[,]𝐶)(𝑥 < 𝑦 → (∃𝑧 ∈ 𝐴 𝑥 < 𝑧 ∨ ∀𝑧 ∈ 𝐴 𝑧 < 𝑦))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ (𝐵[,]𝐶)(∀𝑦 ∈ 𝐴 ¬ 𝑥 < 𝑦 ∧ ∀𝑦 ∈ (𝐵[,]𝐶)(𝑦 < 𝑥 → ∃𝑧 ∈ 𝐴 𝑦 < 𝑧))) | ||
| Theorem | dedekindicclemuub 14862* | Lemma for dedekindicc 14869. Any element of the upper cut is an upper bound for the lower cut. (Contributed by Jim Kingdon, 15-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐶 ∈ 𝑈) ⇒ ⊢ (𝜑 → ∀𝑧 ∈ 𝐿 𝑧 < 𝐶) | ||
| Theorem | dedekindicclemub 14863* | Lemma for dedekindicc 14869. The lower cut has an upper bound. (Contributed by Jim Kingdon, 15-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ (𝐴[,]𝐵)∀𝑦 ∈ 𝐿 𝑦 < 𝑥) | ||
| Theorem | dedekindicclemloc 14864* | Lemma for dedekindicc 14869. The set L is located. (Contributed by Jim Kingdon, 15-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) ⇒ ⊢ (𝜑 → ∀𝑥 ∈ (𝐴[,]𝐵)∀𝑦 ∈ (𝐴[,]𝐵)(𝑥 < 𝑦 → (∃𝑧 ∈ 𝐿 𝑥 < 𝑧 ∨ ∀𝑧 ∈ 𝐿 𝑧 < 𝑦))) | ||
| Theorem | dedekindicclemlub 14865* | Lemma for dedekindicc 14869. The set L has a least upper bound. (Contributed by Jim Kingdon, 15-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 < 𝐵) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ (𝐴[,]𝐵)(∀𝑦 ∈ 𝐿 ¬ 𝑥 < 𝑦 ∧ ∀𝑦 ∈ (𝐴[,]𝐵)(𝑦 < 𝑥 → ∃𝑧 ∈ 𝐿 𝑦 < 𝑧))) | ||
| Theorem | dedekindicclemlu 14866* | Lemma for dedekindicc 14869. There is a number which separates the lower and upper cuts. (Contributed by Jim Kingdon, 15-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 < 𝐵) ⇒ ⊢ (𝜑 → ∃𝑥 ∈ (𝐴[,]𝐵)(∀𝑞 ∈ 𝐿 𝑞 < 𝑥 ∧ ∀𝑟 ∈ 𝑈 𝑥 < 𝑟)) | ||
| Theorem | dedekindicclemeu 14867* | Lemma for dedekindicc 14869. Part of proving uniqueness. (Contributed by Jim Kingdon, 15-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → 𝐶 ∈ (𝐴[,]𝐵)) & ⊢ (𝜑 → (∀𝑞 ∈ 𝐿 𝑞 < 𝐶 ∧ ∀𝑟 ∈ 𝑈 𝐶 < 𝑟)) & ⊢ (𝜑 → 𝐷 ∈ (𝐴[,]𝐵)) & ⊢ (𝜑 → (∀𝑞 ∈ 𝐿 𝑞 < 𝐷 ∧ ∀𝑟 ∈ 𝑈 𝐷 < 𝑟)) & ⊢ (𝜑 → 𝐶 < 𝐷) ⇒ ⊢ (𝜑 → ⊥) | ||
| Theorem | dedekindicclemicc 14868* | Lemma for dedekindicc 14869. Same as dedekindicc 14869, except that we merely show 𝑥 to be an element of (𝐴[,]𝐵). Later we will strengthen that to (𝐴(,)𝐵). (Contributed by Jim Kingdon, 5-Jan-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 < 𝐵) ⇒ ⊢ (𝜑 → ∃!𝑥 ∈ (𝐴[,]𝐵)(∀𝑞 ∈ 𝐿 𝑞 < 𝑥 ∧ ∀𝑟 ∈ 𝑈 𝑥 < 𝑟)) | ||
| Theorem | dedekindicc 14869* | A Dedekind cut identifies a unique real number. Similar to df-inp 7533 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.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝐿 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → 𝑈 ⊆ (𝐴[,]𝐵)) & ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) & ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑈) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) & ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑈 ↔ ∃𝑞 ∈ 𝑈 𝑞 < 𝑟)) & ⊢ (𝜑 → (𝐿 ∩ 𝑈) = ∅) & ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑈))) & ⊢ (𝜑 → 𝐴 < 𝐵) ⇒ ⊢ (𝜑 → ∃!𝑥 ∈ (𝐴(,)𝐵)(∀𝑞 ∈ 𝐿 𝑞 < 𝑥 ∧ ∀𝑟 ∈ 𝑈 𝑥 < 𝑟)) | ||
| Theorem | ivthinclemlm 14870* | Lemma for ivthinc 14879. The lower cut is bounded. (Contributed by Jim Kingdon, 18-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∃𝑞 ∈ (𝐴[,]𝐵)𝑞 ∈ 𝐿) | ||
| Theorem | ivthinclemum 14871* | Lemma for ivthinc 14879. The upper cut is bounded. (Contributed by Jim Kingdon, 18-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∃𝑟 ∈ (𝐴[,]𝐵)𝑟 ∈ 𝑅) | ||
| Theorem | ivthinclemlopn 14872* | Lemma for ivthinc 14879. The lower cut is open. (Contributed by Jim Kingdon, 6-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} & ⊢ (𝜑 → 𝑄 ∈ 𝐿) ⇒ ⊢ (𝜑 → ∃𝑟 ∈ 𝐿 𝑄 < 𝑟) | ||
| Theorem | ivthinclemlr 14873* | Lemma for ivthinc 14879. The lower cut is rounded. (Contributed by Jim Kingdon, 18-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)(𝑞 ∈ 𝐿 ↔ ∃𝑟 ∈ 𝐿 𝑞 < 𝑟)) | ||
| Theorem | ivthinclemuopn 14874* | Lemma for ivthinc 14879. The upper cut is open. (Contributed by Jim Kingdon, 19-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} & ⊢ (𝜑 → 𝑆 ∈ 𝑅) ⇒ ⊢ (𝜑 → ∃𝑞 ∈ 𝑅 𝑞 < 𝑆) | ||
| Theorem | ivthinclemur 14875* | Lemma for ivthinc 14879. The upper cut is rounded. (Contributed by Jim Kingdon, 18-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∀𝑟 ∈ (𝐴[,]𝐵)(𝑟 ∈ 𝑅 ↔ ∃𝑞 ∈ 𝑅 𝑞 < 𝑟)) | ||
| Theorem | ivthinclemdisj 14876* | Lemma for ivthinc 14879. The lower and upper cuts are disjoint. (Contributed by Jim Kingdon, 18-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → (𝐿 ∩ 𝑅) = ∅) | ||
| Theorem | ivthinclemloc 14877* | Lemma for ivthinc 14879. Locatedness. (Contributed by Jim Kingdon, 18-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∀𝑞 ∈ (𝐴[,]𝐵)∀𝑟 ∈ (𝐴[,]𝐵)(𝑞 < 𝑟 → (𝑞 ∈ 𝐿 ∨ 𝑟 ∈ 𝑅))) | ||
| Theorem | ivthinclemex 14878* | Lemma for ivthinc 14879. Existence of a number between the lower cut and the upper cut. (Contributed by Jim Kingdon, 20-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) & ⊢ 𝐿 = {𝑤 ∈ (𝐴[,]𝐵) ∣ (𝐹‘𝑤) < 𝑈} & ⊢ 𝑅 = {𝑤 ∈ (𝐴[,]𝐵) ∣ 𝑈 < (𝐹‘𝑤)} ⇒ ⊢ (𝜑 → ∃!𝑧 ∈ (𝐴(,)𝐵)(∀𝑞 ∈ 𝐿 𝑞 < 𝑧 ∧ ∀𝑟 ∈ 𝑅 𝑧 < 𝑟)) | ||
| Theorem | ivthinc 14879* | 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→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑥) < (𝐹‘𝑦)) ⇒ ⊢ (𝜑 → ∃𝑐 ∈ (𝐴(,)𝐵)(𝐹‘𝑐) = 𝑈) | ||
| Theorem | ivthdec 14880* | The intermediate value theorem, decreasing case, for a strictly monotonic function. (Contributed by Jim Kingdon, 20-Feb-2024.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → (𝐴[,]𝐵) ⊆ 𝐷) & ⊢ (𝜑 → 𝐹 ∈ (𝐷–cn→ℂ)) & ⊢ ((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) → (𝐹‘𝑥) ∈ ℝ) & ⊢ (𝜑 → ((𝐹‘𝐵) < 𝑈 ∧ 𝑈 < (𝐹‘𝐴))) & ⊢ (((𝜑 ∧ 𝑥 ∈ (𝐴[,]𝐵)) ∧ (𝑦 ∈ (𝐴[,]𝐵) ∧ 𝑥 < 𝑦)) → (𝐹‘𝑦) < (𝐹‘𝑥)) ⇒ ⊢ (𝜑 → ∃𝑐 ∈ (𝐴(,)𝐵)(𝐹‘𝑐) = 𝑈) | ||
| Theorem | ivthreinc 14881* | Restating the intermediate value theorem. Given a hypothesis stating the intermediate value theorem (in a strong form which is not provable given our axioms alone), provide a conclusion similar to the theorem as stated in the Metamath Proof Explorer (which is also similar to how we state the theorem for a strictly monotonic function at ivthinc 14879). Being able to have a hypothesis stating the intermediate value theorem will be helpful when it comes time to show that it implies a constructive taboo. This version of the theorem requires that the function 𝐹 is continuous on the entire real line, not just (𝐴[,]𝐵) which may be an unnecessary condition but which is sufficient for the way we want to use it. (Contributed by Jim Kingdon, 7-Jul-2025.) | 
| ⊢ (𝜑 → 𝐴 ∈ ℝ) & ⊢ (𝜑 → 𝐵 ∈ ℝ) & ⊢ (𝜑 → 𝑈 ∈ ℝ) & ⊢ (𝜑 → 𝐴 < 𝐵) & ⊢ (𝜑 → 𝐹 ∈ (ℝ–cn→ℝ)) & ⊢ (𝜑 → ((𝐹‘𝐴) < 𝑈 ∧ 𝑈 < (𝐹‘𝐵))) & ⊢ (𝜑 → ∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0)))) ⇒ ⊢ (𝜑 → ∃𝑐 ∈ (𝐴(,)𝐵)(𝐹‘𝑐) = 𝑈) | ||
| Theorem | hovercncf 14882 | The hover function is continuous. By hover function, we mean a a function which starts out as a line of slope one, is constant at zero from zero to one, and then resumes as a slope of one. (Contributed by Jim Kingdon, 20-Jul-2025.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < )) ⇒ ⊢ 𝐹 ∈ (ℝ–cn→ℝ) | ||
| Theorem | hovera 14883* | A point at which the hover function is less than a given value. (Contributed by Jim Kingdon, 21-Jul-2025.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < )) ⇒ ⊢ (𝑍 ∈ ℝ → (𝐹‘(𝑍 − 1)) < 𝑍) | ||
| Theorem | hoverb 14884* | A point at which the hover function is greater than a given value. (Contributed by Jim Kingdon, 21-Jul-2025.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < )) ⇒ ⊢ (𝑍 ∈ ℝ → 𝑍 < (𝐹‘(𝑍 + 2))) | ||
| Theorem | hoverlt1 14885* | The hover function evaluated at a point less than one. (Contributed by Jim Kingdon, 22-Jul-2025.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < )) ⇒ ⊢ ((𝐶 ∈ ℝ ∧ 𝐶 < 1) → (𝐹‘𝐶) ≤ 0) | ||
| Theorem | hovergt0 14886* | The hover function evaluated at a point greater than zero. (Contributed by Jim Kingdon, 22-Jul-2025.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < )) ⇒ ⊢ ((𝐶 ∈ ℝ ∧ 0 < 𝐶) → 0 ≤ (𝐹‘𝐶)) | ||
| Theorem | ivthdichlem 14887* | Lemma for ivthdich 14889. The result, with a few notational conveniences. (Contributed by Jim Kingdon, 22-Jul-2025.) | 
| ⊢ 𝐹 = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < )) & ⊢ (𝜑 → 𝑍 ∈ ℝ) & ⊢ (𝜑 → ∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0)))) ⇒ ⊢ (𝜑 → (𝑍 ≤ 0 ∨ 0 ≤ 𝑍)) | ||
| Theorem | dich0 14888* | Real number dichotomy stated in terms of two real numbers or a real number and zero. (Contributed by Jim Kingdon, 22-Jul-2025.) | 
| ⊢ (∀𝑧 ∈ ℝ (𝑧 ≤ 0 ∨ 0 ≤ 𝑧) ↔ ∀𝑥 ∈ ℝ ∀𝑦 ∈ ℝ (𝑥 ≤ 𝑦 ∨ 𝑦 ≤ 𝑥)) | ||
| Theorem | ivthdich 14889* | 
The intermediate value theorem implies real number dichotomy.  Because
       real number dichotomy (also known as analytic LLPO) is a constructive
       taboo, this means we will be unable to prove the intermediate value
       theorem as stated here (although versions with additional conditions,
       such as ivthinc 14879 for strictly monotonic functions, can be
proved).
 The proof is via a function which we call the hover function and which is also described in Section 5.1 of [Bauer], p. 493. Consider any real number 𝑧. We want to show that 𝑧 ≤ 0 ∨ 0 ≤ 𝑧. Because of hovercncf 14882, hovera 14883, and hoverb 14884, we are able to apply the intermediate value theorem to get a value 𝑐 such that the hover function at 𝑐 equals 𝑧. By axltwlin 8094, 𝑐 < 1 or 0 < 𝑐, and that leads to 𝑧 ≤ 0 by hoverlt1 14885 or 0 ≤ 𝑧 by hovergt0 14886. (Contributed by Jim Kingdon and Mario Carneiro, 22-Jul-2025.)  | 
| ⊢ (∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0))) → ∀𝑟 ∈ ℝ ∀𝑠 ∈ ℝ (𝑟 ≤ 𝑠 ∨ 𝑠 ≤ 𝑟)) | ||
| Syntax | climc 14890 | The limit operator. | 
| class limℂ | ||
| Syntax | cdv 14891 | The derivative operator. | 
| class D | ||
| Definition | df-limced 14892* | Define the set of limits of a complex function at a point. Under normal circumstances, this will be a singleton or empty, depending on whether the limit exists. (Contributed by Mario Carneiro, 24-Dec-2016.) (Revised by Jim Kingdon, 3-Jun-2023.) | 
| ⊢ limℂ = (𝑓 ∈ (ℂ ↑pm ℂ), 𝑥 ∈ ℂ ↦ {𝑦 ∈ ℂ ∣ ((𝑓:dom 𝑓⟶ℂ ∧ dom 𝑓 ⊆ ℂ) ∧ (𝑥 ∈ ℂ ∧ ∀𝑒 ∈ ℝ+ ∃𝑑 ∈ ℝ+ ∀𝑧 ∈ dom 𝑓((𝑧 # 𝑥 ∧ (abs‘(𝑧 − 𝑥)) < 𝑑) → (abs‘((𝑓‘𝑧) − 𝑦)) < 𝑒)))}) | ||
| Definition | df-dvap 14893* | Define the derivative operator. This acts on functions to produce a function that is defined where the original function is differentiable, with value the derivative of the function at these points. The set 𝑠 here is the ambient topological space under which we are evaluating the continuity of the difference quotient. Although the definition is valid for any subset of ℂ and is well-behaved when 𝑠 contains no isolated points, we will restrict our attention to the cases 𝑠 = ℝ or 𝑠 = ℂ for the majority of the development, these corresponding respectively to real and complex differentiation. (Contributed by Mario Carneiro, 7-Aug-2014.) (Revised by Jim Kingdon, 25-Jun-2023.) | 
| ⊢ D = (𝑠 ∈ 𝒫 ℂ, 𝑓 ∈ (ℂ ↑pm 𝑠) ↦ ∪ 𝑥 ∈ ((int‘((MetOpen‘(abs ∘ − )) ↾t 𝑠))‘dom 𝑓)({𝑥} × ((𝑧 ∈ {𝑤 ∈ dom 𝑓 ∣ 𝑤 # 𝑥} ↦ (((𝑓‘𝑧) − (𝑓‘𝑥)) / (𝑧 − 𝑥))) limℂ 𝑥))) | ||
| Theorem | limcrcl 14894 | Reverse closure for the limit operator. (Contributed by Mario Carneiro, 28-Dec-2016.) | 
| ⊢ (𝐶 ∈ (𝐹 limℂ 𝐵) → (𝐹:dom 𝐹⟶ℂ ∧ dom 𝐹 ⊆ ℂ ∧ 𝐵 ∈ ℂ)) | ||
| Theorem | limccl 14895 | Closure of the limit operator. (Contributed by Mario Carneiro, 25-Dec-2016.) | 
| ⊢ (𝐹 limℂ 𝐵) ⊆ ℂ | ||
| Theorem | ellimc3apf 14896* | 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‘((𝐹‘𝑧) − 𝐶)) < 𝑥)))) | ||
| Theorem | ellimc3ap 14897* | Write the epsilon-delta definition of a limit. (Contributed by Mario Carneiro, 28-Dec-2016.) Use apartness. (Revised by Jim Kingdon, 3-Jun-2023.) | 
| ⊢ (𝜑 → 𝐹:𝐴⟶ℂ) & ⊢ (𝜑 → 𝐴 ⊆ ℂ) & ⊢ (𝜑 → 𝐵 ∈ ℂ) ⇒ ⊢ (𝜑 → (𝐶 ∈ (𝐹 limℂ 𝐵) ↔ (𝐶 ∈ ℂ ∧ ∀𝑥 ∈ ℝ+ ∃𝑦 ∈ ℝ+ ∀𝑧 ∈ 𝐴 ((𝑧 # 𝐵 ∧ (abs‘(𝑧 − 𝐵)) < 𝑦) → (abs‘((𝐹‘𝑧) − 𝐶)) < 𝑥)))) | ||
| Theorem | limcdifap 14898* | It suffices to consider functions which are not defined at 𝐵 to define the limit of a function. In particular, the value of the original function 𝐹 at 𝐵 does not affect the limit of 𝐹. (Contributed by Mario Carneiro, 25-Dec-2016.) (Revised by Jim Kingdon, 3-Jun-2023.) | 
| ⊢ (𝜑 → 𝐹:𝐴⟶ℂ) & ⊢ (𝜑 → 𝐴 ⊆ ℂ) ⇒ ⊢ (𝜑 → (𝐹 limℂ 𝐵) = ((𝐹 ↾ {𝑥 ∈ 𝐴 ∣ 𝑥 # 𝐵}) limℂ 𝐵)) | ||
| Theorem | limcmpted 14899* | Express the limit operator for a function defined by a mapping, via epsilon-delta. (Contributed by Jim Kingdon, 3-Nov-2023.) | 
| ⊢ (𝜑 → 𝐴 ⊆ ℂ) & ⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ ((𝜑 ∧ 𝑧 ∈ 𝐴) → 𝐷 ∈ ℂ) ⇒ ⊢ (𝜑 → (𝐶 ∈ ((𝑧 ∈ 𝐴 ↦ 𝐷) limℂ 𝐵) ↔ (𝐶 ∈ ℂ ∧ ∀𝑥 ∈ ℝ+ ∃𝑦 ∈ ℝ+ ∀𝑧 ∈ 𝐴 ((𝑧 # 𝐵 ∧ (abs‘(𝑧 − 𝐵)) < 𝑦) → (abs‘(𝐷 − 𝐶)) < 𝑥)))) | ||
| Theorem | limcimolemlt 14900* | Lemma for limcimo 14901. (Contributed by Jim Kingdon, 3-Jul-2023.) | 
| ⊢ (𝜑 → 𝐹:𝐴⟶ℂ) & ⊢ (𝜑 → 𝐴 ⊆ ℂ) & ⊢ (𝜑 → 𝐵 ∈ ℂ) & ⊢ (𝜑 → 𝐵 ∈ 𝐶) & ⊢ (𝜑 → 𝐵 ∈ 𝑆) & ⊢ (𝜑 → 𝐶 ∈ (𝐾 ↾t 𝑆)) & ⊢ (𝜑 → 𝑆 ∈ {ℝ, ℂ}) & ⊢ (𝜑 → {𝑞 ∈ 𝐶 ∣ 𝑞 # 𝐵} ⊆ 𝐴) & ⊢ 𝐾 = (MetOpen‘(abs ∘ − )) & ⊢ (𝜑 → 𝐷 ∈ ℝ+) & ⊢ (𝜑 → 𝑋 ∈ (𝐹 limℂ 𝐵)) & ⊢ (𝜑 → 𝑌 ∈ (𝐹 limℂ 𝐵)) & ⊢ (𝜑 → ∀𝑧 ∈ 𝐴 ((𝑧 # 𝐵 ∧ (abs‘(𝑧 − 𝐵)) < 𝐷) → (abs‘((𝐹‘𝑧) − 𝑋)) < ((abs‘(𝑋 − 𝑌)) / 2))) & ⊢ (𝜑 → 𝐺 ∈ ℝ+) & ⊢ (𝜑 → ∀𝑤 ∈ 𝐴 ((𝑤 # 𝐵 ∧ (abs‘(𝑤 − 𝐵)) < 𝐺) → (abs‘((𝐹‘𝑤) − 𝑌)) < ((abs‘(𝑋 − 𝑌)) / 2))) ⇒ ⊢ (𝜑 → (abs‘(𝑋 − 𝑌)) < (abs‘(𝑋 − 𝑌))) | ||
| < Previous Next > | 
| Copyright terms: Public domain | < Previous Next > |