MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  cshnz Structured version   Visualization version   GIF version

Theorem cshnz 14142
Description: A cyclical shift is the empty set if the number of shifts is not an integer. (Contributed by Alexander van der Vekens, 21-May-2018.) (Revised by AV, 17-Nov-2018.)
Assertion
Ref Expression
cshnz 𝑁 ∈ ℤ → (𝑊 cyclShift 𝑁) = ∅)

Proof of Theorem cshnz
Dummy variables 𝑓 𝑙 𝑛 𝑤 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 df-csh 14139 . . 3 cyclShift = (𝑤 ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)}, 𝑛 ∈ ℤ ↦ if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤))))))
2 0ex 5202 . . . 4 ∅ ∈ V
3 ovex 7178 . . . 4 ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤)))) ∈ V
42, 3ifex 4511 . . 3 if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤))))) ∈ V
51, 4dmmpo 7758 . 2 dom cyclShift = ({𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)} × ℤ)
6 id 22 . . 3 𝑁 ∈ ℤ → ¬ 𝑁 ∈ ℤ)
76intnand 489 . 2 𝑁 ∈ ℤ → ¬ (𝑊 ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)} ∧ 𝑁 ∈ ℤ))
8 ndmovg 7320 . 2 ((dom cyclShift = ({𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)} × ℤ) ∧ ¬ (𝑊 ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)} ∧ 𝑁 ∈ ℤ)) → (𝑊 cyclShift 𝑁) = ∅)
95, 7, 8sylancr 587 1 𝑁 ∈ ℤ → (𝑊 cyclShift 𝑁) = ∅)
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 396   = wceq 1528  wcel 2105  {cab 2796  wrex 3136  c0 4288  ifcif 4463  cop 4563   × cxp 5546  dom cdm 5548   Fn wfn 6343  cfv 6348  (class class class)co 7145  0cc0 10525  0cn0 11885  cz 11969  ..^cfzo 13021   mod cmo 13225  chash 13678   ++ cconcat 13910   substr csubstr 13990   prefix cpfx 14020   cyclShift ccsh 14138
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1787  ax-4 1801  ax-5 1902  ax-6 1961  ax-7 2006  ax-8 2107  ax-9 2115  ax-10 2136  ax-11 2151  ax-12 2167  ax-ext 2790  ax-sep 5194  ax-nul 5201  ax-pow 5257  ax-pr 5320  ax-un 7450
This theorem depends on definitions:  df-bi 208  df-an 397  df-or 842  df-3an 1081  df-tru 1531  df-ex 1772  df-nf 1776  df-sb 2061  df-mo 2615  df-eu 2647  df-clab 2797  df-cleq 2811  df-clel 2890  df-nfc 2960  df-ne 3014  df-ral 3140  df-rex 3141  df-rab 3144  df-v 3494  df-sbc 3770  df-csb 3881  df-dif 3936  df-un 3938  df-in 3940  df-ss 3949  df-nul 4289  df-if 4464  df-sn 4558  df-pr 4560  df-op 4564  df-uni 4831  df-iun 4912  df-br 5058  df-opab 5120  df-mpt 5138  df-id 5453  df-xp 5554  df-rel 5555  df-cnv 5556  df-co 5557  df-dm 5558  df-rn 5559  df-res 5560  df-ima 5561  df-iota 6307  df-fun 6350  df-fn 6351  df-f 6352  df-fv 6356  df-ov 7148  df-oprab 7149  df-mpo 7150  df-1st 7678  df-2nd 7679  df-csh 14139
This theorem is referenced by:  0csh0  14143  cshwcl  14148
  Copyright terms: Public domain W3C validator