Skip to content

Commit

Permalink
Merge branch 'main' of github.com:ImperialCollegeLondon/FLT
Browse files Browse the repository at this point in the history
  • Loading branch information
kbuzzard committed Dec 25, 2024
2 parents 8d7fdd3 + 14bd4f6 commit 7a9b622
Show file tree
Hide file tree
Showing 15 changed files with 515 additions and 166 deletions.
30 changes: 30 additions & 0 deletions FLT/Deformations/RepresentationTheory/Irreducible.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,30 @@
import FLT.Deformations.RepresentationTheory.Subrepresentation
import FLT.Mathlib.RepresentationTheory.Basic

namespace Representation

variable {G : Type*} [Group G]

variable {k : Type*} [Field k]

variable {W : Type*} [AddCommMonoid W] [Module k W]

/-!
`IsIrreducible ρ` is the statement that a given Representation ρ is irreducible (also known as simple),
meaning that any subrepresentation must be either the full one (⊤) or zero (⊥)
This notion is only well behaved when the representation is over a field k. If it were defined over
a ring A with a nontrivial ideal J, the subrepresentation JW would often be a non trivial subrepresentation,
so ρ would rarely be irreducible.
-/
class IsIrreducible (ρ : Representation k G W) : Prop where
irreducible : IsSimpleOrder (Subrepresentation ρ)

/-!
`IsAbsolutelyIrreducible ρ` states that a given Representation ρ over a field k
is absolutely irreducible, meaning that all the possible base change extensions are irreducible.
-/
class IsAbsolutelyIrreducible (ρ : Representation k G W) : Prop where
absolutelyIrreducible : ∀ k', ∀ _ : Field k', ∀ _ : Algebra k k', IsIrreducible (k' ⊗ᵣ' ρ)

end Representation
74 changes: 74 additions & 0 deletions FLT/Deformations/RepresentationTheory/Subrepresentation.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,74 @@
import Mathlib.RepresentationTheory.Basic
import FLT.Mathlib.LinearAlgebra.Span.Defs

open Pointwise

universe u

variable {A : Type*} [CommRing A]

variable {G : Type*} [Group G]

variable {W : Type*} [AddCommMonoid W] [Module A W]

variable {ρ : Representation A G W}

variable (ρ) in
structure Subrepresentation where
toSubmodule : Submodule A W
apply_mem_toSubmodule (g : G) ⦃v : W⦄ : v ∈ toSubmodule → ρ g v ∈ toSubmodule

namespace Subrepresentation

lemma toSubmodule_injective : Function.Injective (toSubmodule : Subrepresentation ρ → Submodule A W) := by
rintro ⟨_,_⟩
congr!

instance : SetLike (Subrepresentation ρ) W where
coe ρ' := ρ'.toSubmodule
coe_injective' := SetLike.coe_injective.comp toSubmodule_injective

