| Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > MPE Home > Th. List > eucrct2eupth1 | Structured version Visualization version GIF version | ||
| Description: Removing one edge (𝐼‘(𝐹‘𝑁)) from a nonempty graph 𝐺 with an Eulerian circuit 〈𝐹, 𝑃〉 results in a graph 𝑆 with an Eulerian path 〈𝐻, 𝑄〉. This is the special case of eucrct2eupth 30181 (with 𝐽 = (𝑁 − 1)) where the last segment/edge of the circuit is removed. (Contributed by AV, 11-Mar-2021.) Hypothesis revised using the prefix operation. (Revised by AV, 30-Nov-2022.) |
| Ref | Expression |
|---|---|
| eucrct2eupth1.v | ⊢ 𝑉 = (Vtx‘𝐺) |
| eucrct2eupth1.i | ⊢ 𝐼 = (iEdg‘𝐺) |
| eucrct2eupth1.d | ⊢ (𝜑 → 𝐹(EulerPaths‘𝐺)𝑃) |
| eucrct2eupth1.c | ⊢ (𝜑 → 𝐹(Circuits‘𝐺)𝑃) |
| eucrct2eupth1.s | ⊢ (Vtx‘𝑆) = 𝑉 |
| eucrct2eupth1.g | ⊢ (𝜑 → 0 < (♯‘𝐹)) |
| eucrct2eupth1.n | ⊢ (𝜑 → 𝑁 = ((♯‘𝐹) − 1)) |
| eucrct2eupth1.e | ⊢ (𝜑 → (iEdg‘𝑆) = (𝐼 ↾ (𝐹 “ (0..^𝑁)))) |
| eucrct2eupth1.h | ⊢ 𝐻 = (𝐹 prefix 𝑁) |
| eucrct2eupth1.q | ⊢ 𝑄 = (𝑃 ↾ (0...𝑁)) |
| Ref | Expression |
|---|---|
| eucrct2eupth1 | ⊢ (𝜑 → 𝐻(EulerPaths‘𝑆)𝑄) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | eucrct2eupth1.v | . 2 ⊢ 𝑉 = (Vtx‘𝐺) | |
| 2 | eucrct2eupth1.i | . 2 ⊢ 𝐼 = (iEdg‘𝐺) | |
| 3 | eucrct2eupth1.d | . 2 ⊢ (𝜑 → 𝐹(EulerPaths‘𝐺)𝑃) | |
| 4 | eucrct2eupth1.n | . . 3 ⊢ (𝜑 → 𝑁 = ((♯‘𝐹) − 1)) | |
| 5 | eucrct2eupth1.g | . . . . 5 ⊢ (𝜑 → 0 < (♯‘𝐹)) | |
| 6 | eupthiswlk 30148 | . . . . . 6 ⊢ (𝐹(EulerPaths‘𝐺)𝑃 → 𝐹(Walks‘𝐺)𝑃) | |
| 7 | wlkcl 29550 | . . . . . 6 ⊢ (𝐹(Walks‘𝐺)𝑃 → (♯‘𝐹) ∈ ℕ0) | |
| 8 | nn0z 12570 | . . . . . . . . 9 ⊢ ((♯‘𝐹) ∈ ℕ0 → (♯‘𝐹) ∈ ℤ) | |
| 9 | 8 | anim1i 615 | . . . . . . . 8 ⊢ (((♯‘𝐹) ∈ ℕ0 ∧ 0 < (♯‘𝐹)) → ((♯‘𝐹) ∈ ℤ ∧ 0 < (♯‘𝐹))) |
| 10 | elnnz 12555 | . . . . . . . 8 ⊢ ((♯‘𝐹) ∈ ℕ ↔ ((♯‘𝐹) ∈ ℤ ∧ 0 < (♯‘𝐹))) | |
| 11 | 9, 10 | sylibr 234 | . . . . . . 7 ⊢ (((♯‘𝐹) ∈ ℕ0 ∧ 0 < (♯‘𝐹)) → (♯‘𝐹) ∈ ℕ) |
| 12 | 11 | ex 412 | . . . . . 6 ⊢ ((♯‘𝐹) ∈ ℕ0 → (0 < (♯‘𝐹) → (♯‘𝐹) ∈ ℕ)) |
| 13 | 3, 6, 7, 12 | 4syl 19 | . . . . 5 ⊢ (𝜑 → (0 < (♯‘𝐹) → (♯‘𝐹) ∈ ℕ)) |
| 14 | 5, 13 | mpd 15 | . . . 4 ⊢ (𝜑 → (♯‘𝐹) ∈ ℕ) |
| 15 | fzo0end 13731 | . . . 4 ⊢ ((♯‘𝐹) ∈ ℕ → ((♯‘𝐹) − 1) ∈ (0..^(♯‘𝐹))) | |
| 16 | 14, 15 | syl 17 | . . 3 ⊢ (𝜑 → ((♯‘𝐹) − 1) ∈ (0..^(♯‘𝐹))) |
| 17 | 4, 16 | eqeltrd 2829 | . 2 ⊢ (𝜑 → 𝑁 ∈ (0..^(♯‘𝐹))) |
| 18 | eucrct2eupth1.e | . 2 ⊢ (𝜑 → (iEdg‘𝑆) = (𝐼 ↾ (𝐹 “ (0..^𝑁)))) | |
| 19 | eucrct2eupth1.h | . 2 ⊢ 𝐻 = (𝐹 prefix 𝑁) | |
| 20 | eucrct2eupth1.q | . 2 ⊢ 𝑄 = (𝑃 ↾ (0...𝑁)) | |
| 21 | eucrct2eupth1.s | . 2 ⊢ (Vtx‘𝑆) = 𝑉 | |
| 22 | 1, 2, 3, 17, 18, 19, 20, 21 | eupthres 30151 | 1 ⊢ (𝜑 → 𝐻(EulerPaths‘𝑆)𝑄) |
| Colors of variables: wff setvar class |
| Syntax hints: → wi 4 ∧ wa 395 = wceq 1540 ∈ wcel 2109 class class class wbr 5115 ↾ cres 5648 “ cima 5649 ‘cfv 6519 (class class class)co 7394 0cc0 11086 1c1 11087 < clt 11226 − cmin 11423 ℕcn 12197 ℕ0cn0 12458 ℤcz 12545 ...cfz 13481 ..^cfzo 13628 ♯chash 14305 prefix cpfx 14645 Vtxcvtx 28930 iEdgciedg 28931 Walkscwlks 29531 Circuitsccrcts 29721 EulerPathsceupth 30133 |
| This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1795 ax-4 1809 ax-5 1910 ax-6 1967 ax-7 2008 ax-8 2111 ax-9 2119 ax-10 2142 ax-11 2158 ax-12 2178 ax-ext 2702 ax-rep 5242 ax-sep 5259 ax-nul 5269 ax-pow 5328 ax-pr 5395 ax-un 7718 ax-cnex 11142 ax-resscn 11143 ax-1cn 11144 ax-icn 11145 ax-addcl 11146 ax-addrcl 11147 ax-mulcl 11148 ax-mulrcl 11149 ax-mulcom 11150 ax-addass 11151 ax-mulass 11152 ax-distr 11153 ax-i2m1 11154 ax-1ne0 11155 ax-1rid 11156 ax-rnegex 11157 ax-rrecex 11158 ax-cnre 11159 ax-pre-lttri 11160 ax-pre-lttrn 11161 ax-pre-ltadd 11162 ax-pre-mulgt0 11163 |
| This theorem depends on definitions: df-bi 207 df-an 396 df-or 848 df-ifp 1063 df-3or 1087 df-3an 1088 df-tru 1543 df-fal 1553 df-ex 1780 df-nf 1784 df-sb 2066 df-mo 2534 df-eu 2563 df-clab 2709 df-cleq 2722 df-clel 2804 df-nfc 2880 df-ne 2928 df-nel 3032 df-ral 3047 df-rex 3056 df-reu 3358 df-rab 3412 df-v 3457 df-sbc 3762 df-csb 3871 df-dif 3925 df-un 3927 df-in 3929 df-ss 3939 df-pss 3942 df-nul 4305 df-if 4497 df-pw 4573 df-sn 4598 df-pr 4600 df-op 4604 df-uni 4880 df-int 4919 df-iun 4965 df-br 5116 df-opab 5178 df-mpt 5197 df-tr 5223 df-id 5541 df-eprel 5546 df-po 5554 df-so 5555 df-fr 5599 df-we 5601 df-xp 5652 df-rel 5653 df-cnv 5654 df-co 5655 df-dm 5656 df-rn 5657 df-res 5658 df-ima 5659 df-pred 6282 df-ord 6343 df-on 6344 df-lim 6345 df-suc 6346 df-iota 6472 df-fun 6521 df-fn 6522 df-f 6523 df-f1 6524 df-fo 6525 df-f1o 6526 df-fv 6527 df-riota 7351 df-ov 7397 df-oprab 7398 df-mpo 7399 df-om 7851 df-1st 7977 df-2nd 7978 df-frecs 8269 df-wrecs 8300 df-recs 8349 df-rdg 8387 df-1o 8443 df-er 8682 df-map 8805 df-pm 8806 df-en 8923 df-dom 8924 df-sdom 8925 df-fin 8926 df-card 9910 df-pnf 11228 df-mnf 11229 df-xr 11230 df-ltxr 11231 df-le 11232 df-sub 11425 df-neg 11426 df-nn 12198 df-n0 12459 df-z 12546 df-uz 12810 df-fz 13482 df-fzo 13629 df-hash 14306 df-word 14489 df-substr 14616 df-pfx 14646 df-wlks 29534 df-trls 29627 df-eupth 30134 |
| This theorem is referenced by: eucrct2eupth 30181 |
| Copyright terms: Public domain | W3C validator |