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

Theorem 2ffzeq 13561
Description: Two functions over 0-based finite set of sequential integers are equal if and only if their domains have the same length and the function values are the same at each position. (Contributed by Alexander van der Vekens, 30-Jun-2018.)
Assertion
Ref Expression
2ffzeq ((𝑀 ∈ ℕ0𝐹:(0...𝑀)⟶𝑋𝑃:(0...𝑁)⟶𝑌) → (𝐹 = 𝑃 ↔ (𝑀 = 𝑁 ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖))))
Distinct variable groups:   𝑖,𝐹   𝑖,𝑀   𝑃,𝑖
Allowed substitution hints:   𝑁(𝑖)   𝑋(𝑖)   𝑌(𝑖)

Proof of Theorem 2ffzeq
StepHypRef Expression
1 ffn 6668 . . . . 5 (𝐹:(0...𝑀)⟶𝑋𝐹 Fn (0...𝑀))
2 ffn 6668 . . . . 5 (𝑃:(0...𝑁)⟶𝑌𝑃 Fn (0...𝑁))
31, 2anim12i 613 . . . 4 ((𝐹:(0...𝑀)⟶𝑋𝑃:(0...𝑁)⟶𝑌) → (𝐹 Fn (0...𝑀) ∧ 𝑃 Fn (0...𝑁)))
433adant1 1130 . . 3 ((𝑀 ∈ ℕ0𝐹:(0...𝑀)⟶𝑋𝑃:(0...𝑁)⟶𝑌) → (𝐹 Fn (0...𝑀) ∧ 𝑃 Fn (0...𝑁)))
5 eqfnfv2 6983 . . 3 ((𝐹 Fn (0...𝑀) ∧ 𝑃 Fn (0...𝑁)) → (𝐹 = 𝑃 ↔ ((0...𝑀) = (0...𝑁) ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖))))
64, 5syl 17 . 2 ((𝑀 ∈ ℕ0𝐹:(0...𝑀)⟶𝑋𝑃:(0...𝑁)⟶𝑌) → (𝐹 = 𝑃 ↔ ((0...𝑀) = (0...𝑁) ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖))))
7 elnn0uz 12807 . . . . . . 7 (𝑀 ∈ ℕ0𝑀 ∈ (ℤ‘0))
8 fzopth 13477 . . . . . . 7 (𝑀 ∈ (ℤ‘0) → ((0...𝑀) = (0...𝑁) ↔ (0 = 0 ∧ 𝑀 = 𝑁)))
97, 8sylbi 216 . . . . . 6 (𝑀 ∈ ℕ0 → ((0...𝑀) = (0...𝑁) ↔ (0 = 0 ∧ 𝑀 = 𝑁)))
10 simpr 485 . . . . . 6 ((0 = 0 ∧ 𝑀 = 𝑁) → 𝑀 = 𝑁)
119, 10syl6bi 252 . . . . 5 (𝑀 ∈ ℕ0 → ((0...𝑀) = (0...𝑁) → 𝑀 = 𝑁))
1211anim1d 611 . . . 4 (𝑀 ∈ ℕ0 → (((0...𝑀) = (0...𝑁) ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖)) → (𝑀 = 𝑁 ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖))))
13 oveq2 7364 . . . . 5 (𝑀 = 𝑁 → (0...𝑀) = (0...𝑁))
1413anim1i 615 . . . 4 ((𝑀 = 𝑁 ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖)) → ((0...𝑀) = (0...𝑁) ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖)))
1512, 14impbid1 224 . . 3 (𝑀 ∈ ℕ0 → (((0...𝑀) = (0...𝑁) ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖)) ↔ (𝑀 = 𝑁 ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖))))
16153ad2ant1 1133 . 2 ((𝑀 ∈ ℕ0𝐹:(0...𝑀)⟶𝑋𝑃:(0...𝑁)⟶𝑌) → (((0...𝑀) = (0...𝑁) ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖)) ↔ (𝑀 = 𝑁 ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖))))
176, 16bitrd 278 1 ((𝑀 ∈ ℕ0𝐹:(0...𝑀)⟶𝑋𝑃:(0...𝑁)⟶𝑌) → (𝐹 = 𝑃 ↔ (𝑀 = 𝑁 ∧ ∀𝑖 ∈ (0...𝑀)(𝐹𝑖) = (𝑃𝑖))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 205  wa 396  w3a 1087   = wceq 1541  wcel 2106  wral 3064   Fn wfn 6491  wf 6492  cfv 6496  (class class class)co 7356  0cc0 11050  0cn0 12412  cuz 12762  ...cfz 13423
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1797  ax-4 1811  ax-5 1913  ax-6 1971  ax-7 2011  ax-8 2108  ax-9 2116  ax-10 2137  ax-11 2154  ax-12 2171  ax-ext 2707  ax-sep 5256  ax-nul 5263  ax-pow 5320  ax-pr 5384  ax-un 7671  ax-cnex 11106  ax-resscn 11107  ax-1cn 11108  ax-icn 11109  ax-addcl 11110  ax-addrcl 11111  ax-mulcl 11112  ax-mulrcl 11113  ax-mulcom 11114  ax-addass 11115  ax-mulass 11116  ax-distr 11117  ax-i2m1 11118  ax-1ne0 11119  ax-1rid 11120  ax-rnegex 11121  ax-rrecex 11122  ax-cnre 11123  ax-pre-lttri 11124  ax-pre-lttrn 11125  ax-pre-ltadd 11126  ax-pre-mulgt0 11127
This theorem depends on definitions:  df-bi 206  df-an 397  df-or 846  df-3or 1088  df-3an 1089  df-tru 1544  df-fal 1554  df-ex 1782  df-nf 1786  df-sb 2068  df-mo 2538  df-eu 2567  df-clab 2714  df-cleq 2728  df-clel 2814  df-nfc 2889  df-ne 2944  df-nel 3050  df-ral 3065  df-rex 3074  df-reu 3354  df-rab 3408  df-v 3447  df-sbc 3740  df-csb 3856  df-dif 3913  df-un 3915  df-in 3917  df-ss 3927  df-pss 3929  df-nul 4283  df-if 4487  df-pw 4562  df-sn 4587  df-pr 4589  df-op 4593  df-uni 4866  df-iun 4956  df-br 5106  df-opab 5168  df-mpt 5189  df-tr 5223  df-id 5531  df-eprel 5537  df-po 5545  df-so 5546  df-fr 5588  df-we 5590  df-xp 5639  df-rel 5640  df-cnv 5641  df-co 5642  df-dm 5643  df-rn 5644  df-res 5645  df-ima 5646  df-pred 6253  df-ord 6320  df-on 6321  df-lim 6322  df-suc 6323  df-iota 6448  df-fun 6498  df-fn 6499  df-f 6500  df-f1 6501  df-fo 6502  df-f1o 6503  df-fv 6504  df-riota 7312  df-ov 7359  df-oprab 7360  df-mpo 7361  df-om 7802  df-1st 7920  df-2nd 7921  df-frecs 8211  df-wrecs 8242  df-recs 8316  df-rdg 8355  df-er 8647  df-en 8883  df-dom 8884  df-sdom 8885  df-pnf 11190  df-mnf 11191  df-xr 11192  df-ltxr 11193  df-le 11194  df-sub 11386  df-neg 11387  df-nn 12153  df-n0 12413  df-z 12499  df-uz 12763  df-fz 13424
This theorem is referenced by:  wlkeq  28529
  Copyright terms: Public domain W3C validator