| Intuitionistic Logic Explorer | 
      
      
      < Previous  
      Next >
      
       Nearby theorems  | 
  ||
| Mirrors > Home > ILE Home > Th. List > fzoshftral | GIF version | ||
| Description: Shift the scanning order inside of a quantification over a half-open integer range, analogous to fzshftral 10183. (Contributed by Alexander van der Vekens, 23-Sep-2018.) | 
| Ref | Expression | 
|---|---|
| fzoshftral | ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → (∀𝑗 ∈ (𝑀..^𝑁)𝜑 ↔ ∀𝑘 ∈ ((𝑀 + 𝐾)..^(𝑁 + 𝐾))[(𝑘 − 𝐾) / 𝑗]𝜑)) | 
| Step | Hyp | Ref | Expression | 
|---|---|---|---|
| 1 | fzoval 10223 | . . . 4 ⊢ (𝑁 ∈ ℤ → (𝑀..^𝑁) = (𝑀...(𝑁 − 1))) | |
| 2 | 1 | 3ad2ant2 1021 | . . 3 ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → (𝑀..^𝑁) = (𝑀...(𝑁 − 1))) | 
| 3 | 2 | raleqdv 2699 | . 2 ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → (∀𝑗 ∈ (𝑀..^𝑁)𝜑 ↔ ∀𝑗 ∈ (𝑀...(𝑁 − 1))𝜑)) | 
| 4 | peano2zm 9364 | . . 3 ⊢ (𝑁 ∈ ℤ → (𝑁 − 1) ∈ ℤ) | |
| 5 | fzshftral 10183 | . . 3 ⊢ ((𝑀 ∈ ℤ ∧ (𝑁 − 1) ∈ ℤ ∧ 𝐾 ∈ ℤ) → (∀𝑗 ∈ (𝑀...(𝑁 − 1))𝜑 ↔ ∀𝑘 ∈ ((𝑀 + 𝐾)...((𝑁 − 1) + 𝐾))[(𝑘 − 𝐾) / 𝑗]𝜑)) | |
| 6 | 4, 5 | syl3an2 1283 | . 2 ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → (∀𝑗 ∈ (𝑀...(𝑁 − 1))𝜑 ↔ ∀𝑘 ∈ ((𝑀 + 𝐾)...((𝑁 − 1) + 𝐾))[(𝑘 − 𝐾) / 𝑗]𝜑)) | 
| 7 | zaddcl 9366 | . . . . . 6 ⊢ ((𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → (𝑁 + 𝐾) ∈ ℤ) | |
| 8 | 7 | 3adant1 1017 | . . . . 5 ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → (𝑁 + 𝐾) ∈ ℤ) | 
| 9 | fzoval 10223 | . . . . 5 ⊢ ((𝑁 + 𝐾) ∈ ℤ → ((𝑀 + 𝐾)..^(𝑁 + 𝐾)) = ((𝑀 + 𝐾)...((𝑁 + 𝐾) − 1))) | |
| 10 | 8, 9 | syl 14 | . . . 4 ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → ((𝑀 + 𝐾)..^(𝑁 + 𝐾)) = ((𝑀 + 𝐾)...((𝑁 + 𝐾) − 1))) | 
| 11 | zcn 9331 | . . . . . . . 8 ⊢ (𝑁 ∈ ℤ → 𝑁 ∈ ℂ) | |
| 12 | 11 | adantr 276 | . . . . . . 7 ⊢ ((𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → 𝑁 ∈ ℂ) | 
| 13 | zcn 9331 | . . . . . . . 8 ⊢ (𝐾 ∈ ℤ → 𝐾 ∈ ℂ) | |
| 14 | 13 | adantl 277 | . . . . . . 7 ⊢ ((𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → 𝐾 ∈ ℂ) | 
| 15 | 1cnd 8042 | . . . . . . 7 ⊢ ((𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → 1 ∈ ℂ) | |
| 16 | 12, 14, 15 | addsubd 8358 | . . . . . 6 ⊢ ((𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → ((𝑁 + 𝐾) − 1) = ((𝑁 − 1) + 𝐾)) | 
| 17 | 16 | 3adant1 1017 | . . . . 5 ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → ((𝑁 + 𝐾) − 1) = ((𝑁 − 1) + 𝐾)) | 
| 18 | 17 | oveq2d 5938 | . . . 4 ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → ((𝑀 + 𝐾)...((𝑁 + 𝐾) − 1)) = ((𝑀 + 𝐾)...((𝑁 − 1) + 𝐾))) | 
| 19 | 10, 18 | eqtr2d 2230 | . . 3 ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → ((𝑀 + 𝐾)...((𝑁 − 1) + 𝐾)) = ((𝑀 + 𝐾)..^(𝑁 + 𝐾))) | 
| 20 | 19 | raleqdv 2699 | . 2 ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → (∀𝑘 ∈ ((𝑀 + 𝐾)...((𝑁 − 1) + 𝐾))[(𝑘 − 𝐾) / 𝑗]𝜑 ↔ ∀𝑘 ∈ ((𝑀 + 𝐾)..^(𝑁 + 𝐾))[(𝑘 − 𝐾) / 𝑗]𝜑)) | 
| 21 | 3, 6, 20 | 3bitrd 214 | 1 ⊢ ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → (∀𝑗 ∈ (𝑀..^𝑁)𝜑 ↔ ∀𝑘 ∈ ((𝑀 + 𝐾)..^(𝑁 + 𝐾))[(𝑘 − 𝐾) / 𝑗]𝜑)) | 
| Colors of variables: wff set class | 
| Syntax hints: → wi 4 ∧ wa 104 ↔ wb 105 ∧ w3a 980 = wceq 1364 ∈ wcel 2167 ∀wral 2475 [wsbc 2989 (class class class)co 5922 ℂcc 7877 1c1 7880 + caddc 7882 − cmin 8197 ℤcz 9326 ...cfz 10083 ..^cfzo 10217 | 
| 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 1461 ax-7 1462 ax-gen 1463 ax-ie1 1507 ax-ie2 1508 ax-8 1518 ax-10 1519 ax-11 1520 ax-i12 1521 ax-bndl 1523 ax-4 1524 ax-17 1540 ax-i9 1544 ax-ial 1548 ax-i5r 1549 ax-13 2169 ax-14 2170 ax-ext 2178 ax-sep 4151 ax-pow 4207 ax-pr 4242 ax-un 4468 ax-setind 4573 ax-cnex 7970 ax-resscn 7971 ax-1cn 7972 ax-1re 7973 ax-icn 7974 ax-addcl 7975 ax-addrcl 7976 ax-mulcl 7977 ax-addcom 7979 ax-addass 7981 ax-distr 7983 ax-i2m1 7984 ax-0lt1 7985 ax-0id 7987 ax-rnegex 7988 ax-cnre 7990 ax-pre-ltirr 7991 ax-pre-ltwlin 7992 ax-pre-lttrn 7993 ax-pre-ltadd 7995 | 
| This theorem depends on definitions: df-bi 117 df-3or 981 df-3an 982 df-tru 1367 df-fal 1370 df-nf 1475 df-sb 1777 df-eu 2048 df-mo 2049 df-clab 2183 df-cleq 2189 df-clel 2192 df-nfc 2328 df-ne 2368 df-nel 2463 df-ral 2480 df-rex 2481 df-reu 2482 df-rab 2484 df-v 2765 df-sbc 2990 df-csb 3085 df-dif 3159 df-un 3161 df-in 3163 df-ss 3170 df-pw 3607 df-sn 3628 df-pr 3629 df-op 3631 df-uni 3840 df-int 3875 df-iun 3918 df-br 4034 df-opab 4095 df-mpt 4096 df-id 4328 df-xp 4669 df-rel 4670 df-cnv 4671 df-co 4672 df-dm 4673 df-rn 4674 df-res 4675 df-ima 4676 df-iota 5219 df-fun 5260 df-fn 5261 df-f 5262 df-fv 5266 df-riota 5877 df-ov 5925 df-oprab 5926 df-mpo 5927 df-1st 6198 df-2nd 6199 df-pnf 8063 df-mnf 8064 df-xr 8065 df-ltxr 8066 df-le 8067 df-sub 8199 df-neg 8200 df-inn 8991 df-n0 9250 df-z 9327 df-uz 9602 df-fz 10084 df-fzo 10218 | 
| This theorem is referenced by: (None) | 
| Copyright terms: Public domain | W3C validator |