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

Theorem minmar1fval 21255
Description: First substitution for the definition of a matrix for a minor. (Contributed by AV, 31-Dec-2018.)
Hypotheses
Ref Expression
minmar1fval.a 𝐴 = (𝑁 Mat 𝑅)
minmar1fval.b 𝐵 = (Base‘𝐴)
minmar1fval.q 𝑄 = (𝑁 minMatR1 𝑅)
minmar1fval.o 1 = (1r𝑅)
minmar1fval.z 0 = (0g𝑅)
Assertion
Ref Expression
minmar1fval 𝑄 = (𝑚𝐵 ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗)))))
Distinct variable groups:   𝐵,𝑚   𝑖,𝑁,𝑗,𝑘,𝑙,𝑚   𝑅,𝑖,𝑗,𝑘,𝑙,𝑚
Allowed substitution hints:   𝐴(𝑖,𝑗,𝑘,𝑚,𝑙)   𝐵(𝑖,𝑗,𝑘,𝑙)   𝑄(𝑖,𝑗,𝑘,𝑚,𝑙)   1 (𝑖,𝑗,𝑘,𝑚,𝑙)   0 (𝑖,𝑗,𝑘,𝑚,𝑙)

Proof of Theorem minmar1fval
Dummy variables 𝑛 𝑟 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 minmar1fval.q . 2 𝑄 = (𝑁 minMatR1 𝑅)
2 oveq12 7165 . . . . . . . 8 ((𝑛 = 𝑁𝑟 = 𝑅) → (𝑛 Mat 𝑟) = (𝑁 Mat 𝑅))
3 minmar1fval.a . . . . . . . 8 𝐴 = (𝑁 Mat 𝑅)
42, 3syl6eqr 2874 . . . . . . 7 ((𝑛 = 𝑁𝑟 = 𝑅) → (𝑛 Mat 𝑟) = 𝐴)
54fveq2d 6674 . . . . . 6 ((𝑛 = 𝑁𝑟 = 𝑅) → (Base‘(𝑛 Mat 𝑟)) = (Base‘𝐴))
6 minmar1fval.b . . . . . 6 𝐵 = (Base‘𝐴)
75, 6syl6eqr 2874 . . . . 5 ((𝑛 = 𝑁𝑟 = 𝑅) → (Base‘(𝑛 Mat 𝑟)) = 𝐵)
8 simpl 485 . . . . . 6 ((𝑛 = 𝑁𝑟 = 𝑅) → 𝑛 = 𝑁)
9 fveq2 6670 . . . . . . . . . . 11 (𝑟 = 𝑅 → (1r𝑟) = (1r𝑅))
10 minmar1fval.o . . . . . . . . . . 11 1 = (1r𝑅)
119, 10syl6eqr 2874 . . . . . . . . . 10 (𝑟 = 𝑅 → (1r𝑟) = 1 )
12 fveq2 6670 . . . . . . . . . . 11 (𝑟 = 𝑅 → (0g𝑟) = (0g𝑅))
13 minmar1fval.z . . . . . . . . . . 11 0 = (0g𝑅)
1412, 13syl6eqr 2874 . . . . . . . . . 10 (𝑟 = 𝑅 → (0g𝑟) = 0 )
1511, 14ifeq12d 4487 . . . . . . . . 9 (𝑟 = 𝑅 → if(𝑗 = 𝑙, (1r𝑟), (0g𝑟)) = if(𝑗 = 𝑙, 1 , 0 ))
1615ifeq1d 4485 . . . . . . . 8 (𝑟 = 𝑅 → if(𝑖 = 𝑘, if(𝑗 = 𝑙, (1r𝑟), (0g𝑟)), (𝑖𝑚𝑗)) = if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗)))
1716adantl 484 . . . . . . 7 ((𝑛 = 𝑁𝑟 = 𝑅) → if(𝑖 = 𝑘, if(𝑗 = 𝑙, (1r𝑟), (0g𝑟)), (𝑖𝑚𝑗)) = if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗)))
188, 8, 17mpoeq123dv 7229 . . . . . 6 ((𝑛 = 𝑁𝑟 = 𝑅) → (𝑖𝑛, 𝑗𝑛 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, (1r𝑟), (0g𝑟)), (𝑖𝑚𝑗))) = (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗))))
198, 8, 18mpoeq123dv 7229 . . . . 5 ((𝑛 = 𝑁𝑟 = 𝑅) → (𝑘𝑛, 𝑙𝑛 ↦ (𝑖𝑛, 𝑗𝑛 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, (1r𝑟), (0g𝑟)), (𝑖𝑚𝑗)))) = (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗)))))
207, 19mpteq12dv 5151 . . . 4 ((𝑛 = 𝑁𝑟 = 𝑅) → (𝑚 ∈ (Base‘(𝑛 Mat 𝑟)) ↦ (𝑘𝑛, 𝑙𝑛 ↦ (𝑖𝑛, 𝑗𝑛 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, (1r𝑟), (0g𝑟)), (𝑖𝑚𝑗))))) = (𝑚𝐵 ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗))))))
21 df-minmar1 21244 . . . 4 minMatR1 = (𝑛 ∈ V, 𝑟 ∈ V ↦ (𝑚 ∈ (Base‘(𝑛 Mat 𝑟)) ↦ (𝑘𝑛, 𝑙𝑛 ↦ (𝑖𝑛, 𝑗𝑛 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, (1r𝑟), (0g𝑟)), (𝑖𝑚𝑗))))))
226fvexi 6684 . . . . 5 𝐵 ∈ V
2322mptex 6986 . . . 4 (𝑚𝐵 ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗))))) ∈ V
2420, 21, 23ovmpoa 7305 . . 3 ((𝑁 ∈ V ∧ 𝑅 ∈ V) → (𝑁 minMatR1 𝑅) = (𝑚𝐵 ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗))))))
2521mpondm0 7386 . . . . 5 (¬ (𝑁 ∈ V ∧ 𝑅 ∈ V) → (𝑁 minMatR1 𝑅) = ∅)
26 mpt0 6490 . . . . 5 (𝑚 ∈ ∅ ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗))))) = ∅
2725, 26syl6eqr 2874 . . . 4 (¬ (𝑁 ∈ V ∧ 𝑅 ∈ V) → (𝑁 minMatR1 𝑅) = (𝑚 ∈ ∅ ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗))))))
283fveq2i 6673 . . . . . . 7 (Base‘𝐴) = (Base‘(𝑁 Mat 𝑅))
296, 28eqtri 2844 . . . . . 6 𝐵 = (Base‘(𝑁 Mat 𝑅))
30 matbas0pc 21018 . . . . . 6 (¬ (𝑁 ∈ V ∧ 𝑅 ∈ V) → (Base‘(𝑁 Mat 𝑅)) = ∅)
3129, 30syl5eq 2868 . . . . 5 (¬ (𝑁 ∈ V ∧ 𝑅 ∈ V) → 𝐵 = ∅)
3231mpteq1d 5155 . . . 4 (¬ (𝑁 ∈ V ∧ 𝑅 ∈ V) → (𝑚𝐵 ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗))))) = (𝑚 ∈ ∅ ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗))))))
3327, 32eqtr4d 2859 . . 3 (¬ (𝑁 ∈ V ∧ 𝑅 ∈ V) → (𝑁 minMatR1 𝑅) = (𝑚𝐵 ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗))))))
3424, 33pm2.61i 184 . 2 (𝑁 minMatR1 𝑅) = (𝑚𝐵 ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗)))))
351, 34eqtri 2844 1 𝑄 = (𝑚𝐵 ↦ (𝑘𝑁, 𝑙𝑁 ↦ (𝑖𝑁, 𝑗𝑁 ↦ if(𝑖 = 𝑘, if(𝑗 = 𝑙, 1 , 0 ), (𝑖𝑚𝑗)))))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wa 398   = wceq 1537  wcel 2114  Vcvv 3494  c0 4291  ifcif 4467  cmpt 5146  cfv 6355  (class class class)co 7156  cmpo 7158  Basecbs 16483  0gc0g 16713  1rcur 19251   Mat cmat 21016   minMatR1 cminmar1 21242
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 2793  ax-rep 5190  ax-sep 5203  ax-nul 5210  ax-pow 5266  ax-pr 5330
This theorem depends on definitions:  df-bi 209  df-an 399  df-or 844  df-3an 1085  df-tru 1540  df-ex 1781  df-nf 1785  df-sb 2070  df-mo 2622  df-eu 2654  df-clab 2800  df-cleq 2814  df-clel 2893  df-nfc 2963  df-ne 3017  df-ral 3143  df-rex 3144  df-reu 3145  df-rab 3147  df-v 3496  df-sbc 3773  df-csb 3884  df-dif 3939  df-un 3941  df-in 3943  df-ss 3952  df-nul 4292  df-if 4468  df-sn 4568  df-pr 4570  df-op 4574  df-uni 4839  df-iun 4921  df-br 5067  df-opab 5129  df-mpt 5147  df-id 5460  df-xp 5561  df-rel 5562  df-cnv 5563  df-co 5564  df-dm 5565  df-rn 5566  df-res 5567  df-ima 5568  df-iota 6314  df-fun 6357  df-fn 6358  df-f 6359  df-f1 6360  df-fo 6361  df-f1o 6362  df-fv 6363  df-ov 7159  df-oprab 7160  df-mpo 7161  df-slot 16487  df-base 16489  df-mat 21017  df-minmar1 21244
This theorem is referenced by:  minmar1val0  21256
  Copyright terms: Public domain W3C validator