![]() |
Intuitionistic Logic Explorer |
< Previous
Next >
Nearby theorems |
|
Mirrors > Home > ILE Home > Th. List > ivthdich | GIF version |
Description: 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 14797 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 14800, hovera 14801, and hoverb 14802, we are able to apply the intermediate value theorem to get a value 𝑐 such that the hover function at 𝑐 equals 𝑧. By axltwlin 8087, 𝑐 < 1 or 0 < 𝑐, and that leads to 𝑧 ≤ 0 by hoverlt1 14803 or 0 ≤ 𝑧 by hovergt0 14804. (Contributed by Jim Kingdon and Mario Carneiro, 22-Jul-2025.) |
Ref | Expression |
---|---|
ivthdich | ⊢ (∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0))) → ∀𝑟 ∈ ℝ ∀𝑠 ∈ ℝ (𝑟 ≤ 𝑠 ∨ 𝑠 ≤ 𝑟)) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | breq2 4033 | . . . . . . . . . 10 ⊢ (𝑥 = 𝑞 → (𝑎 < 𝑥 ↔ 𝑎 < 𝑞)) | |
2 | breq1 4032 | . . . . . . . . . 10 ⊢ (𝑥 = 𝑞 → (𝑥 < 𝑏 ↔ 𝑞 < 𝑏)) | |
3 | fveqeq2 5563 | . . . . . . . . . 10 ⊢ (𝑥 = 𝑞 → ((𝑓‘𝑥) = 0 ↔ (𝑓‘𝑞) = 0)) | |
4 | 1, 2, 3 | 3anbi123d 1323 | . . . . . . . . 9 ⊢ (𝑥 = 𝑞 → ((𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0) ↔ (𝑎 < 𝑞 ∧ 𝑞 < 𝑏 ∧ (𝑓‘𝑞) = 0))) |
5 | 4 | cbvrexv 2727 | . . . . . . . 8 ⊢ (∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0) ↔ ∃𝑞 ∈ ℝ (𝑎 < 𝑞 ∧ 𝑞 < 𝑏 ∧ (𝑓‘𝑞) = 0)) |
6 | 5 | imbi2i 226 | . . . . . . 7 ⊢ (((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0)) ↔ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑞 ∈ ℝ (𝑎 < 𝑞 ∧ 𝑞 < 𝑏 ∧ (𝑓‘𝑞) = 0))) |
7 | 6 | 2ralbii 2502 | . . . . . 6 ⊢ (∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0)) ↔ ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑞 ∈ ℝ (𝑎 < 𝑞 ∧ 𝑞 < 𝑏 ∧ (𝑓‘𝑞) = 0))) |
8 | 7 | imbi2i 226 | . . . . 5 ⊢ ((𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0))) ↔ (𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑞 ∈ ℝ (𝑎 < 𝑞 ∧ 𝑞 < 𝑏 ∧ (𝑓‘𝑞) = 0)))) |
9 | 8 | albii 1481 | . . . 4 ⊢ (∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0))) ↔ ∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑞 ∈ ℝ (𝑎 < 𝑞 ∧ 𝑞 < 𝑏 ∧ (𝑓‘𝑞) = 0)))) |
10 | preq1 3695 | . . . . . . . . 9 ⊢ (𝑡 = 𝑥 → {𝑡, 0} = {𝑥, 0}) | |
11 | 10 | infeq1d 7071 | . . . . . . . 8 ⊢ (𝑡 = 𝑥 → inf({𝑡, 0}, ℝ, < ) = inf({𝑥, 0}, ℝ, < )) |
12 | oveq1 5925 | . . . . . . . 8 ⊢ (𝑡 = 𝑥 → (𝑡 − 1) = (𝑥 − 1)) | |
13 | 11, 12 | preq12d 3703 | . . . . . . 7 ⊢ (𝑡 = 𝑥 → {inf({𝑡, 0}, ℝ, < ), (𝑡 − 1)} = {inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}) |
14 | 13 | supeq1d 7046 | . . . . . 6 ⊢ (𝑡 = 𝑥 → sup({inf({𝑡, 0}, ℝ, < ), (𝑡 − 1)}, ℝ, < ) = sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < )) |
15 | 14 | cbvmptv 4125 | . . . . 5 ⊢ (𝑡 ∈ ℝ ↦ sup({inf({𝑡, 0}, ℝ, < ), (𝑡 − 1)}, ℝ, < )) = (𝑥 ∈ ℝ ↦ sup({inf({𝑥, 0}, ℝ, < ), (𝑥 − 1)}, ℝ, < )) |
16 | simpr 110 | . . . . 5 ⊢ ((∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑞 ∈ ℝ (𝑎 < 𝑞 ∧ 𝑞 < 𝑏 ∧ (𝑓‘𝑞) = 0))) ∧ 𝑧 ∈ ℝ) → 𝑧 ∈ ℝ) | |
17 | 9 | biimpri 133 | . . . . . 6 ⊢ (∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑞 ∈ ℝ (𝑎 < 𝑞 ∧ 𝑞 < 𝑏 ∧ (𝑓‘𝑞) = 0))) → ∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0)))) |
18 | 17 | adantr 276 | . . . . 5 ⊢ ((∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑞 ∈ ℝ (𝑎 < 𝑞 ∧ 𝑞 < 𝑏 ∧ (𝑓‘𝑞) = 0))) ∧ 𝑧 ∈ ℝ) → ∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0)))) |
19 | 15, 16, 18 | ivthdichlem 14805 | . . . 4 ⊢ ((∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑞 ∈ ℝ (𝑎 < 𝑞 ∧ 𝑞 < 𝑏 ∧ (𝑓‘𝑞) = 0))) ∧ 𝑧 ∈ ℝ) → (𝑧 ≤ 0 ∨ 0 ≤ 𝑧)) |
20 | 9, 19 | sylanb 284 | . . 3 ⊢ ((∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0))) ∧ 𝑧 ∈ ℝ) → (𝑧 ≤ 0 ∨ 0 ≤ 𝑧)) |
21 | 20 | ralrimiva 2567 | . 2 ⊢ (∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0))) → ∀𝑧 ∈ ℝ (𝑧 ≤ 0 ∨ 0 ≤ 𝑧)) |
22 | dich0 14806 | . 2 ⊢ (∀𝑧 ∈ ℝ (𝑧 ≤ 0 ∨ 0 ≤ 𝑧) ↔ ∀𝑟 ∈ ℝ ∀𝑠 ∈ ℝ (𝑟 ≤ 𝑠 ∨ 𝑠 ≤ 𝑟)) | |
23 | 21, 22 | sylib 122 | 1 ⊢ (∀𝑓(𝑓 ∈ (ℝ–cn→ℝ) → ∀𝑎 ∈ ℝ ∀𝑏 ∈ ℝ ((𝑎 < 𝑏 ∧ (𝑓‘𝑎) < 0 ∧ 0 < (𝑓‘𝑏)) → ∃𝑥 ∈ ℝ (𝑎 < 𝑥 ∧ 𝑥 < 𝑏 ∧ (𝑓‘𝑥) = 0))) → ∀𝑟 ∈ ℝ ∀𝑠 ∈ ℝ (𝑟 ≤ 𝑠 ∨ 𝑠 ≤ 𝑟)) |
Colors of variables: wff set class |
Syntax hints: → wi 4 ∧ wa 104 ∨ wo 709 ∧ w3a 980 ∀wal 1362 = wceq 1364 ∈ wcel 2164 ∀wral 2472 ∃wrex 2473 {cpr 3619 class class class wbr 4029 ↦ cmpt 4090 ‘cfv 5254 (class class class)co 5918 supcsup 7041 infcinf 7042 ℝcr 7871 0cc0 7872 1c1 7873 < clt 8054 ≤ cle 8055 − cmin 8190 –cn→ccncf 14725 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-ia1 106 ax-ia2 107 ax-ia3 108 ax-in1 615 ax-in2 616 ax-io 710 ax-5 1458 ax-7 1459 ax-gen 1460 ax-ie1 1504 ax-ie2 1505 ax-8 1515 ax-10 1516 ax-11 1517 ax-i12 1518 ax-bndl 1520 ax-4 1521 ax-17 1537 ax-i9 1541 ax-ial 1545 ax-i5r 1546 ax-13 2166 ax-14 2167 ax-ext 2175 ax-coll 4144 ax-sep 4147 ax-nul 4155 ax-pow 4203 ax-pr 4238 ax-un 4464 ax-setind 4569 ax-iinf 4620 ax-cnex 7963 ax-resscn 7964 ax-1cn 7965 ax-1re 7966 ax-icn 7967 ax-addcl 7968 ax-addrcl 7969 ax-mulcl 7970 ax-mulrcl 7971 ax-addcom 7972 ax-mulcom 7973 ax-addass 7974 ax-mulass 7975 ax-distr 7976 ax-i2m1 7977 ax-0lt1 7978 ax-1rid 7979 ax-0id 7980 ax-rnegex 7981 ax-precex 7982 ax-cnre 7983 ax-pre-ltirr 7984 ax-pre-ltwlin 7985 ax-pre-lttrn 7986 ax-pre-apti 7987 ax-pre-ltadd 7988 ax-pre-mulgt0 7989 ax-pre-mulext 7990 ax-arch 7991 ax-caucvg 7992 ax-addf 7994 |
This theorem depends on definitions: df-bi 117 df-stab 832 df-dc 836 df-3or 981 df-3an 982 df-tru 1367 df-fal 1370 df-nf 1472 df-sb 1774 df-eu 2045 df-mo 2046 df-clab 2180 df-cleq 2186 df-clel 2189 df-nfc 2325 df-ne 2365 df-nel 2460 df-ral 2477 df-rex 2478 df-reu 2479 df-rmo 2480 df-rab 2481 df-v 2762 df-sbc 2986 df-csb 3081 df-dif 3155 df-un 3157 df-in 3159 df-ss 3166 df-nul 3447 df-if 3558 df-pw 3603 df-sn 3624 df-pr 3625 df-op 3627 df-uni 3836 df-int 3871 df-iun 3914 df-br 4030 df-opab 4091 df-mpt 4092 df-tr 4128 df-id 4324 df-po 4327 df-iso 4328 df-iord 4397 df-on 4399 df-ilim 4400 df-suc 4402 df-iom 4623 df-xp 4665 df-rel 4666 df-cnv 4667 df-co 4668 df-dm 4669 df-rn 4670 df-res 4671 df-ima 4672 df-iota 5215 df-fun 5256 df-fn 5257 df-f 5258 df-f1 5259 df-fo 5260 df-f1o 5261 df-fv 5262 df-isom 5263 df-riota 5873 df-ov 5921 df-oprab 5922 df-mpo 5923 df-1st 6193 df-2nd 6194 df-recs 6358 df-frec 6444 df-map 6704 df-sup 7043 df-inf 7044 df-pnf 8056 df-mnf 8057 df-xr 8058 df-ltxr 8059 df-le 8060 df-sub 8192 df-neg 8193 df-reap 8594 df-ap 8601 df-div 8692 df-inn 8983 df-2 9041 df-3 9042 df-4 9043 df-n0 9241 df-z 9318 df-uz 9593 df-q 9685 df-rp 9720 df-xneg 9838 df-xadd 9839 df-ioo 9958 df-seqfrec 10519 df-exp 10610 df-cj 10986 df-re 10987 df-im 10988 df-rsqrt 11142 df-abs 11143 df-rest 12852 df-topgen 12871 df-psmet 14039 df-xmet 14040 df-met 14041 df-bl 14042 df-mopn 14043 df-top 14166 df-topon 14179 df-bases 14211 df-cn 14356 df-cnp 14357 df-tx 14421 df-cncf 14726 |
This theorem is referenced by: (None) |
Copyright terms: Public domain | W3C validator |