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

Theorem 0csh0 14776
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 14772 . . . 4 cyclShift = (𝑤 ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)}, 𝑛 ∈ ℤ ↦ if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤))))))
21a1i 11 . . 3 (𝑁 ∈ ℤ → cyclShift = (𝑤 ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)}, 𝑛 ∈ ℤ ↦ if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤)))))))
3 iftrue 4535 . . . 4 (𝑤 = ∅ → if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤))))) = ∅)
43ad2antrl 727 . . 3 ((𝑁 ∈ ℤ ∧ (𝑤 = ∅ ∧ 𝑛 = 𝑁)) → if(𝑤 = ∅, ∅, ((𝑤 substr ⟨(𝑛 mod (♯‘𝑤)), (♯‘𝑤)⟩) ++ (𝑤 prefix (𝑛 mod (♯‘𝑤))))) = ∅)
5 0nn0 12518 . . . . . 6 0 ∈ ℕ0
6 f0 6778 . . . . . . 7 ∅:∅⟶V
7 ffn 6722 . . . . . . . 8 (∅:∅⟶V → ∅ Fn ∅)
8 fzo0 13689 . . . . . . . . . 10 (0..^0) = ∅
98eqcomi 2737 . . . . . . . . 9 ∅ = (0..^0)
109fneq2i 6652 . . . . . . . 8 (∅ Fn ∅ ↔ ∅ Fn (0..^0))
117, 10sylib 217 . . . . . . 7 (∅:∅⟶V → ∅ Fn (0..^0))
126, 11ax-mp 5 . . . . . 6 ∅ Fn (0..^0)
13 id 22 . . . . . . 7 (0 ∈ ℕ0 → 0 ∈ ℕ0)
14 oveq2 7428 . . . . . . . . 9 (𝑙 = 0 → (0..^𝑙) = (0..^0))
1514fneq2d 6648 . . . . . . . 8 (𝑙 = 0 → (∅ Fn (0..^𝑙) ↔ ∅ Fn (0..^0)))
1615adantl 481 . . . . . . 7 ((0 ∈ ℕ0𝑙 = 0) → (∅ Fn (0..^𝑙) ↔ ∅ Fn (0..^0)))
1713, 16rspcedv 3602 . . . . . 6 (0 ∈ ℕ0 → (∅ Fn (0..^0) → ∃𝑙 ∈ ℕ0 ∅ Fn (0..^𝑙)))
185, 12, 17mp2 9 . . . . 5 𝑙 ∈ ℕ0 ∅ Fn (0..^𝑙)
19 0ex 5307 . . . . . 6 ∅ ∈ V
20 fneq1 6645 . . . . . . 7 (𝑓 = ∅ → (𝑓 Fn (0..^𝑙) ↔ ∅ Fn (0..^𝑙)))
2120rexbidv 3175 . . . . . 6 (𝑓 = ∅ → (∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙) ↔ ∃𝑙 ∈ ℕ0 ∅ Fn (0..^𝑙)))
2219, 21elab 3667 . . . . 5 (∅ ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)} ↔ ∃𝑙 ∈ ℕ0 ∅ Fn (0..^𝑙))
2318, 22mpbir 230 . . . 4 ∅ ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)}
2423a1i 11 . . 3 (𝑁 ∈ ℤ → ∅ ∈ {𝑓 ∣ ∃𝑙 ∈ ℕ0 𝑓 Fn (0..^𝑙)})
25 id 22 . . 3 (𝑁 ∈ ℤ → 𝑁 ∈ ℤ)
2619a1i 11 . . 3 (𝑁 ∈ ℤ → ∅ ∈ V)
272, 4, 24, 25, 26ovmpod 7573 . 2 (𝑁 ∈ ℤ → (∅ cyclShift 𝑁) = ∅)
28 cshnz 14775 . 2 𝑁 ∈ ℤ → (∅ cyclShift 𝑁) = ∅)
2927, 28pm2.61i 182 1 (∅ cyclShift 𝑁) = ∅
Colors of variables: wff setvar class
Syntax hints:  wb 205   = wceq 1534  wcel 2099  {cab 2705  wrex 3067  Vcvv 3471  c0 4323  ifcif 4529  cop 4635   Fn wfn 6543  wf 6544  cfv 6548  (class class class)co 7420  cmpo 7422  0cc0 11139  0cn0 12503  cz 12589  ..^cfzo 13660   mod cmo 13867  chash 14322   ++ cconcat 14553   substr csubstr 14623   prefix cpfx 14653   cyclShift ccsh 14771
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1790  ax-4 1804  ax-5 1906  ax-6 1964  ax-7 2004  ax-8 2101  ax-9 2109  ax-10 2130  ax-11 2147  ax-12 2167  ax-ext 2699  ax-sep 5299  ax-nul 5306  ax-pow 5365  ax-pr 5429  ax-un 7740  ax-cnex 11195  ax-resscn 11196  ax-1cn 11197  ax-icn 11198  ax-addcl 11199  ax-addrcl 11200  ax-mulcl 11201  ax-mulrcl 11202  ax-mulcom 11203  ax-addass 11204  ax-mulass 11205  ax-distr 11206  ax-i2m1 11207  ax-1ne0 11208  ax-1rid 11209  ax-rnegex 11210  ax-rrecex 11211  ax-cnre 11212  ax-pre-lttri 11213  ax-pre-lttrn 11214  ax-pre-ltadd 11215  ax-pre-mulgt0 11216
This theorem depends on definitions:  df-bi 206  df-an 396  df-or 847  df-3or 1086  df-3an 1087  df-tru 1537  df-fal 1547  df-ex 1775  df-nf 1779  df-sb 2061  df-mo 2530  df-eu 2559  df-clab 2706  df-cleq 2720  df-clel 2806  df-nfc 2881  df-ne 2938  df-nel 3044  df-ral 3059  df-rex 3068  df-reu 3374  df-rab 3430  df-v 3473  df-sbc 3777  df-csb 3893  df-dif 3950  df-un 3952  df-in 3954  df-ss 3964  df-pss 3966  df-nul 4324  df-if 4530  df-pw 4605  df-sn 4630  df-pr 4632  df-op 4636  df-uni 4909  df-iun 4998  df-br 5149  df-opab 5211  df-mpt 5232  df-tr 5266  df-id 5576  df-eprel 5582  df-po 5590  df-so 5591  df-fr 5633  df-we 5635  df-xp 5684  df-rel 5685  df-cnv 5686  df-co 5687  df-dm 5688  df-rn 5689  df-res 5690  df-ima 5691  df-pred 6305  df-ord 6372  df-on 6373  df-lim 6374  df-suc 6375  df-iota 6500  df-fun 6550  df-fn 6551  df-f 6552  df-f1 6553  df-fo 6554  df-f1o 6555  df-fv 6556  df-riota 7376  df-ov 7423  df-oprab 7424  df-mpo 7425  df-om 7871  df-1st 7993  df-2nd 7994  df-frecs 8287  df-wrecs 8318  df-recs 8392  df-rdg 8431  df-er 8725  df-en 8965  df-dom 8966  df-sdom 8967  df-pnf 11281  df-mnf 11282  df-xr 11283  df-ltxr 11284  df-le 11285  df-sub 11477  df-neg 11478  df-nn 12244  df-n0 12504  df-z 12590  df-uz 12854  df-fz 13518  df-fzo 13661  df-csh 14772
This theorem is referenced by:  cshw0  14777  cshwmodn  14778  cshwn  14780  cshwlen  14782  repswcshw  14795
  Copyright terms: Public domain W3C validator