@@ -3,6 +3,7 @@ Copyright (c) 2022 Oliver Nash. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Oliver Nash
5
5
-/
6
+ import data.nat.totient
6
7
import algebra.ring.add_aut
7
8
import group_theory.divisible
8
9
import group_theory.order_of_element
@@ -67,6 +68,10 @@ lemma coe_nsmul {n : ℕ} {x : 𝕜} : (↑(n • x) : add_circle p) = n • (x
67
68
68
69
lemma coe_zsmul {n : ℤ} {x : 𝕜} : (↑(n • x) : add_circle p) = n • (x : add_circle p) := rfl
69
70
71
+ lemma coe_add (x y : 𝕜) : (↑(x + y) : add_circle p) = (x : add_circle p) + (y : add_circle p) := rfl
72
+
73
+ lemma coe_sub (x y : 𝕜) : (↑(x - y) : add_circle p) = (x : add_circle p) - (y : add_circle p) := rfl
74
+
70
75
lemma coe_neg {x : 𝕜} : (↑(-x) : add_circle p) = -(x : add_circle p) := rfl
71
76
72
77
lemma coe_eq_zero_iff {x : 𝕜} : (x : add_circle p) = 0 ↔ ∃ (n : ℤ), n • p = x :=
@@ -280,6 +285,72 @@ begin
280
285
simpa [zero_add] using (equiv_Ico p 0 u).2 .2 , },
281
286
end
282
287
288
+ lemma add_order_of_eq_pos_iff {u : add_circle p} {n : ℕ} (h : 0 < n) :
289
+ add_order_of u = n ↔ ∃ m < n, gcd m n = 1 ∧ ↑(↑m / ↑n * p) = u :=
290
+ begin
291
+ refine ⟨λ hu, _, _⟩,
292
+ { rw ← hu at h,
293
+ obtain ⟨m, h₀, h₁, h₂⟩ := exists_gcd_eq_one_of_is_of_fin_add_order (add_order_of_pos_iff.mp h),
294
+ refine ⟨m, _, _, _⟩;
295
+ rwa ← hu, },
296
+ { rintros ⟨m, h₀, h₁, rfl⟩,
297
+ exact add_order_of_div_of_gcd_eq_one h h₁, },
298
+ end
299
+
300
+ variables (p)
301
+
302
+ /-- The natural bijection between points of order `n` and natural numbers less than and coprime to
303
+ `n`. The inverse of the map sends `m ↦ (m/n * p : add_circle p)` where `m` is coprime to `n` and
304
+ satisfies `0 ≤ m < n`. -/
305
+ def set_add_order_of_equiv {n : ℕ} (hn : 0 < n) :
306
+ {u : add_circle p | add_order_of u = n} ≃ {m | m < n ∧ gcd m n = 1 } :=
307
+ { to_fun := λ u, by
308
+ { let h := (add_order_of_eq_pos_iff hn).mp u.property,
309
+ exact ⟨classical.some h, classical.some (classical.some_spec h),
310
+ (classical.some_spec (classical.some_spec h)).1 ⟩, },
311
+ inv_fun := λ m, ⟨↑((m : 𝕜) / n * p), add_order_of_div_of_gcd_eq_one hn (m.property.2 )⟩,
312
+ left_inv := λ u, subtype.ext
313
+ (classical.some_spec (classical.some_spec $ (add_order_of_eq_pos_iff hn).mp u.2 )).2 ,
314
+ right_inv :=
315
+ begin
316
+ rintros ⟨m, hm₁, hm₂⟩,
317
+ let u : {u : add_circle p | add_order_of u = n} :=
318
+ ⟨↑((m : 𝕜) / n * p), add_order_of_div_of_gcd_eq_one hn hm₂⟩,
319
+ let h := (add_order_of_eq_pos_iff hn).mp u.property,
320
+ ext,
321
+ let m' := classical.some h,
322
+ change m' = m,
323
+ obtain ⟨h₁ : m' < n, h₂ : gcd m' n = 1 , h₃ : quotient_add_group.mk ((m' : 𝕜) / n * p) =
324
+ quotient_add_group.mk ((m : 𝕜) / n * p)⟩ := classical.some_spec h,
325
+ replace h₃ := congr_arg (coe : Ico 0 (0 + p) → 𝕜) (congr_arg (equiv_Ico p 0 ) h₃),
326
+ simpa only [coe_equiv_Ico_mk_apply, mul_left_inj' hp.out.ne', mul_div_cancel _ hp.out.ne',
327
+ int.fract_div_nat_cast_eq_div_nat_cast_mod,
328
+ div_left_inj' (nat.cast_ne_zero.mpr hn.ne' : (n : 𝕜) ≠ 0 ), nat.cast_inj,
329
+ (nat.mod_eq_iff_lt hn.ne').mpr hm₁, (nat.mod_eq_iff_lt hn.ne').mpr h₁] using h₃,
330
+ end }
331
+
332
+ @[simp] lemma card_add_order_of_eq_totient {n : ℕ} :
333
+ nat.card {u : add_circle p // add_order_of u = n} = n.totient :=
334
+ begin
335
+ rcases n.eq_zero_or_pos with rfl | hn,
336
+ { simp only [nat.totient_zero, add_order_of_eq_zero_iff],
337
+ rcases em (∃ (u : add_circle p), ¬ is_of_fin_add_order u) with ⟨u, hu⟩ | h,
338
+ { haveI : infinite {u : add_circle p // ¬is_of_fin_add_order u},
339
+ { erw infinite_coe_iff,
340
+ exact infinite_not_is_of_fin_add_order hu, },
341
+ exact nat.card_eq_zero_of_infinite, },
342
+ { haveI : is_empty {u : add_circle p // ¬is_of_fin_add_order u}, { simpa using h, },
343
+ exact nat.card_of_is_empty, }, },
344
+ { rw [← coe_set_of, nat.card_congr (set_add_order_of_equiv p hn),
345
+ n.totient_eq_card_lt_and_coprime],
346
+ simpa only [@nat.coprime_comm _ n], },
347
+ end
348
+
349
+ lemma finite_set_of_add_order_eq {n : ℕ} (hn : 0 < n) :
350
+ {u : add_circle p | add_order_of u = n}.finite :=
351
+ finite_coe_iff.mp $ nat.finite_of_card_ne_zero $ by simpa only [coe_set_of,
352
+ card_add_order_of_eq_totient p] using (nat.totient_pos hn).ne'
353
+
283
354
end finite_order_points
284
355
285
356
end linear_ordered_field
0 commit comments