@@ -20,10 +20,24 @@ Main definitions:
20
20
21
21
- `DirichletCharacter`: The type representing a Dirichlet character.
22
22
- `changeLevel`: Extend the Dirichlet character χ of level `n` to level `m`, where `n` divides `m`.
23
+ - `conductor`: The conductor of a Dirichlet character.
24
+ - `isPrimitive`: If the level is equal to the conductor.
23
25
24
26
## TODO
25
27
26
- - properties of the conductor
28
+ - reduction, even, odd
29
+ - add
30
+ `lemma unitsMap_surjective {n m : ℕ} (h : n ∣ m) (hm : m ≠ 0) : Function.Surjective (unitsMap h)`
31
+ and then
32
+ ```
33
+ lemma changeLevel_injective {d : ℕ} (h : d ∣ n) (hn : n ≠ 0) :
34
+ Function.Injective (changeLevel (R := R) h)
35
+ ```
36
+ and
37
+ ```
38
+ lemma changeLevel_one_iff {d : ℕ} {χ : DirichletCharacter R n} {χ' : DirichletCharacter R d}
39
+ (hdn : d ∣ n) (hn : n ≠ 0) (h : χ = changeLevel hdn χ') : χ = 1 ↔ χ' = 1
40
+ ```
27
41
28
42
## Tags
29
43
@@ -101,6 +115,31 @@ lemma same_level : FactorsThrough χ n := ⟨dvd_refl n, χ, (changeLevel_self
101
115
102
116
end FactorsThrough
103
117
118
+ lemma level_one (χ : DirichletCharacter R 1 ) : χ = 1 := by
119
+ ext
120
+ simp [units_eq_one]
121
+
122
+ lemma level_one' (hn : n = 1 ) : χ = 1 := by
123
+ subst hn
124
+ exact level_one _
125
+
126
+ instance : Subsingleton (DirichletCharacter R 1 ) := by
127
+ refine subsingleton_iff.mpr (fun χ χ' ↦ ?_)
128
+ simp [level_one]
129
+
130
+ noncomputable instance : Inhabited (DirichletCharacter R n) := ⟨1 ⟩
131
+
132
+ noncomputable instance : Unique (DirichletCharacter R 1 ) := Unique.mk' (DirichletCharacter R 1 )
133
+
134
+ lemma changeLevel_one {d : ℕ} (h : d ∣ n) :
135
+ changeLevel h (1 : DirichletCharacter R d) = 1 := by
136
+ simp [changeLevel]
137
+
138
+ lemma factorsThrough_one_iff : FactorsThrough χ 1 ↔ χ = 1 := by
139
+ refine ⟨fun ⟨_, χ₀, hχ₀⟩ ↦ ?_,
140
+ fun h ↦ ⟨one_dvd n, 1 , by rw [h, changeLevel_one]⟩⟩
141
+ rwa [level_one χ₀, changeLevel_one] at hχ₀
142
+
104
143
/-- The set of natural numbers for which a Dirichlet character is periodic. -/
105
144
def conductorSet : Set ℕ := {x : ℕ | FactorsThrough χ x}
106
145
@@ -114,4 +153,50 @@ lemma mem_conductorSet_dvd {x : ℕ} (hx : x ∈ conductorSet χ) : x ∣ n := h
114
153
The Dirichlet character `χ` can then alternatively be reformulated on `ℤ/nℤ`. -/
115
154
noncomputable def conductor : ℕ := sInf (conductorSet χ)
116
155
156
+ lemma conductor_mem_conductorSet : conductor χ ∈ conductorSet χ :=
157
+ Nat.sInf_mem (Set.nonempty_of_mem (level_mem_conductorSet χ))
158
+
159
+ lemma conductor_dvd_level : conductor χ ∣ n := (conductor_mem_conductorSet χ).dvd
160
+
161
+ lemma factorsThrough_conductor : FactorsThrough χ (conductor χ) := conductor_mem_conductorSet χ
162
+
163
+ lemma conductor_ne_zero (hn : n ≠ 0 ) : conductor χ ≠ 0 :=
164
+ fun h ↦ hn <| Nat.eq_zero_of_zero_dvd <| h ▸ conductor_dvd_level _
165
+
166
+ lemma conductor_one (hn : n ≠ 0 ) : conductor (1 : DirichletCharacter R n) = 1 := by
167
+ suffices : FactorsThrough (1 : DirichletCharacter R n) 1
168
+ · have h : conductor (1 : DirichletCharacter R n) ≤ 1 :=
169
+ Nat.sInf_le <| (mem_conductorSet_iff _).mpr this
170
+ exact Nat.le_antisymm h (Nat.pos_of_ne_zero <| conductor_ne_zero _ hn)
171
+ · exact (factorsThrough_one_iff _).mpr rfl
172
+
173
+ variable {χ}
174
+
175
+ lemma eq_one_iff_conductor_eq_one (hn : n ≠ 0 ) : χ = 1 ↔ conductor χ = 1 := by
176
+ refine ⟨fun h ↦ h ▸ conductor_one hn, fun hχ ↦ ?_⟩
177
+ obtain ⟨h', χ₀, h⟩ := factorsThrough_conductor χ
178
+ exact (level_one' χ₀ hχ ▸ h).trans <| changeLevel_one h'
179
+
180
+ lemma conductor_eq_zero_iff_level_eq_zero : conductor χ = 0 ↔ n = 0 := by
181
+ refine ⟨(conductor_ne_zero χ).mtr, ?_⟩
182
+ rintro rfl
183
+ exact Nat.sInf_eq_zero.mpr <| Or.inl <| level_mem_conductorSet χ
184
+
185
+ variable (χ)
186
+
187
+ /-- A character is primitive if its level is equal to its conductor. -/
188
+ def isPrimitive : Prop := conductor χ = n
189
+
190
+ lemma isPrimitive_def : isPrimitive χ ↔ conductor χ = n := Iff.rfl
191
+
192
+ lemma isPrimitive_one_level_one : isPrimitive (1 : DirichletCharacter R 1 ) :=
193
+ Nat.dvd_one.mp (conductor_dvd_level _)
194
+
195
+ lemma isPritive_one_level_zero : isPrimitive (1 : DirichletCharacter R 0 ) :=
196
+ conductor_eq_zero_iff_level_eq_zero.mpr rfl
197
+
198
+ lemma conductor_one_dvd (n : ℕ) : conductor (1 : DirichletCharacter R 1 ) ∣ n := by
199
+ rw [(isPrimitive_def _).mp isPrimitive_one_level_one]
200
+ apply one_dvd _
201
+
117
202
end DirichletCharacter
0 commit comments