|
| 1 | +/- |
| 2 | +Copyright (c) 2018 Johannes Hölzl. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Johannes Hölzl |
| 5 | +-/ |
| 6 | +import group_theory.coset |
| 7 | +open set function |
| 8 | + |
| 9 | +variables {α : Type*} {s : set α} {a a₁ a₂ b c: α} |
| 10 | + |
| 11 | +-- TODO this lemma isn't used anywhere in this file, and should be moved elsewhere. |
| 12 | +namespace finset |
| 13 | +open finset |
| 14 | + |
| 15 | +lemma mem_range_iff_mem_finset_range_of_mod_eq [decidable_eq α] {f : ℤ → α} {a : α} {n : ℕ} |
| 16 | + (hn : 0 < n) (h : ∀i, f (i % n) = f i) : |
| 17 | + a ∈ set.range f ↔ a ∈ (finset.range n).image (λi, f i) := |
| 18 | +suffices (∃i, f (i % n) = a) ↔ ∃i, i < n ∧ f ↑i = a, by simpa [h], |
| 19 | +have hn' : 0 < (n : ℤ), from int.coe_nat_lt.mpr hn, |
| 20 | +iff.intro |
| 21 | + (assume ⟨i, hi⟩, |
| 22 | + have 0 ≤ i % ↑n, from int.mod_nonneg _ (ne_of_gt hn'), |
| 23 | + ⟨int.to_nat (i % n), |
| 24 | + by rw [←int.coe_nat_lt, int.to_nat_of_nonneg this]; exact ⟨int.mod_lt_of_pos i hn', hi⟩⟩) |
| 25 | + (assume ⟨i, hi, ha⟩, |
| 26 | + ⟨i, by rw [int.mod_eq_of_lt (int.coe_zero_le _) (int.coe_nat_lt_coe_nat_of_lt hi), ha]⟩) |
| 27 | + |
| 28 | +end finset |
| 29 | + |
| 30 | +section order_of |
| 31 | +variables [group α] [fintype α] [decidable_eq α] |
| 32 | + |
| 33 | +lemma exists_gpow_eq_one (a : α) : ∃i≠0, a ^ (i:ℤ) = 1 := |
| 34 | +have ¬ injective (λi, a ^ i), |
| 35 | + from not_injective_int_fintype, |
| 36 | +let ⟨i, j, a_eq, ne⟩ := show ∃(i j : ℤ), a ^ i = a ^ j ∧ i ≠ j, |
| 37 | + by rw [injective] at this; simpa [classical.not_forall] in |
| 38 | +have a ^ (i - j) = 1, |
| 39 | + by simp [gpow_add, gpow_neg, a_eq], |
| 40 | +⟨i - j, sub_ne_zero.mpr ne, this⟩ |
| 41 | + |
| 42 | +lemma exists_pow_eq_one (a : α) : ∃i≠0, a ^ i = 1 := |
| 43 | +let ⟨i, hi, eq⟩ := exists_gpow_eq_one a in |
| 44 | +begin |
| 45 | + cases i, |
| 46 | + { exact ⟨i, by simp [int.of_nat_eq_coe, *] at *, eq⟩ }, |
| 47 | + { exact ⟨i + 1, dec_trivial, inv_eq_one.1 eq⟩ } |
| 48 | +end |
| 49 | + |
| 50 | +/-- `order_of a` is the order of the element `a`, i.e. the `n ≥ 1`, s.t. `a ^ n = 1` -/ |
| 51 | +def order_of (a : α) : ℕ := nat.find (exists_pow_eq_one a) |
| 52 | + |
| 53 | +lemma pow_order_of_eq_one (a : α) : a ^ order_of a = 1 := |
| 54 | +let ⟨h₁, h₂⟩ := nat.find_spec (exists_pow_eq_one a) in h₂ |
| 55 | + |
| 56 | +lemma order_of_ne_zero (a : α) : order_of a ≠ 0 := |
| 57 | +let ⟨h₁, h₂⟩ := nat.find_spec (exists_pow_eq_one a) in h₁ |
| 58 | + |
| 59 | +private lemma pow_injective_aux {n m : ℕ} (a : α) (h : n ≤ m) |
| 60 | + (hn : n < order_of a) (hm : m < order_of a) (eq : a ^ n = a ^ m) : n = m := |
| 61 | +decidable.by_contradiction $ assume ne : n ≠ m, |
| 62 | + have h₁ : m - n ≠ 0, by simp [nat.sub_eq_iff_eq_add h, ne.symm], |
| 63 | + have h₂ : a ^ (m - n) = 1, by simp [pow_sub _ h, eq], |
| 64 | + have le : order_of a ≤ m - n, from nat.find_min' (exists_pow_eq_one a) ⟨h₁, h₂⟩, |
| 65 | + have lt : m - n < order_of a, |
| 66 | + from (nat.sub_lt_left_iff_lt_add h).mpr $ nat.lt_add_left _ _ _ hm, |
| 67 | + lt_irrefl _ (lt_of_le_of_lt le lt) |
| 68 | + |
| 69 | +lemma pow_injective_of_lt_order_of {n m : ℕ} (a : α) |
| 70 | + (hn : n < order_of a) (hm : m < order_of a) (eq : a ^ n = a ^ m) : n = m := |
| 71 | +(le_total n m).elim |
| 72 | + (assume h, pow_injective_aux a h hn hm eq) |
| 73 | + (assume h, (pow_injective_aux a h hm hn eq.symm).symm) |
| 74 | + |
| 75 | +lemma order_of_le_card_univ : order_of a ≤ fintype.card α := |
| 76 | +finset.card_le_of_inj_on ((^) a) |
| 77 | + (assume n _, fintype.complete _) |
| 78 | + (assume i j, pow_injective_of_lt_order_of a) |
| 79 | + |
| 80 | +lemma pow_eq_mod_order_of {n : ℕ} : a ^ n = a ^ (n % order_of a) := |
| 81 | +calc a ^ n = a ^ (n % order_of a + order_of a * (n / order_of a)) : |
| 82 | + by rw [nat.mod_add_div] |
| 83 | + ... = a ^ (n % order_of a) : |
| 84 | + by simp [pow_add, pow_mul, pow_order_of_eq_one] |
| 85 | + |
| 86 | +lemma gpow_eq_mod_order_of {i : ℤ} : a ^ i = a ^ (i % order_of a) := |
| 87 | +calc a ^ i = a ^ (i % order_of a + order_of a * (i / order_of a)) : |
| 88 | + by rw [int.mod_add_div] |
| 89 | + ... = a ^ (i % order_of a) : |
| 90 | + by simp [gpow_add, gpow_mul, pow_order_of_eq_one] |
| 91 | + |
| 92 | +lemma mem_range_gpow_iff_mem_range_order_of {a a' : α} : |
| 93 | + a' ∈ range ((^) a : ℤ → α) ↔ a' ∈ (finset.range (order_of a)).image ((^) a : ℕ → α) := |
| 94 | +finset.mem_range_iff_mem_finset_range_of_mod_eq |
| 95 | + (nat.pos_iff_ne_zero.mpr (order_of_ne_zero a)) |
| 96 | + (assume i, gpow_eq_mod_order_of.symm) |
| 97 | + |
| 98 | +instance decidable_range_gpow : decidable_pred (range ((^) a : ℤ → α)) := |
| 99 | +assume a', decidable_of_iff' |
| 100 | + (a' ∈ (finset.range (order_of a)).image ((^) a)) |
| 101 | + mem_range_gpow_iff_mem_range_order_of |
| 102 | + |
| 103 | +section |
| 104 | +local attribute [instance] set_fintype |
| 105 | + |
| 106 | +lemma order_eq_card_range_gpow : order_of a = fintype.card (range ((^) a : ℤ → α)) := |
| 107 | +begin |
| 108 | + refine (finset.card_eq_of_bijective _ _ _ _).symm, |
| 109 | + { exact λn hn, ⟨gpow a n, mem_range_self n⟩ }, |
| 110 | + { exact assume ⟨_, i, rfl⟩ _, |
| 111 | + have pos: (0:int) < order_of a, |
| 112 | + from int.coe_nat_lt.mpr $ nat.pos_iff_ne_zero.mpr $ order_of_ne_zero a, |
| 113 | + have 0 ≤ i % (order_of a), |
| 114 | + from int.mod_nonneg _ $ ne_of_gt pos, |
| 115 | + ⟨int.to_nat (i % order_of a), |
| 116 | + by rw [← int.coe_nat_lt, int.to_nat_of_nonneg this]; |
| 117 | + exact ⟨int.mod_lt_of_pos _ pos, subtype.eq gpow_eq_mod_order_of.symm⟩⟩ }, |
| 118 | + { intros, exact finset.mem_univ _ }, |
| 119 | + { exact assume i j hi hj eq, pow_injective_of_lt_order_of a hi hj $ by simpa using eq } |
| 120 | +end |
| 121 | + |
| 122 | +section classical |
| 123 | +local attribute [instance] classical.prop_decidable |
| 124 | + |
| 125 | +/- TODO: use cardinal theory, introduce `card : set α → ℕ`, or setup decidability for cosets -/ |
| 126 | +lemma order_of_dvd_card_univ : order_of a ∣ fintype.card α := |
| 127 | +let ⟨equiv⟩ := (gpowers.is_subgroup a).group_equiv_left_cosets_times_subgroup in |
| 128 | +have ft_prod : fintype (left_cosets (gpowers a) × (gpowers a)), |
| 129 | + from fintype.of_equiv α equiv, |
| 130 | +have ft_s : fintype (gpowers a), |
| 131 | + from @fintype.fintype_prod_right _ _ _ ft_prod |
| 132 | + ⟨⟨(gpowers a), @is_subgroup.subgroup_mem_left_cosets α _ _ (gpowers.is_subgroup a)⟩⟩, |
| 133 | +have ft_cosets : fintype (left_cosets (gpowers a)), |
| 134 | + from @fintype.fintype_prod_left _ _ _ ft_prod ⟨⟨1, is_submonoid.one_mem (gpowers a)⟩⟩, |
| 135 | +have ft : fintype (left_cosets (gpowers a) × (gpowers a)), |
| 136 | + from @prod.fintype _ _ ft_cosets ft_s, |
| 137 | +have eq₁ : fintype.card α = @fintype.card _ ft_cosets * @fintype.card _ ft_s, |
| 138 | + from calc fintype.card α = @fintype.card _ ft_prod : |
| 139 | + (@fintype.card_eq _ _ _ ft_prod).2 (gpowers.is_subgroup a).group_equiv_left_cosets_times_subgroup |
| 140 | + ... = @fintype.card _ (@prod.fintype _ _ ft_cosets ft_s) : |
| 141 | + congr_arg (@fintype.card _) $ subsingleton.elim _ _ |
| 142 | + ... = @fintype.card _ ft_cosets * @fintype.card _ ft_s : |
| 143 | + @fintype.card_prod _ _ ft_cosets ft_s, |
| 144 | +have eq₂ : order_of a = @fintype.card _ ft_s, |
| 145 | + from calc order_of a = _ : order_eq_card_range_gpow |
| 146 | + ... = _ : congr_arg (@fintype.card _) $ subsingleton.elim _ _, |
| 147 | +dvd.intro (@fintype.card (left_cosets (gpowers a)) ft_cosets) $ |
| 148 | + by rw [eq₁, eq₂, mul_comm] |
| 149 | + |
| 150 | +end classical |
| 151 | + |
| 152 | +end |
| 153 | + |
| 154 | +end order_of |
0 commit comments