Skip to content

Commit 217555e

Browse files
Fix and silence some warnings (#164)
This makes it somewhat easier to spot legitimate output. Co-authored-by: Pietro Monticone <[email protected]>
1 parent d7d2013 commit 217555e

8 files changed

+22
-13
lines changed

FLT/AutomorphicForm/QuaternionAlgebra.lean

+9-9
Original file line numberDiff line numberDiff line change
@@ -20,7 +20,7 @@ suppress_compilation
2020

2121
variable (F : Type*) [Field F] [NumberField F]
2222

23-
variable (D : Type*) [Ring D] [Algebra F D] [FiniteDimensional F D]
23+
variable (D : Type*) [Ring D] [Algebra F D]
2424

2525
open DedekindDomain
2626

@@ -56,7 +56,7 @@ end missing_instances
5656

5757
instance : TopologicalSpace (D ⊗[F] (FiniteAdeleRing (𝓞 F) F)) := actionTopology (FiniteAdeleRing (𝓞 F) F) _
5858
instance : IsActionTopology (FiniteAdeleRing (𝓞 F) F) (D ⊗[F] (FiniteAdeleRing (𝓞 F) F)) := ⟨rfl⟩
59-
instance : TopologicalRing (D ⊗[F] (FiniteAdeleRing (𝓞 F) F)) :=
59+
instance [FiniteDimensional F D] : TopologicalRing (D ⊗[F] (FiniteAdeleRing (𝓞 F) F)) :=
6060
-- this def would be a dangerous instance
6161
-- (it can't guess R) but it's just a Prop so we can easily add it here
6262
ActionTopology.Module.topologicalRing (FiniteAdeleRing (𝓞 F) F) _
@@ -93,16 +93,16 @@ attribute [coe] AutomorphicForm.toFun
9393
theorem ext (φ ψ : AutomorphicForm F D M) (h : ∀ x, φ x = ψ x) : φ = ψ := by
9494
cases φ; cases ψ; simp only [mk.injEq]; ext; apply h
9595

96-
def zero : (AutomorphicForm F D M) where
96+
def zero [FiniteDimensional F D] : (AutomorphicForm F D M) where
9797
toFun := 0
9898
left_invt := by simp
9999
loc_cst := by use ⊤; simp
100100

101-
instance : Zero (AutomorphicForm F D M) where
101+
instance [FiniteDimensional F D] : Zero (AutomorphicForm F D M) where
102102
zero := zero
103103

104104
@[simp]
105-
theorem zero_apply (x : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) :
105+
theorem zero_apply [FiniteDimensional F D] (x : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) :
106106
(0 : AutomorphicForm F D M) x = 0 := rfl
107107

108108
def neg (φ : AutomorphicForm F D M) : AutomorphicForm F D M where
@@ -147,7 +147,7 @@ instance : Add (AutomorphicForm F D M) where
147147
theorem add_apply (φ ψ : AutomorphicForm F D M) (x : (D ⊗[F] (FiniteAdeleRing (𝓞 F) F))ˣ) :
148148
(φ + ψ) x = (φ x) + (ψ x) := rfl
149149

150-
instance addCommGroup : AddCommGroup (AutomorphicForm F D M) where
150+
instance addCommGroup [FiniteDimensional F D] : AddCommGroup (AutomorphicForm F D M) where
151151
add := (· + ·)
152152
add_assoc := by intros; ext; simp [add_assoc];
153153
zero := 0
@@ -184,7 +184,7 @@ theorem toConjAct_open {G : Type*} [Group G] [TopologicalSpace G] [TopologicalGr
184184
group
185185
exact hu
186186

187-
instance : SMul (Dfx F D) (AutomorphicForm F D M) where
187+
instance [FiniteDimensional F D] : SMul (Dfx F D) (AutomorphicForm F D M) where
188188
smul g φ := { -- (g • f) (x) := f(xg) -- x(gf)=(xg)f
189189
toFun := fun x => φ (x * g)
190190
left_invt := by
@@ -204,10 +204,10 @@ instance : SMul (Dfx F D) (AutomorphicForm F D M) where
204204
}
205205

206206
@[simp]
207-
theorem sMul_eval (g : Dfx F D) (f : AutomorphicForm F D M) (x : (D ⊗[F] FiniteAdeleRing (𝓞 F) F)ˣ) :
207+
theorem sMul_eval [FiniteDimensional F D] (g : Dfx F D) (f : AutomorphicForm F D M) (x : (D ⊗[F] FiniteAdeleRing (𝓞 F) F)ˣ) :
208208
(g • f) x = f (x * g) := rfl
209209

210-
instance : MulAction (Dfx F D) (AutomorphicForm F D M) where
210+
instance [FiniteDimensional F D] : MulAction (Dfx F D) (AutomorphicForm F D M) where
211211
smul := (· • ·)
212212
one_smul := by intros; ext; simp only [sMul_eval, mul_one]
213213
mul_smul := by intros; ext; simp only [sMul_eval, mul_assoc]

FLT/HIMExperiments/ContinuousSMul_topology.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -94,7 +94,7 @@ lemma induced_continuous_smul [τA : TopologicalSpace A] [ContinuousSMul R A] (h
9494
@ContinuousSMul S B _ _ (TopologicalSpace.induced g τA) := by
9595
convert Inducing.continuousSMul (inducing_induced g) hf (fun {c} {x} ↦ map_smulₛₗ g c x)
9696

97-
#check Prod.continuousSMul -- exists and is an instance :-)
97+
-- #check Prod.continuousSMul -- exists and is an instance :-)
9898
--#check Induced.continuousSMul -- doesn't exist
9999

100100
end continuous_smul

FLT/HIMExperiments/FGModuleTopology.lean

+2
Original file line numberDiff line numberDiff line change
@@ -31,6 +31,8 @@ or $p$-adics).
3131
3232
-/
3333

34+
set_option linter.unusedSectionVars false
35+
3436
section basics
3537

3638
variable (R : Type*) [TopologicalSpace R] [Semiring R]

FLT/HIMExperiments/dual_topology.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -94,7 +94,7 @@ namespace dual_topology
9494

9595
section basics
9696

97-
variable (R : Type*) [Monoid R] [TopologicalSpace R] [ContinuousMul R]
97+
variable (R : Type*) [Monoid R] [TopologicalSpace R]
9898
variable (A : Type*) [SMul R A]
9999

100100
abbrev dualTopology : TopologicalSpace A :=
@@ -121,10 +121,10 @@ lemma isDualTopology [τA : TopologicalSpace A] [IsDualTopology R A] :
121121
end basics
122122

123123
-- Non-commutative variables
124-
variable (R : Type*) [Monoid R] [τR: TopologicalSpace R] [ContinuousMul R]
124+
variable (R : Type*) [Monoid R] [τR: TopologicalSpace R]
125125

126126

127-
lemma Module.topology_self : τR = dualTopology R R := by
127+
lemma Module.topology_self [ContinuousMul R] : τR = dualTopology R R := by
128128
refine le_antisymm (le_iInf (fun i ↦ ?_)) <| sInf_le ⟨MulActionHom.id R, induced_id⟩
129129
rw [← continuous_iff_le_induced,
130130
show i = ⟨fun r ↦ r • i 1, fun _ _ ↦ mul_assoc _ _ _⟩ by

FLT/HIMExperiments/module_topology.lean

+2
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,8 @@ and target.
2626
2727
-/
2828

29+
set_option linter.unusedSectionVars false
30+
2931
-- This was an early theorem I proved when I didn't know what was true or not, and
3032
-- was just experimenting.
3133

FLT/HIMExperiments/right_module_topology.lean

+2
Original file line numberDiff line numberDiff line change
@@ -35,6 +35,8 @@ functions from `M` (now considered only as an index set, so with no topology) to
3535
3636
-/
3737

38+
set_option linter.unusedSectionVars false
39+
3840
-- See FLT.ForMathlib.ActionTopology for a parallel development.
3941
namespace right_module_topology
4042

FLT/MathlibExperiments/Frobenius.lean

+2
Original file line numberDiff line numberDiff line change
@@ -116,6 +116,8 @@ for "helping me to understand the math proof and to formalize it."
116116
117117
-/
118118

119+
set_option linter.unusedSectionVars false
120+
119121
open Classical
120122

121123
section FiniteFrobeniusDef

FLT/MathlibExperiments/Frobenius2.lean

+1
Original file line numberDiff line numberDiff line change
@@ -181,6 +181,7 @@ open scoped algebraMap
181181
noncomputable local instance : Algebra A[X] B[X] :=
182182
RingHom.toAlgebra (Polynomial.mapRingHom (Algebra.toRingHom))
183183

184+
omit [Fintype (B ≃ₐ[A] B)] [DecidableEq (Ideal B)] in
184185
@[simp, norm_cast]
185186
lemma coe_monomial (n : ℕ) (a : A) : ((monomial n a : A[X]) : B[X]) = monomial n (a : B) :=
186187
map_monomial _

0 commit comments

Comments
 (0)