Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 422e70f

Browse files
committed
feat(topology/instances/add_circle): expand API for finite-order points on the circle (#17986)
1 parent 46a64b5 commit 422e70f

File tree

4 files changed

+130
-8
lines changed

4 files changed

+130
-8
lines changed

src/data/set/finite.lean

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -945,6 +945,17 @@ lemma infinite.exists_not_mem_finset {s : set α} (hs : s.infinite) (f : finset
945945
let ⟨a, has, haf⟩ := (hs.diff (to_finite f)).nonempty
946946
in ⟨a, has, λ h, haf $ finset.mem_coe.1 h⟩
947947

948+
lemma not_inj_on_infinite_finite_image {f : α → β} {s : set α}
949+
(h_inf : s.infinite) (h_fin : (f '' s).finite) :
950+
¬ inj_on f s :=
951+
begin
952+
haveI : finite (f '' s) := finite_coe_iff.mpr h_fin,
953+
haveI : infinite s := infinite_coe_iff.mpr h_inf,
954+
have := not_injective_infinite_finite
955+
((f '' s).cod_restrict (s.restrict f) $ λ x, ⟨x, x.property, rfl⟩),
956+
contrapose! this,
957+
rwa [injective_cod_restrict, ← inj_on_iff_injective],
958+
end
948959

949960
/-! ### Order properties -/
950961

src/data/set/intervals/infinite.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -52,7 +52,7 @@ instance [no_min_order α] {a : α} : infinite (Iic a) := no_min_order.infinite
5252
lemma Iic_infinite [no_min_order α] (a : α) : (Iic a).infinite := infinite_coe_iff.1 Iic.infinite
5353

5454
instance [no_max_order α] {a : α} : infinite (Ioi a) := no_max_order.infinite
55-
lemma Ioi_infinite [no_min_order α] (a : α) : (Iio a).infinite := infinite_coe_iff.1 Iio.infinite
55+
lemma Ioi_infinite [no_max_order α] (a : α) : (Ioi a).infinite := infinite_coe_iff.1 Ioi.infinite
5656

5757
instance [no_max_order α] {a : α} : infinite (Ici a) := no_max_order.infinite
5858
lemma Ici_infinite [no_max_order α] (a : α) : (Ici a).infinite := infinite_coe_iff.1 Ici.infinite

src/group_theory/order_of_element.lean

Lines changed: 47 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Julian Kuelshammer
66
import algebra.hom.iterate
77
import data.nat.modeq
88
import data.set.pointwise.basic
9+
import data.set.intervals.infinite
910
import dynamics.periodic_pts
1011
import group_theory.index
1112

@@ -68,6 +69,19 @@ lemma is_of_fin_order_iff_pow_eq_one (x : G) :
6869
is_of_fin_order x ↔ ∃ n, 0 < n ∧ x ^ n = 1 :=
6970
by { convert iff.rfl, simp [is_periodic_pt_mul_iff_pow_eq_one] }
7071

72+
/-- See also `injective_pow_iff_not_is_of_fin_order`. -/
73+
@[to_additive not_is_of_fin_add_order_of_injective_nsmul "See also
74+
`injective_nsmul_iff_not_is_of_fin_add_order`."]
75+
lemma not_is_of_fin_order_of_injective_pow {x : G} (h : injective (λ (n : ℕ), x^n)) :
76+
¬ is_of_fin_order x :=
77+
begin
78+
simp_rw [is_of_fin_order_iff_pow_eq_one, not_exists, not_and],
79+
intros n hn_pos hnx,
80+
rw ← pow_zero x at hnx,
81+
rw h hnx at hn_pos,
82+
exact irrefl 0 hn_pos,
83+
end
84+
7185
/-- Elements of finite order are of finite order in submonoids.-/
7286
@[to_additive is_of_fin_add_order_iff_coe "Elements of finite order are of finite order in
7387
submonoids."]
@@ -380,9 +394,11 @@ lemma mem_powers_iff_mem_range_order_of' [decidable_eq G] (hx : 0 < order_of x)
380394
y ∈ submonoid.powers x ↔ y ∈ (finset.range (order_of x)).image ((^) x : ℕ → G) :=
381395
finset.mem_range_iff_mem_finset_range_of_mod_eq' hx (λ i, pow_eq_mod_order_of.symm)
382396

397+
@[to_additive]
383398
lemma pow_eq_one_iff_modeq : x ^ n = 1 ↔ n ≡ 0 [MOD (order_of x)] :=
384399
by rw [modeq_zero_iff_dvd, order_of_dvd_iff_pow_eq_one]
385400

401+
@[to_additive]
386402
lemma pow_eq_pow_iff_modeq : x ^ n = x ^ m ↔ n ≡ m [MOD (order_of x)] :=
387403
begin
388404
wlog hmn : m ≤ n,
@@ -391,6 +407,34 @@ begin
391407
exact ⟨λ h, nat.modeq.add_left _ h, λ h, nat.modeq.add_left_cancel' _ h⟩,
392408
end
393409

410+
@[simp, to_additive injective_nsmul_iff_not_is_of_fin_add_order]
411+
lemma injective_pow_iff_not_is_of_fin_order {x : G} :
412+
injective (λ (n : ℕ), x^n) ↔ ¬ is_of_fin_order x :=
413+
begin
414+
refine ⟨λ h, not_is_of_fin_order_of_injective_pow h, λ h n m hnm, _⟩,
415+
rwa [pow_eq_pow_iff_modeq, order_of_eq_zero_iff.mpr h, modeq_zero_iff] at hnm,
416+
end
417+
418+
@[to_additive infinite_not_is_of_fin_add_order]
419+
lemma infinite_not_is_of_fin_order {x : G} (h : ¬ is_of_fin_order x) :
420+
{y : G | ¬ is_of_fin_order y}.infinite :=
421+
begin
422+
let s := {n | 0 < n}.image (λ (n : ℕ), x^n),
423+
have hs : s ⊆ {y : G | ¬ is_of_fin_order y},
424+
{ rintros - ⟨n, hn : 0 < n, rfl⟩ (contra : is_of_fin_order (x^n)),
425+
apply h,
426+
rw is_of_fin_order_iff_pow_eq_one at contra ⊢,
427+
obtain ⟨m, hm, hm'⟩ := contra,
428+
exact ⟨n * m, mul_pos hn hm, by rwa pow_mul⟩, },
429+
suffices : s.infinite, { exact this.mono hs, },
430+
contrapose! h,
431+
have : ¬ injective (λ (n : ℕ), x^n),
432+
{ have := set.not_inj_on_infinite_finite_image (set.Ioi_infinite 0) (set.not_infinite.mp h),
433+
contrapose! this,
434+
exact set.inj_on_of_injective this _, },
435+
rwa [injective_pow_iff_not_is_of_fin_order, not_not] at this,
436+
end
437+
394438
end cancel_monoid
395439

396440
section group
@@ -541,13 +585,9 @@ variables [left_cancel_monoid G] [add_left_cancel_monoid A]
541585
@[to_additive]
542586
lemma exists_pow_eq_one [finite G] (x : G) : is_of_fin_order x :=
543587
begin
544-
refine (is_of_fin_order_iff_pow_eq_one _).mpr _,
545-
obtain ⟨i, j, a_eq, ne⟩ : ∃(i j : ℕ), x ^ i = x ^ j ∧ i ≠ j :=
546-
by simpa only [not_forall, exists_prop, injective]
547-
using (not_injective_infinite_finite (λi:ℕ, x^i)),
548-
wlog h'' : j ≤ i,
549-
refine ⟨i - j, tsub_pos_of_lt (lt_of_le_of_ne h'' ne.symm), mul_right_injective (x^j) _⟩,
550-
rw [mul_one, ← pow_add, ← a_eq, add_tsub_cancel_of_le h''],
588+
have : (set.univ : set G).finite := set.univ.to_finite,
589+
contrapose! this,
590+
exact set.infinite.mono (set.subset_univ _) (infinite_not_is_of_fin_order this),
551591
end
552592

553593
@[to_additive add_order_of_le_card_univ]

src/topology/instances/add_circle.lean

Lines changed: 71 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,7 @@ Copyright (c) 2022 Oliver Nash. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Oliver Nash
55
-/
6+
import data.nat.totient
67
import algebra.ring.add_aut
78
import group_theory.divisible
89
import group_theory.order_of_element
@@ -67,6 +68,10 @@ lemma coe_nsmul {n : ℕ} {x : 𝕜} : (↑(n • x) : add_circle p) = n • (x
6768

6869
lemma coe_zsmul {n : ℤ} {x : 𝕜} : (↑(n • x) : add_circle p) = n • (x : add_circle p) := rfl
6970

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+
7075
lemma coe_neg {x : 𝕜} : (↑(-x) : add_circle p) = -(x : add_circle p) := rfl
7176

7277
lemma coe_eq_zero_iff {x : 𝕜} : (x : add_circle p) = 0 ↔ ∃ (n : ℤ), n • p = x :=
@@ -280,6 +285,72 @@ begin
280285
simpa [zero_add] using (equiv_Ico p 0 u).2.2, },
281286
end
282287

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+
283354
end finite_order_points
284355

285356
end linear_ordered_field

0 commit comments

Comments
 (0)