def toRepresentation (ρ' : Subrepresentation ρ): Representation A G ρ'.toSubmodule where
toFun g := (ρ g).restrict (ρ'.apply_mem_toSubmodule g)
map_one' := by ext; simp
map_mul' x y := by ext; simp

instance : Max (Subrepresentation ρ) where
max ρ₁ ρ₂ := .mk (ρ₁.toSubmodule ⊔ ρ₂.toSubmodule) <| by
simp only [Submodule.forall_mem_sup, map_add]
intro g x₁ hx₁ x₂ hx₂
exact Submodule.mem_sup.mpr
⟨ρ g x₁, ρ₁.apply_mem_toSubmodule g hx₁, ρ g x₂, ρ₂.apply_mem_toSubmodule g hx₂, rfl⟩

instance : Min (Subrepresentation ρ) where
min ρ₁ ρ₂ := .mk (ρ₁.toSubmodule ⊓ ρ₂.toSubmodule) <| by
simp only [Submodule.mem_inf, and_imp]
rintro g x hx₁ hx₂
exact ⟨ρ₁.apply_mem_toSubmodule g hx₁, ρ₂.apply_mem_toSubmodule g hx₂⟩


@[simp, norm_cast]
lemma coe_sup (ρ₁ ρ₂ : Subrepresentation ρ) : ↑(ρ₁ ⊔ ρ₂) = (ρ₁ : Set W) + (ρ₂ : Set W) :=
Submodule.coe_sup ρ₁.toSubmodule ρ₂.toSubmodule

@[simp, norm_cast]
lemma coe_inf (ρ₁ ρ₂ : Subrepresentation ρ) : ↑(ρ₁ ⊓ ρ₂) = (ρ₁ ∩ ρ₂ : Set W) := rfl

@[simp]
lemma toSubmodule_sup (ρ₁ ρ₂ : Subrepresentation ρ) :
(ρ₁ ⊔ ρ₂).toSubmodule = ρ₁.toSubmodule ⊔ ρ₂.toSubmodule := rfl

@[simp]
lemma toSubmodule_inf (ρ₁ ρ₂ : Subrepresentation ρ) :
(ρ₁ ⊓ ρ₂).toSubmodule = ρ₁.toSubmodule ⊓ ρ₂.toSubmodule := rfl

instance : Lattice (Subrepresentation ρ) :=
toSubmodule_injective.lattice _ toSubmodule_sup toSubmodule_inf

instance : BoundedOrder (Subrepresentation ρ) where
top := ⟨⊤, by simp⟩
le_top _ _ := by simp
bot := ⟨⊥, by simp⟩
bot_le _ _ := by simp +contextual

end Subrepresentation
3 changes: 0 additions & 3 deletions FLT/FLT_files.lean
Original file line number Diff line number Diff line change
Expand Up @@ -27,9 +27,6 @@ import FLT.Mathlib.Algebra.Module.Equiv.Defs
import FLT.Mathlib.Algebra.Module.LinearMap.Defs
import FLT.Mathlib.Algebra.Order.Hom.Monoid
import FLT.Mathlib.Algebra.Order.Monoid.Unbundled.TypeTags
import FLT.Mathlib.Analysis.Complex.Basic
import FLT.Mathlib.Data.Complex.Basic
import FLT.Mathlib.Data.ENNReal.Inv
import FLT.MathlibExperiments.Coalgebra.Monoid
import FLT.MathlibExperiments.Coalgebra.Sweedler
import FLT.MathlibExperiments.Coalgebra.TensorProduct
Expand Down
1 change: 0 additions & 1 deletion FLT/HaarMeasure/DistribHaarChar/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Andrew Yang, Yaël Dillies, Javier López-Contreras
-/
import FLT.HaarMeasure.DomMulActMeasure
import FLT.Mathlib.Data.ENNReal.Inv

/-!
# The distributive character of Haar measures
Expand Down
2 changes: 0 additions & 2 deletions FLT/HaarMeasure/DistribHaarChar/RealComplex.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,8 +8,6 @@ import Mathlib.MeasureTheory.Measure.Lebesgue.EqHaar
import Mathlib.RingTheory.Complex
import Mathlib.RingTheory.Norm.Transitivity
import FLT.HaarMeasure.DistribHaarChar.Basic
import FLT.Mathlib.Analysis.Complex.Basic
import FLT.Mathlib.Data.Complex.Basic
import FLT.Mathlib.LinearAlgebra.Determinant

/-!
Expand Down
8 changes: 0 additions & 8 deletions FLT/Mathlib/Analysis/Complex/Basic.lean

This file was deleted.

19 changes: 0 additions & 19 deletions FLT/Mathlib/Data/Complex/Basic.lean

This file was deleted.

41 changes: 0 additions & 41 deletions FLT/Mathlib/Data/ENNReal/Inv.lean

This file was deleted.

28 changes: 28 additions & 0 deletions FLT/Mathlib/LinearAlgebra/Span/Defs.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,28 @@
import Mathlib.LinearAlgebra.Span.Defs

open Pointwise

variable {R : Type*} [Semiring R]

variable {M : Type*} [AddCommMonoid M] [Module R M]

variable {p p' : Submodule R M}

variable {P : M → Prop}

namespace Submodule

@[simp high]
lemma forall_mem_sup : (∀ x ∈ p ⊔ p', P x) ↔ (∀ x₁ ∈ p, ∀ x₂ ∈ p', P (x₁ + x₂)) := by
simp [mem_sup]
aesop

@[simp high]
lemma exists_mem_sup : (∃ x ∈ p ⊔ p', P x) ↔ (∃ x₁ ∈ p, ∃ x₂ ∈ p', P (x₁ + x₂)) := by
simp [mem_sup]

@[simp, norm_cast]
lemma coe_sup' : ↑(p ⊔ p') = (p : Set M) + (p' : Set M) := by
simp [coe_sup]

end Submodule
38 changes: 38 additions & 0 deletions FLT/Mathlib/RepresentationTheory/Basic.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,38 @@
import Mathlib.LinearAlgebra.Matrix.GeneralLinearGroup.Defs
import Mathlib.RepresentationTheory.Basic

open LinearMap
open scoped TensorProduct

namespace Representation

variable {R V G ι: Type*} [CommRing R] [AddCommMonoid V] [Module R V] [Module.Free R V]
[Module.Finite R V] [Group G] [DecidableEq ι] [Fintype ι]

variable (ρ : Representation R G V) (𝓑 : Basis ι R V)

omit [Module.Free R V] [Module.Finite R V] in
@[simp]
lemma comp_def (g h : G): ρ g ∘ₗ ρ h = ρ g * ρ h := rfl

noncomputable def gl_map_of_basis
: G →* Matrix.GeneralLinearGroup ι R where
toFun g := {
val := LinearMap.toMatrix 𝓑 𝓑 (ρ g)
inv := LinearMap.toMatrix 𝓑 𝓑 (ρ g⁻¹)
val_inv := by rw [← toMatrix_comp, comp_def, ← map_mul]; simp
inv_val := by rw [← toMatrix_comp, comp_def, ← map_mul]; simp
}
map_one' := by aesop
map_mul' := by rintro x y; simp [LinearMap.toMatrix_mul]; norm_cast

noncomputable def baseChange (R' : Type*) [CommRing R'] [Algebra R R'] (ρ : Representation R G V)
: Representation R' G (R' ⊗[R] V) where
toFun g := LinearMap.baseChange R' (ρ g)
map_one' := by aesop
map_mul' := by aesop

scoped notation ρ "⊗ᵣ" ρ' => tprod ρ ρ'
scoped notation R' "⊗ᵣ'" ρ => baseChange R' ρ

end Representation
12 changes: 12 additions & 0 deletions blueprint/lean_decls
Original file line number Diff line number Diff line change
Expand Up @@ -89,6 +89,18 @@ Rat.AdeleRing.discrete
NumberField.AdeleRing.discrete
Rat.AdeleRing.cocompact
NumberField.AdeleRing.cocompact
distribHaarChar_real
distribHaarChar_complex
distribHaarChar_padic
addHaarScalarFactor_eq_distribHaarChar_det
distribHaarChar_algebra
addHaarScalarFactor_prod
addHaarScalarFactor_pi_finite
distribHaarChar_prod
distribHaarChar_pi_finite
addHaarScalarFactor_restricted_product
distribHaarChar_restricted_product
addHaarScalarFactor_eq_one
TotallyDefiniteQuaternionAlgebra.AutomorphicForm
TotallyDefiniteQuaternionAlgebra.AutomorphicForm.addCommGroup
TotallyDefiniteQuaternionAlgebra.AutomorphicForm.module
Expand Down
Loading

0 comments on commit 7a9b622

Please sign in to comment.