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

Theorem divmulasscom 11365
Description: An associative/commutative law for division and multiplication. (Contributed by AV, 10-Jul-2021.)
Assertion
Ref Expression
divmulasscom (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → ((𝐴 · (𝐵 / 𝐷)) · 𝐶) = (𝐵 · ((𝐴 · 𝐶) / 𝐷)))

Proof of Theorem divmulasscom
StepHypRef Expression
1 divmulass 11364 . 2 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → ((𝐴 · (𝐵 / 𝐷)) · 𝐶) = ((𝐴 · 𝐵) · (𝐶 / 𝐷)))
2 mulcom 10666 . . . . 5 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ) → (𝐴 · 𝐵) = (𝐵 · 𝐴))
323adant3 1129 . . . 4 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) → (𝐴 · 𝐵) = (𝐵 · 𝐴))
43adantr 484 . . 3 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → (𝐴 · 𝐵) = (𝐵 · 𝐴))
54oveq1d 7170 . 2 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → ((𝐴 · 𝐵) · (𝐶 / 𝐷)) = ((𝐵 · 𝐴) · (𝐶 / 𝐷)))
6 simpl2 1189 . . . 4 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → 𝐵 ∈ ℂ)
7 simpl1 1188 . . . 4 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → 𝐴 ∈ ℂ)
8 simp3 1135 . . . . . . 7 ((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) → 𝐶 ∈ ℂ)
98anim1i 617 . . . . . 6 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → (𝐶 ∈ ℂ ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)))
10 3anass 1092 . . . . . 6 ((𝐶 ∈ ℂ ∧ 𝐷 ∈ ℂ ∧ 𝐷 ≠ 0) ↔ (𝐶 ∈ ℂ ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)))
119, 10sylibr 237 . . . . 5 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → (𝐶 ∈ ℂ ∧ 𝐷 ∈ ℂ ∧ 𝐷 ≠ 0))
12 divcl 11347 . . . . 5 ((𝐶 ∈ ℂ ∧ 𝐷 ∈ ℂ ∧ 𝐷 ≠ 0) → (𝐶 / 𝐷) ∈ ℂ)
1311, 12syl 17 . . . 4 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → (𝐶 / 𝐷) ∈ ℂ)
146, 7, 13mulassd 10707 . . 3 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → ((𝐵 · 𝐴) · (𝐶 / 𝐷)) = (𝐵 · (𝐴 · (𝐶 / 𝐷))))
158adantr 484 . . . . . 6 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → 𝐶 ∈ ℂ)
16 simpr 488 . . . . . 6 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0))
17 divass 11359 . . . . . 6 ((𝐴 ∈ ℂ ∧ 𝐶 ∈ ℂ ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → ((𝐴 · 𝐶) / 𝐷) = (𝐴 · (𝐶 / 𝐷)))
187, 15, 16, 17syl3anc 1368 . . . . 5 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → ((𝐴 · 𝐶) / 𝐷) = (𝐴 · (𝐶 / 𝐷)))
1918eqcomd 2764 . . . 4 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → (𝐴 · (𝐶 / 𝐷)) = ((𝐴 · 𝐶) / 𝐷))
2019oveq2d 7171 . . 3 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → (𝐵 · (𝐴 · (𝐶 / 𝐷))) = (𝐵 · ((𝐴 · 𝐶) / 𝐷)))
2114, 20eqtrd 2793 . 2 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → ((𝐵 · 𝐴) · (𝐶 / 𝐷)) = (𝐵 · ((𝐴 · 𝐶) / 𝐷)))
221, 5, 213eqtrd 2797 1 (((𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ∧ 𝐶 ∈ ℂ) ∧ (𝐷 ∈ ℂ ∧ 𝐷 ≠ 0)) → ((𝐴 · (𝐵 / 𝐷)) · 𝐶) = (𝐵 · ((𝐴 · 𝐶) / 𝐷)))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 399  w3a 1084   = wceq 1538  wcel 2111  wne 2951  (class class class)co 7155  cc 10578  0cc0 10580   · cmul 10585   / cdiv 11340
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 1911  ax-6 1970  ax-7 2015  ax-8 2113  ax-9 2121  ax-10 2142  ax-11 2158  ax-12 2175  ax-ext 2729  ax-sep 5172  ax-nul 5179  ax-pow 5237  ax-pr 5301  ax-un 7464  ax-resscn 10637  ax-1cn 10638  ax-icn 10639  ax-addcl 10640  ax-addrcl 10641  ax-mulcl 10642  ax-mulrcl 10643  ax-mulcom 10644  ax-addass 10645  ax-mulass 10646  ax-distr 10647  ax-i2m1 10648  ax-1ne0 10649  ax-1rid 10650  ax-rnegex 10651  ax-rrecex 10652  ax-cnre 10653  ax-pre-lttri 10654  ax-pre-lttrn 10655  ax-pre-ltadd 10656  ax-pre-mulgt0 10657
This theorem depends on definitions:  df-bi 210  df-an 400  df-or 845  df-3or 1085  df-3an 1086  df-tru 1541  df-fal 1551  df-ex 1782  df-nf 1786  df-sb 2070  df-mo 2557  df-eu 2588  df-clab 2736  df-cleq 2750  df-clel 2830  df-nfc 2901  df-ne 2952  df-nel 3056  df-ral 3075  df-rex 3076  df-reu 3077  df-rmo 3078  df-rab 3079  df-v 3411  df-sbc 3699  df-csb 3808  df-dif 3863  df-un 3865  df-in 3867  df-ss 3877  df-nul 4228  df-if 4424  df-pw 4499  df-sn 4526  df-pr 4528  df-op 4532  df-uni 4802  df-br 5036  df-opab 5098  df-mpt 5116  df-id 5433  df-po 5446  df-so 5447  df-xp 5533  df-rel 5534  df-cnv 5535  df-co 5536  df-dm 5537  df-rn 5538  df-res 5539  df-ima 5540  df-iota 6298  df-fun 6341  df-fn 6342  df-f 6343  df-f1 6344  df-fo 6345  df-f1o 6346  df-fv 6347  df-riota 7113  df-ov 7158  df-oprab 7159  df-mpo 7160  df-er 8304  df-en 8533  df-dom 8534  df-sdom 8535  df-pnf 10720  df-mnf 10721  df-xr 10722  df-ltxr 10723  df-le 10724  df-sub 10915  df-neg 10916  df-div 11341
This theorem is referenced by:  cncongr2  16069
  Copyright terms: Public domain W3C validator