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

Theorem cshnz 14831
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 14828 . . 3 cyclShift = (𝑤 ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)}, 𝑛 ∈ ℤ ↦ if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤))))))
2 0ex 5306 . . . 4 ∅ ∈ V
3 ovex 7465 . . . 4 ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤)))) ∈ V
42, 3ifex 4575 . . 3 if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤))))) ∈ V
51, 4dmmpo 8097 . 2 dom cyclShift = ({𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)} × ℤ)
6 id 22 . . 3 𝑁 ∈ ℤ → ¬ 𝑁 ∈ ℤ)
76intnand 488 . 2 𝑁 ∈ ℤ → ¬ (𝑊 ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)} ∧ 𝑁 ∈ ℤ))
8 ndmovg 7617 . 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 395   = wceq 1539  wcel 2107  {cab 2713  wrex 3069  c0 4332  ifcif 4524  cop 4631   × cxp 5682  dom cdm 5684   Fn wfn 6555  cfv 6560  (class class class)co 7432  0cc0 11156  0cn0 12528  cz 12615  ..^cfzo 13695   mod cmo 13910  chash 14370   ++ cconcat 14609   substr csubstr 14679   prefix cpfx 14709   cyclShift ccsh 14827
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1794  ax-4 1808  ax-5 1909  ax-6 1966  ax-7 2006  ax-8 2109  ax-9 2117  ax-10 2140  ax-11 2156  ax-12 2176  ax-ext 2707  ax-sep 5295  ax-nul 5305  ax-pr 5431  ax-un 7756
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 848  df-3an 1088  df-tru 1542  df-fal 1552  df-ex 1779  df-nf 1783  df-sb 2064  df-mo 2539  df-eu 2568  df-clab 2714  df-cleq 2728  df-clel 2815  df-nfc 2891  df-ne 2940  df-ral 3061  df-rex 3070  df-rab 3436  df-v 3481  df-sbc 3788  df-csb 3899  df-dif 3953  df-un 3955  df-in 3957  df-ss 3967  df-nul 4333  df-if 4525  df-sn 4626  df-pr 4628  df-op 4632  df-uni 4907  df-iun 4992  df-br 5143  df-opab 5205  df-mpt 5225  df-id 5577  df-xp 5690  df-rel 5691  df-cnv 5692  df-co 5693  df-dm 5694  df-rn 5695  df-res 5696  df-ima 5697  df-iota 6513  df-fun 6562  df-fn 6563  df-f 6564  df-fv 6568  df-ov 7435  df-oprab 7436  df-mpo 7437  df-1st 8015  df-2nd 8016  df-csh 14828
This theorem is referenced by:  0csh0  14832  cshwcl  14837
  Copyright terms: Public domain W3C validator