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

Theorem odzval 15334
Description: Value of the order function. This is a function of functions; the inner argument selects the base (i.e. mod 𝑁 for some 𝑁, often prime) and the outer argument selects the integer or equivalence class (if you want to think about it that way) from the integers mod 𝑁. In order to ensure the supremum is well-defined, we only define the expression when 𝐴 and 𝑁 are coprime. (Contributed by Mario Carneiro, 23-Feb-2014.) (Revised by AV, 26-Sep-2020.)
Assertion
Ref Expression
odzval ((𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ (𝐴 gcd 𝑁) = 1) → ((od𝑁)‘𝐴) = inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝐴𝑛) − 1)}, ℝ, < ))
Distinct variable groups:   𝑛,𝑁   𝐴,𝑛

Proof of Theorem odzval
Dummy variables 𝑚 𝑥 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 oveq2 6557 . . . . . . . . 9 (𝑚 = 𝑁 → (𝑥 gcd 𝑚) = (𝑥 gcd 𝑁))
21eqeq1d 2612 . . . . . . . 8 (𝑚 = 𝑁 → ((𝑥 gcd 𝑚) = 1 ↔ (𝑥 gcd 𝑁) = 1))
32rabbidv 3164 . . . . . . 7 (𝑚 = 𝑁 → {𝑥 ∈ ℤ ∣ (𝑥 gcd 𝑚) = 1} = {𝑥 ∈ ℤ ∣ (𝑥 gcd 𝑁) = 1})
4 oveq1 6556 . . . . . . . . 9 (𝑛 = 𝑥 → (𝑛 gcd 𝑁) = (𝑥 gcd 𝑁))
54eqeq1d 2612 . . . . . . . 8 (𝑛 = 𝑥 → ((𝑛 gcd 𝑁) = 1 ↔ (𝑥 gcd 𝑁) = 1))
65cbvrabv 3172 . . . . . . 7 {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} = {𝑥 ∈ ℤ ∣ (𝑥 gcd 𝑁) = 1}
73, 6syl6eqr 2662 . . . . . 6 (𝑚 = 𝑁 → {𝑥 ∈ ℤ ∣ (𝑥 gcd 𝑚) = 1} = {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1})
8 breq1 4586 . . . . . . . 8 (𝑚 = 𝑁 → (𝑚 ∥ ((𝑥𝑛) − 1) ↔ 𝑁 ∥ ((𝑥𝑛) − 1)))
98rabbidv 3164 . . . . . . 7 (𝑚 = 𝑁 → {𝑛 ∈ ℕ ∣ 𝑚 ∥ ((𝑥𝑛) − 1)} = {𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)})
109infeq1d 8266 . . . . . 6 (𝑚 = 𝑁 → inf({𝑛 ∈ ℕ ∣ 𝑚 ∥ ((𝑥𝑛) − 1)}, ℝ, < ) = inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)}, ℝ, < ))
117, 10mpteq12dv 4663 . . . . 5 (𝑚 = 𝑁 → (𝑥 ∈ {𝑥 ∈ ℤ ∣ (𝑥 gcd 𝑚) = 1} ↦ inf({𝑛 ∈ ℕ ∣ 𝑚 ∥ ((𝑥𝑛) − 1)}, ℝ, < )) = (𝑥 ∈ {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} ↦ inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)}, ℝ, < )))
12 df-odz 15308 . . . . 5 od = (𝑚 ∈ ℕ ↦ (𝑥 ∈ {𝑥 ∈ ℤ ∣ (𝑥 gcd 𝑚) = 1} ↦ inf({𝑛 ∈ ℕ ∣ 𝑚 ∥ ((𝑥𝑛) − 1)}, ℝ, < )))
13 zex 11263 . . . . . 6 ℤ ∈ V
1413mptrabex 6392 . . . . 5 (𝑥 ∈ {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} ↦ inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)}, ℝ, < )) ∈ V
1511, 12, 14fvmpt 6191 . . . 4 (𝑁 ∈ ℕ → (od𝑁) = (𝑥 ∈ {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} ↦ inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)}, ℝ, < )))
1615fveq1d 6105 . . 3 (𝑁 ∈ ℕ → ((od𝑁)‘𝐴) = ((𝑥 ∈ {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} ↦ inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)}, ℝ, < ))‘𝐴))
17 oveq1 6556 . . . . . 6 (𝑛 = 𝐴 → (𝑛 gcd 𝑁) = (𝐴 gcd 𝑁))
1817eqeq1d 2612 . . . . 5 (𝑛 = 𝐴 → ((𝑛 gcd 𝑁) = 1 ↔ (𝐴 gcd 𝑁) = 1))
1918elrab 3331 . . . 4 (𝐴 ∈ {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} ↔ (𝐴 ∈ ℤ ∧ (𝐴 gcd 𝑁) = 1))
20 oveq1 6556 . . . . . . . . 9 (𝑥 = 𝐴 → (𝑥𝑛) = (𝐴𝑛))
2120oveq1d 6564 . . . . . . . 8 (𝑥 = 𝐴 → ((𝑥𝑛) − 1) = ((𝐴𝑛) − 1))
2221breq2d 4595 . . . . . . 7 (𝑥 = 𝐴 → (𝑁 ∥ ((𝑥𝑛) − 1) ↔ 𝑁 ∥ ((𝐴𝑛) − 1)))
2322rabbidv 3164 . . . . . 6 (𝑥 = 𝐴 → {𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)} = {𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝐴𝑛) − 1)})
2423infeq1d 8266 . . . . 5 (𝑥 = 𝐴 → inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)}, ℝ, < ) = inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝐴𝑛) − 1)}, ℝ, < ))
25 eqid 2610 . . . . 5 (𝑥 ∈ {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} ↦ inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)}, ℝ, < )) = (𝑥 ∈ {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} ↦ inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)}, ℝ, < ))
26 ltso 9997 . . . . . 6 < Or ℝ
2726infex 8282 . . . . 5 inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝐴𝑛) − 1)}, ℝ, < ) ∈ V
2824, 25, 27fvmpt 6191 . . . 4 (𝐴 ∈ {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} → ((𝑥 ∈ {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} ↦ inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)}, ℝ, < ))‘𝐴) = inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝐴𝑛) − 1)}, ℝ, < ))
2919, 28sylbir 224 . . 3 ((𝐴 ∈ ℤ ∧ (𝐴 gcd 𝑁) = 1) → ((𝑥 ∈ {𝑛 ∈ ℤ ∣ (𝑛 gcd 𝑁) = 1} ↦ inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝑥𝑛) − 1)}, ℝ, < ))‘𝐴) = inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝐴𝑛) − 1)}, ℝ, < ))
3016, 29sylan9eq 2664 . 2 ((𝑁 ∈ ℕ ∧ (𝐴 ∈ ℤ ∧ (𝐴 gcd 𝑁) = 1)) → ((od𝑁)‘𝐴) = inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝐴𝑛) − 1)}, ℝ, < ))
31303impb 1252 1 ((𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ (𝐴 gcd 𝑁) = 1) → ((od𝑁)‘𝐴) = inf({𝑛 ∈ ℕ ∣ 𝑁 ∥ ((𝐴𝑛) − 1)}, ℝ, < ))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 383  w3a 1031   = wceq 1475  wcel 1977  {crab 2900   class class class wbr 4583  cmpt 4643  cfv 5804  (class class class)co 6549  infcinf 8230  cr 9814  1c1 9816   < clt 9953  cmin 10145  cn 10897  cz 11254  cexp 12722  cdvds 14821   gcd cgcd 15054  odcodz 15306
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1713  ax-4 1728  ax-5 1827  ax-6 1875  ax-7 1922  ax-8 1979  ax-9 1986  ax-10 2006  ax-11 2021  ax-12 2034  ax-13 2234  ax-ext 2590  ax-rep 4699  ax-sep 4709  ax-nul 4717  ax-pow 4769  ax-pr 4833  ax-un 6847  ax-cnex 9871  ax-resscn 9872  ax-pre-lttri 9889  ax-pre-lttrn 9890
This theorem depends on definitions:  df-bi 196  df-or 384  df-an 385  df-3or 1032  df-3an 1033  df-tru 1478  df-ex 1696  df-nf 1701  df-sb 1868  df-eu 2462  df-mo 2463  df-clab 2597  df-cleq 2603  df-clel 2606  df-nfc 2740  df-ne 2782  df-nel 2783  df-ral 2901  df-rex 2902  df-reu 2903  df-rmo 2904  df-rab 2905  df-v 3175  df-sbc 3403  df-csb 3500  df-dif 3543  df-un 3545  df-in 3547  df-ss 3554  df-nul 3875  df-if 4037  df-pw 4110  df-sn 4126  df-pr 4128  df-op 4132  df-uni 4373  df-iun 4457  df-br 4584  df-opab 4644  df-mpt 4645  df-id 4953  df-po 4959  df-so 4960  df-xp 5044  df-rel 5045  df-cnv 5046  df-co 5047  df-dm 5048  df-rn 5049  df-res 5050  df-ima 5051  df-iota 5768  df-fun 5806  df-fn 5807  df-f 5808  df-f1 5809  df-fo 5810  df-f1o 5811  df-fv 5812  df-ov 6552  df-er 7629  df-en 7842  df-dom 7843  df-sdom 7844  df-sup 8231  df-inf 8232  df-pnf 9955  df-mnf 9956  df-ltxr 9958  df-neg 10148  df-z 11255  df-odz 15308
This theorem is referenced by:  odzcllem  15335  odzdvds  15338
  Copyright terms: Public domain W3C validator