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

Theorem 0csh0 14157
Description: Cyclically shifting an empty set/word always results in the empty word/set. (Contributed by AV, 25-Oct-2018.) (Revised by AV, 17-Nov-2018.)
Assertion
Ref Expression
0csh0 (∅ cyclShift 𝑁) = ∅

Proof of Theorem 0csh0
Dummy variables 𝑓 𝑙 𝑛 𝑤 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 df-csh 14153 . . . 4 cyclShift = (𝑤 ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)}, 𝑛 ∈ ℤ ↦ if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤))))))
21a1i 11 . . 3 (𝑁 ∈ ℤ → cyclShift = (𝑤 ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)}, 𝑛 ∈ ℤ ↦ if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤)))))))
3 iftrue 4475 . . . 4 (𝑤 = ∅ → if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤))))) = ∅)
43ad2antrl 726 . . 3 ((𝑁 ∈ ℤ ∧ (𝑤 = ∅ ∧ 𝑛 = 𝑁)) → if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤))))) = ∅)
5 0nn0 11915 . . . . . 6 0 ∈ ℕ0
6 f0 6562 . . . . . . 7 ∅:∅⟶V
7 ffn 6516 . . . . . . . 8 (∅:∅⟶V → ∅ Fn ∅)
8 fzo0 13064 . . . . . . . . . 10 (0..^0) = ∅
98eqcomi 2832 . . . . . . . . 9 ∅ = (0..^0)
109fneq2i 6453 . . . . . . . 8 (∅ Fn ∅ ↔ ∅ Fn (0..^0))
117, 10sylib 220 . . . . . . 7 (∅:∅⟶V → ∅ Fn (0..^0))
126, 11ax-mp 5 . . . . . 6 ∅ Fn (0..^0)
13 id 22 . . . . . . 7 (0 ∈ ℕ0 → 0 ∈ ℕ0)
14 oveq2 7166 . . . . . . . . 9 (𝑙 = 0 → (0..^𝑙) = (0..^0))
1514fneq2d 6449 . . . . . . . 8 (𝑙 = 0 → (∅ Fn (0..^𝑙) ↔ ∅ Fn (0..^0)))
1615adantl 484 . . . . . . 7 ((0 ∈ ℕ0𝑙 = 0) → (∅ Fn (0..^𝑙) ↔ ∅ Fn (0..^0)))
1713, 16rspcedv 3618 . . . . . 6 (0 ∈ ℕ0 → (∅ Fn (0..^0) → ∃𝑙 ∈ ℕ0 ∅ Fn (0..^𝑙)))
185, 12, 17mp2 9 . . . . 5 𝑙 ∈ ℕ0 ∅ Fn (0..^𝑙)
19 0ex 5213 . . . . . 6 ∅ ∈ V
20 fneq1 6446 . . . . . . 7 (𝑓 = ∅ → (𝑓 Fn (0..^𝑙) ↔ ∅ Fn (0..^𝑙)))
2120rexbidv 3299 . . . . . 6 (𝑓 = ∅ → (∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙) ↔ ∃𝑙 ∈ ℕ0 ∅ Fn (0..^𝑙)))
2219, 21elab 3669 . . . . 5 (∅ ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)} ↔ ∃𝑙 ∈ ℕ0 ∅ Fn (0..^𝑙))
2318, 22mpbir 233 . . . 4 ∅ ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)}
2423a1i 11 . . 3 (𝑁 ∈ ℤ → ∅ ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)})
25 id 22 . . 3 (𝑁 ∈ ℤ → 𝑁 ∈ ℤ)
2619a1i 11 . . 3 (𝑁 ∈ ℤ → ∅ ∈ V)
272, 4, 24, 25, 26ovmpod 7304 . 2 (𝑁 ∈ ℤ → (∅ cyclShift 𝑁) = ∅)
28 cshnz 14156 . 2 𝑁 ∈ ℤ → (∅ cyclShift 𝑁) = ∅)
2927, 28pm2.61i 184 1 (∅ cyclShift 𝑁) = ∅
Colors of variables: wff setvar class
Syntax hints:  wb 208   = wceq 1537  wcel 2114  {cab 2801  wrex 3141  Vcvv 3496  c0 4293  ifcif 4469  cop 4575   Fn wfn 6352  wf 6353  cfv 6357  (class class class)co 7158  cmpo 7160  0cc0 10539  0cn0 11900  cz 11984  ..^cfzo 13036   mod cmo 13240  chash 13693   ++ cconcat 13924   substr csubstr 14004   prefix cpfx 14034   cyclShift ccsh 14152
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1796  ax-4 1810  ax-5 1911  ax-6 1970  ax-7 2015  ax-8 2116  ax-9 2124  ax-10 2145  ax-11 2161  ax-12 2177  ax-ext 2795  ax-sep 5205  ax-nul 5212  ax-pow 5268  ax-pr 5332  ax-un 7463  ax-cnex 10595  ax-resscn 10596  ax-1cn 10597  ax-icn 10598  ax-addcl 10599  ax-addrcl 10600  ax-mulcl 10601  ax-mulrcl 10602  ax-mulcom 10603  ax-addass 10604  ax-mulass 10605  ax-distr 10606  ax-i2m1 10607  ax-1ne0 10608  ax-1rid 10609  ax-rnegex 10610  ax-rrecex 10611  ax-cnre 10612  ax-pre-lttri 10613  ax-pre-lttrn 10614  ax-pre-ltadd 10615  ax-pre-mulgt0 10616
This theorem depends on definitions:  df-bi 209  df-an 399  df-or 844  df-3or 1084  df-3an 1085  df-tru 1540  df-ex 1781  df-nf 1785  df-sb 2070  df-mo 2622  df-eu 2654  df-clab 2802  df-cleq 2816  df-clel 2895  df-nfc 2965  df-ne 3019  df-nel 3126  df-ral 3145  df-rex 3146  df-reu 3147  df-rab 3149  df-v 3498  df-sbc 3775  df-csb 3886  df-dif 3941  df-un 3943  df-in 3945  df-ss 3954  df-pss 3956  df-nul 4294  df-if 4470  df-pw 4543  df-sn 4570  df-pr 4572  df-tp 4574  df-op 4576  df-uni 4841  df-iun 4923  df-br 5069  df-opab 5131  df-mpt 5149  df-tr 5175  df-id 5462  df-eprel 5467  df-po 5476  df-so 5477  df-fr 5516  df-we 5518  df-xp 5563  df-rel 5564  df-cnv 5565  df-co 5566  df-dm 5567  df-rn 5568  df-res 5569  df-ima 5570  df-pred 6150  df-ord 6196  df-on 6197  df-lim 6198  df-suc 6199  df-iota 6316  df-fun 6359  df-fn 6360  df-f 6361  df-f1 6362  df-fo 6363  df-f1o 6364  df-fv 6365  df-riota 7116  df-ov 7161  df-oprab 7162  df-mpo 7163  df-om 7583  df-1st 7691  df-2nd 7692  df-wrecs 7949  df-recs 8010  df-rdg 8048  df-er 8291  df-en 8512  df-dom 8513  df-sdom 8514  df-pnf 10679  df-mnf 10680  df-xr 10681  df-ltxr 10682  df-le 10683  df-sub 10874  df-neg 10875  df-nn 11641  df-n0 11901  df-z 11985  df-uz 12247  df-fz 12896  df-fzo 13037  df-csh 14153
This theorem is referenced by:  cshw0  14158  cshwmodn  14159  cshwn  14161  cshwlen  14163  repswcshw  14176
  Copyright terms: Public domain W3C validator