Skip to content

Commit d8ad8d8

Browse files
committed
feat: port Algebra.Module.Zlattice (#4944)
Maybe this file should be called `ZLattice`....
1 parent b0a2167 commit d8ad8d8

File tree

2 files changed

+264
-0
lines changed

2 files changed

+264
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -255,6 +255,7 @@ import Mathlib.Algebra.Module.Submodule.Lattice
255255
import Mathlib.Algebra.Module.Submodule.Pointwise
256256
import Mathlib.Algebra.Module.Torsion
257257
import Mathlib.Algebra.Module.ULift
258+
import Mathlib.Algebra.Module.Zlattice
258259
import Mathlib.Algebra.MonoidAlgebra.Basic
259260
import Mathlib.Algebra.MonoidAlgebra.Degree
260261
import Mathlib.Algebra.MonoidAlgebra.Division
Lines changed: 263 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,263 @@
1+
/-
2+
Copyright (c) 2023 Xavier Roblot. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Xavier Roblot
5+
6+
! This file was ported from Lean 3 source module algebra.module.zlattice
7+
! leanprover-community/mathlib commit a3e83f0fa4391c8740f7d773a7a9b74e311ae2a3
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.MeasureTheory.Group.FundamentalDomain
12+
13+
/-!
14+
# ℤ-lattices
15+
16+
Let `E` be a finite dimensional vector space over a `NormedLinearOrderedField` `K` with a solid
17+
norm and that is also a `FloorRing`, e.g. `ℚ` or `ℝ`. A (full) ℤ-lattice `L` of `E` is a discrete
18+
subgroup of `E` such that `L` spans `E` over `K`.
19+
20+
The ℤ-lattice `L` can be defined in two ways:
21+
* For `b` a basis of `E`, then `Submodule.span ℤ (Set.range b)` is a ℤ-lattice of `E`.
22+
* As an `AddSubgroup E` with the additional properties:
23+
* `∀ r : ℝ, (L ∩ Metric.closedBall 0 r).finite`, that is `L` is discrete
24+
* `Submodule.span ℝ (L : Set E) = ⊤`, that is `L` spans `E` over `K`.
25+
26+
## Main result
27+
* `Zspan.isAddFundamentalDomain`: for a ℤ-lattice `Submodule.span ℤ (Set.range b)`, proves that
28+
the set defined by `Zspan.fundamentalDomain` is a fundamental domain.
29+
30+
-/
31+
32+
33+
open scoped BigOperators
34+
35+
noncomputable section
36+
37+
namespace Zspan
38+
39+
open MeasureTheory MeasurableSet Submodule
40+
41+
variable {E ι : Type _}
42+
43+
section NormedLatticeField
44+
45+
variable {K : Type _} [NormedLinearOrderedField K]
46+
47+
variable [NormedAddCommGroup E] [NormedSpace K E]
48+
49+
variable (b : Basis ι K E)
50+
51+
/-- The fundamental domain of the ℤ-lattice spanned by `b`. See `Zspan.isAddFundamentalDomain`
52+
for the proof that it is a fundamental domain. -/
53+
def fundamentalDomain : Set E := {m | ∀ i, b.repr m i ∈ Set.Ico (0 : K) 1}
54+
#align zspan.fundamental_domain Zspan.fundamentalDomain
55+
56+
@[simp]
57+
theorem mem_fundamentalDomain {m : E} :
58+
m ∈ fundamentalDomain b ↔ ∀ i, b.repr m i ∈ Set.Ico (0 : K) 1 :=
59+
Iff.rfl
60+
#align zspan.mem_fundamental_domain Zspan.mem_fundamentalDomain
61+
62+
variable [FloorRing K]
63+
64+
section Fintype
65+
66+
variable [Fintype ι]
67+
68+
/-- The map that sends a vector of `E` to the element of the ℤ-lattice spanned by `b` obtained
69+
by rounding down its coordinates on the basis `b`. -/
70+
def floor (m : E) : span ℤ (Set.range b) := ∑ i, ⌊b.repr m i⌋ • b.restrictScalars ℤ i
71+
#align zspan.floor Zspan.floor
72+
73+
/-- The map that sends a vector of `E` to the element of the ℤ-lattice spanned by `b` obtained
74+
by rounding up its coordinates on the basis `b`. -/
75+
def ceil (m : E) : span ℤ (Set.range b) := ∑ i, ⌈b.repr m i⌉ • b.restrictScalars ℤ i
76+
#align zspan.ceil Zspan.ceil
77+
78+
@[simp]
79+
theorem repr_floor_apply (m : E) (i : ι) : b.repr (floor b m) i = ⌊b.repr m i⌋ := by
80+
classical simp only [floor, zsmul_eq_smul_cast K, b.repr.map_smul, Finsupp.single_apply,
81+
Finset.sum_apply', Basis.repr_self, Finsupp.smul_single', mul_one, Finset.sum_ite_eq', coe_sum,
82+
Finset.mem_univ, if_true, coe_smul_of_tower, Basis.restrictScalars_apply, LinearEquiv.map_sum]
83+
#align zspan.repr_floor_apply Zspan.repr_floor_apply
84+
85+
@[simp]
86+
theorem repr_ceil_apply (m : E) (i : ι) : b.repr (ceil b m) i = ⌈b.repr m i⌉ := by
87+
classical simp only [ceil, zsmul_eq_smul_cast K, b.repr.map_smul, Finsupp.single_apply,
88+
Finset.sum_apply', Basis.repr_self, Finsupp.smul_single', mul_one, Finset.sum_ite_eq', coe_sum,
89+
Finset.mem_univ, if_true, coe_smul_of_tower, Basis.restrictScalars_apply, LinearEquiv.map_sum]
90+
#align zspan.repr_ceil_apply Zspan.repr_ceil_apply
91+
92+
@[simp]
93+
theorem floor_eq_self_of_mem (m : E) (h : m ∈ span ℤ (Set.range b)) : (floor b m : E) = m := by
94+
apply b.ext_elem
95+
simp_rw [repr_floor_apply b]
96+
intro i
97+
obtain ⟨z, hz⟩ := (b.mem_span_iff_repr_mem ℤ _).mp h i
98+
rw [← hz]
99+
exact congr_arg (Int.cast : ℤ → K) (Int.floor_intCast z)
100+
#align zspan.floor_eq_self_of_mem Zspan.floor_eq_self_of_mem
101+
102+
@[simp]
103+
theorem ceil_eq_self_of_mem (m : E) (h : m ∈ span ℤ (Set.range b)) : (ceil b m : E) = m := by
104+
apply b.ext_elem
105+
simp_rw [repr_ceil_apply b]
106+
intro i
107+
obtain ⟨z, hz⟩ := (b.mem_span_iff_repr_mem ℤ _).mp h i
108+
rw [← hz]
109+
exact congr_arg (Int.cast : ℤ → K) (Int.ceil_intCast z)
110+
#align zspan.ceil_eq_self_of_mem Zspan.ceil_eq_self_of_mem
111+
112+
/-- The map that sends a vector `E` to the `fundamentalDomain` of the lattice,
113+
see `Zspan.fract_mem_fundamentalDomain`. -/
114+
def fract (m : E) : E := m - floor b m
115+
#align zspan.fract Zspan.fract
116+
117+
theorem fract_apply (m : E) : fract b m = m - floor b m := rfl
118+
#align zspan.fract_apply Zspan.fract_apply
119+
120+
@[simp]
121+
theorem repr_fract_apply (m : E) (i : ι) : b.repr (fract b m) i = Int.fract (b.repr m i) := by
122+
rw [fract, LinearEquiv.map_sub, Finsupp.coe_sub, Pi.sub_apply, repr_floor_apply, Int.fract]
123+
#align zspan.repr_fract_apply Zspan.repr_fract_apply
124+
125+
@[simp]
126+
theorem fract_fract (m : E) : fract b (fract b m) = fract b m :=
127+
Basis.ext_elem b fun _ => by classical simp only [repr_fract_apply, Int.fract_fract]
128+
#align zspan.fract_fract Zspan.fract_fract
129+
130+
@[simp]
131+
theorem fract_zspan_add (m : E) {v : E} (h : v ∈ span ℤ (Set.range b)) :
132+
fract b (v + m) = fract b m := by
133+
classical
134+
refine (Basis.ext_elem_iff b).mpr fun i => ?_
135+
simp_rw [repr_fract_apply, Int.fract_eq_fract]
136+
use (b.restrictScalars ℤ).repr ⟨v, h⟩ i
137+
rw [map_add, Finsupp.coe_add, Pi.add_apply, add_tsub_cancel_right,
138+
← eq_intCast (algebraMap ℤ K) _, Basis.restrictScalars_repr_apply, coe_mk]
139+
#align zspan.fract_zspan_add Zspan.fract_zspan_add
140+
141+
@[simp]
142+
theorem fract_add_zspan (m : E) {v : E} (h : v ∈ span ℤ (Set.range b)) :
143+
fract b (m + v) = fract b m := by rw [add_comm, fract_zspan_add b m h]
144+
#align zspan.fract_add_zspan Zspan.fract_add_zspan
145+
146+
variable {b}
147+
148+
theorem fract_eq_self {x : E} : fract b x = x ↔ x ∈ fundamentalDomain b := by
149+
classical simp only [Basis.ext_elem_iff b, repr_fract_apply, Int.fract_eq_self,
150+
mem_fundamentalDomain, Set.mem_Ico]
151+
#align zspan.fract_eq_self Zspan.fract_eq_self
152+
153+
variable (b)
154+
155+
theorem fract_mem_fundamentalDomain (x : E) : fract b x ∈ fundamentalDomain b :=
156+
fract_eq_self.mp (fract_fract b _)
157+
#align zspan.fract_mem_fundamental_domain Zspan.fract_mem_fundamentalDomain
158+
159+
theorem fract_eq_fract (m n : E) : fract b m = fract b n ↔ -m + n ∈ span ℤ (Set.range b) := by
160+
classical
161+
rw [eq_comm, Basis.ext_elem_iff b]
162+
simp_rw [repr_fract_apply, Int.fract_eq_fract, eq_comm, Basis.mem_span_iff_repr_mem,
163+
sub_eq_neg_add, map_add, LinearEquiv.map_neg, Finsupp.coe_add, Finsupp.coe_neg, Pi.add_apply,
164+
Pi.neg_apply, ← eq_intCast (algebraMap ℤ K) _, Set.mem_range]
165+
#align zspan.fract_eq_fract Zspan.fract_eq_fract
166+
167+
theorem norm_fract_le [HasSolidNorm K] (m : E) : ‖fract b m‖ ≤ ∑ i, ‖b i‖ := by
168+
classical
169+
calc
170+
‖fract b m‖ = ‖∑ i, b.repr (fract b m) i • b i‖ := by rw [b.sum_repr]
171+
_ = ‖∑ i, Int.fract (b.repr m i) • b i‖ := by simp_rw [repr_fract_apply]
172+
_ ≤ ∑ i, ‖Int.fract (b.repr m i) • b i‖ := (norm_sum_le _ _)
173+
_ = ∑ i, ‖Int.fract (b.repr m i)‖ * ‖b i‖ := by simp_rw [norm_smul]
174+
_ ≤ ∑ i, ‖b i‖ := Finset.sum_le_sum fun i _ => ?_
175+
suffices ‖Int.fract ((b.repr m) i)‖ ≤ 1 by
176+
convert mul_le_mul_of_nonneg_right this (norm_nonneg _ : 0 ≤ ‖b i‖)
177+
exact (one_mul _).symm
178+
rw [(norm_one.symm : 1 = ‖(1 : K)‖)]
179+
apply norm_le_norm_of_abs_le_abs
180+
rw [abs_one, Int.abs_fract]
181+
exact le_of_lt (Int.fract_lt_one _)
182+
#align zspan.norm_fract_le Zspan.norm_fract_le
183+
184+
section Unique
185+
186+
variable [Unique ι]
187+
188+
@[simp]
189+
theorem coe_floor_self (k : K) : (floor (Basis.singleton ι K) k : K) = ⌊k⌋ :=
190+
Basis.ext_elem _ fun _ => by rw [repr_floor_apply, Basis.singleton_repr, Basis.singleton_repr]
191+
#align zspan.coe_floor_self Zspan.coe_floor_self
192+
193+
@[simp]
194+
theorem coe_fract_self (k : K) : (fract (Basis.singleton ι K) k : K) = Int.fract k :=
195+
Basis.ext_elem _ fun _ => by rw [repr_fract_apply, Basis.singleton_repr, Basis.singleton_repr]
196+
#align zspan.coe_fract_self Zspan.coe_fract_self
197+
198+
end Unique
199+
200+
end Fintype
201+
202+
theorem fundamentalDomain_bounded [Finite ι] [HasSolidNorm K] :
203+
Metric.Bounded (fundamentalDomain b) := by
204+
cases nonempty_fintype ι
205+
use 2 * ∑ j, ‖b j‖
206+
intro x hx y hy
207+
refine le_trans (dist_le_norm_add_norm x y) ?_
208+
rw [← fract_eq_self.mpr hx, ← fract_eq_self.mpr hy]
209+
refine (add_le_add (norm_fract_le b x) (norm_fract_le b y)).trans ?_
210+
rw [← two_mul]
211+
#align zspan.fundamental_domain_bounded Zspan.fundamentalDomain_bounded
212+
213+
theorem vadd_mem_fundamentalDomain [Fintype ι] (y : span ℤ (Set.range b)) (x : E) :
214+
y +ᵥ x ∈ fundamentalDomain b ↔ y = -floor b x := by
215+
rw [Subtype.ext_iff, ← add_right_inj x, AddSubgroupClass.coe_neg, ← sub_eq_add_neg, ← fract_apply,
216+
← fract_zspan_add b _ (Subtype.mem y), add_comm, ← vadd_eq_add, ← vadd_def, eq_comm, ←
217+
fract_eq_self]
218+
#align zspan.vadd_mem_fundamental_domain Zspan.vadd_mem_fundamentalDomain
219+
220+
theorem exist_unique_vadd_mem_fundamentalDomain [Finite ι] (x : E) :
221+
∃! v : span ℤ (Set.range b), v +ᵥ x ∈ fundamentalDomain b := by
222+
cases nonempty_fintype ι
223+
refine ⟨-floor b x, ?_, fun y h => ?_⟩
224+
· exact (vadd_mem_fundamentalDomain b (-floor b x) x).mpr rfl
225+
· exact (vadd_mem_fundamentalDomain b y x).mp h
226+
#align zspan.exist_unique_vadd_mem_fundamental_domain Zspan.exist_unique_vadd_mem_fundamentalDomain
227+
228+
end NormedLatticeField
229+
230+
section Real
231+
232+
variable [NormedAddCommGroup E] [NormedSpace ℝ E]
233+
234+
variable (b : Basis ι ℝ E)
235+
236+
@[measurability]
237+
theorem fundamentalDomain_measurableSet [MeasurableSpace E] [OpensMeasurableSpace E] [Finite ι] :
238+
MeasurableSet (fundamentalDomain b) := by
239+
haveI : FiniteDimensional ℝ E := FiniteDimensional.of_fintype_basis b
240+
let f := (Finsupp.linearEquivFunOnFinite ℝ ℝ ι).toLinearMap.comp b.repr.toLinearMap
241+
let D : Set (ι → ℝ) := Set.pi Set.univ fun _ : ι => Set.Ico (0 : ℝ) 1
242+
rw [(_ : fundamentalDomain b = f ⁻¹' D)]
243+
· refine measurableSet_preimage (LinearMap.continuous_of_finiteDimensional f).measurable ?_
244+
exact MeasurableSet.pi Set.countable_univ fun _ _ => measurableSet_Ico
245+
· ext
246+
simp only [fundamentalDomain, Set.mem_setOf_eq, LinearMap.coe_comp,
247+
LinearEquiv.coe_toLinearMap, Set.mem_preimage, Function.comp_apply, Set.mem_univ_pi,
248+
Finsupp.linearEquivFunOnFinite_apply]
249+
#align zspan.fundamental_domain_measurable_set Zspan.fundamentalDomain_measurableSet
250+
251+
/-- For a ℤ-lattice `Submodule.span ℤ (Set.range b)`, proves that the set defined
252+
by `Zspan.fundamentalDomain` is a fundamental domain. -/
253+
protected theorem isAddFundamentalDomain [Finite ι] [MeasurableSpace E] [OpensMeasurableSpace E]
254+
(μ : Measure E) :
255+
IsAddFundamentalDomain (span ℤ (Set.range b)).toAddSubgroup (fundamentalDomain b) μ := by
256+
cases nonempty_fintype ι
257+
exact IsAddFundamentalDomain.mk' (nullMeasurableSet (fundamentalDomain_measurableSet b))
258+
fun x => exist_unique_vadd_mem_fundamentalDomain b x
259+
#align zspan.is_add_fundamental_domain Zspan.isAddFundamentalDomain
260+
261+
end Real
262+
263+
end Zspan

0 commit comments

Comments
 (0)