Loogle!
Result
Found 51 declarations mentioning SlashAction.map.
- SlashAction.map π Mathlib.NumberTheory.ModularForms.SlashActions
{Ξ² : Type u_1} {G : Type u_2} {Ξ± : Type u_3} {instβ : Monoid G} {instβΒΉ : AddMonoid Ξ±} [self : SlashAction Ξ² G Ξ±] : Ξ² β G β Ξ± β Ξ± - SlashAction.slash_one π Mathlib.NumberTheory.ModularForms.SlashActions
{Ξ² : Type u_1} {G : Type u_2} {Ξ± : Type u_3} {instβ : Monoid G} {instβΒΉ : AddMonoid Ξ±} [self : SlashAction Ξ² G Ξ±] (k : Ξ²) (a : Ξ±) : SlashAction.map k 1 a = a - SlashAction.zero_slash π Mathlib.NumberTheory.ModularForms.SlashActions
{Ξ² : Type u_1} {G : Type u_2} {Ξ± : Type u_3} {instβ : Monoid G} {instβΒΉ : AddMonoid Ξ±} [self : SlashAction Ξ² G Ξ±] (k : Ξ²) (g : G) : SlashAction.map k g 0 = 0 - SlashAction.slash_mul π Mathlib.NumberTheory.ModularForms.SlashActions
{Ξ² : Type u_1} {G : Type u_2} {Ξ± : Type u_3} {instβ : Monoid G} {instβΒΉ : AddMonoid Ξ±} [self : SlashAction Ξ² G Ξ±] (k : Ξ²) (g h : G) (a : Ξ±) : SlashAction.map k (g * h) a = SlashAction.map k h (SlashAction.map k g a) - SlashAction.add_slash π Mathlib.NumberTheory.ModularForms.SlashActions
{Ξ² : Type u_1} {G : Type u_2} {Ξ± : Type u_3} {instβ : Monoid G} {instβΒΉ : AddMonoid Ξ±} [self : SlashAction Ξ² G Ξ±] (k : Ξ²) (g : G) (a b : Ξ±) : SlashAction.map k g (a + b) = SlashAction.map k g a + SlashAction.map k g b - SlashAction.neg_slash π Mathlib.NumberTheory.ModularForms.SlashActions
{Ξ² : Type u_1} {G : Type u_2} {Ξ± : Type u_3} [Monoid G] [AddGroup Ξ±] [SlashAction Ξ² G Ξ±] (k : Ξ²) (g : G) (a : Ξ±) : SlashAction.map k g (-a) = -SlashAction.map k g a - ModularForm.is_invariant_const π Mathlib.NumberTheory.ModularForms.SlashActions
(A : Matrix.SpecialLinearGroup (Fin 2) β€) (x : β) : SlashAction.map 0 A (Function.const UpperHalfPlane x) = Function.const UpperHalfPlane x - ModularForm.is_invariant_one π Mathlib.NumberTheory.ModularForms.SlashActions
(A : Matrix.SpecialLinearGroup (Fin 2) β€) : SlashAction.map 0 A 1 = 1 - ModularForm.SL_smul_slash π Mathlib.NumberTheory.ModularForms.SlashActions
{Ξ± : Type u_1} [SMul Ξ± β] [IsScalarTower Ξ± β β] (k : β€) (A : Matrix.SpecialLinearGroup (Fin 2) β€) (f : UpperHalfPlane β β) (c : Ξ±) : SlashAction.map k A (c β’ f) = c β’ SlashAction.map k A f - ModularForm.mul_slash_SL2 π Mathlib.NumberTheory.ModularForms.SlashActions
(k1 k2 : β€) (A : Matrix.SpecialLinearGroup (Fin 2) β€) (f g : UpperHalfPlane β β) : SlashAction.map (k1 + k2) A (f * g) = SlashAction.map k1 A f * SlashAction.map k2 A g - ModularForm.smul_slash π Mathlib.NumberTheory.ModularForms.SlashActions
(k : β€) (A : GL (Fin 2) β) (f : UpperHalfPlane β β) (c : β) : SlashAction.map k A (c β’ f) = (UpperHalfPlane.Ο A) c β’ SlashAction.map k A f - ModularForm.slash_action_eq'_iff π Mathlib.NumberTheory.ModularForms.SlashActions
(k : β€) (f : UpperHalfPlane β β) (Ξ³ : Matrix.SpecialLinearGroup (Fin 2) β€) (z : UpperHalfPlane) : SlashAction.map k Ξ³ f z = f z β f (Ξ³ β’ z) = (β(βΞ³ 1 0) * βz + β(βΞ³ 1 1)) ^ k * f z - ModularForm.slash_apply π Mathlib.NumberTheory.ModularForms.SlashActions
{k : β€} (f : UpperHalfPlane β β) (g : GL (Fin 2) β) (Ο : UpperHalfPlane) : SlashAction.map k g f Ο = (UpperHalfPlane.Ο g) (f (g β’ Ο)) * β|β(Matrix.GeneralLinearGroup.det g)| ^ (k - 1) * UpperHalfPlane.denom g βΟ ^ (-k) - ModularForm.slash_def π Mathlib.NumberTheory.ModularForms.SlashActions
{k : β€} (f : UpperHalfPlane β β) (g : GL (Fin 2) β) : SlashAction.map k g f = fun Ο => (UpperHalfPlane.Ο g) (f (g β’ Ο)) * β|β(Matrix.GeneralLinearGroup.det g)| ^ (k - 1) * UpperHalfPlane.denom g βΟ ^ (-k) - ModularForm.mul_slash π Mathlib.NumberTheory.ModularForms.SlashActions
(k1 k2 : β€) (A : GL (Fin 2) β) (f g : UpperHalfPlane β β) : SlashAction.map (k1 + k2) A (f * g) = |β(Matrix.GeneralLinearGroup.det A)| β’ (SlashAction.map k1 A f * SlashAction.map k2 A g) - ModularForm.is_invariant_one' π Mathlib.NumberTheory.ModularForms.SlashActions
(A : Matrix.SpecialLinearGroup (Fin 2) β€) : SlashAction.map 0 (Matrix.SpecialLinearGroup.toGL ((Matrix.SpecialLinearGroup.map (Int.castRingHom β)) A)) 1 = 1 - ModularForm.SL_slash_apply π Mathlib.NumberTheory.ModularForms.SlashActions
{k : β€} (f : UpperHalfPlane β β) (Ξ³ : Matrix.SpecialLinearGroup (Fin 2) β€) (Ο : UpperHalfPlane) : SlashAction.map k Ξ³ f Ο = f (Ξ³ β’ Ο) * UpperHalfPlane.denom (Matrix.SpecialLinearGroup.toGL ((Matrix.SpecialLinearGroup.map (Int.castRingHom β)) Ξ³)) βΟ ^ (-k) - ModularForm.SL_slash_def π Mathlib.NumberTheory.ModularForms.SlashActions
{k : β€} (f : UpperHalfPlane β β) (Ξ³ : Matrix.SpecialLinearGroup (Fin 2) β€) : SlashAction.map k Ξ³ f = fun Ο => f (Ξ³ β’ Ο) * UpperHalfPlane.denom (Matrix.SpecialLinearGroup.toGL ((Matrix.SpecialLinearGroup.map (Int.castRingHom β)) Ξ³)) βΟ ^ (-k) - ModularForm.SL_slash π Mathlib.NumberTheory.ModularForms.SlashActions
{k : β€} (f : UpperHalfPlane β β) (Ξ³ : Matrix.SpecialLinearGroup (Fin 2) β€) : SlashAction.map k Ξ³ f = SlashAction.map k (Matrix.SpecialLinearGroup.toGL ((Matrix.SpecialLinearGroup.map (Int.castRingHom β)) Ξ³)) f - UpperHalfPlane.IsBoundedAtImInfty.slash π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
{g : GL (Fin 2) β} {f : UpperHalfPlane β β} (k : β€) (hg : βg 1 0 = 0) (hf : UpperHalfPlane.IsBoundedAtImInfty f) : UpperHalfPlane.IsBoundedAtImInfty (SlashAction.map k g f) - UpperHalfPlane.IsZeroAtImInfty.slash π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
{g : GL (Fin 2) β} {f : UpperHalfPlane β β} (k : β€) (hg : βg 1 0 = 0) (hf : UpperHalfPlane.IsZeroAtImInfty f) : UpperHalfPlane.IsZeroAtImInfty (SlashAction.map k g f) - OnePoint.IsBoundedAt.smul_iff π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
{c : OnePoint β} {f : UpperHalfPlane β β} {k : β€} {g : GL (Fin 2) β} : (g β’ c).IsBoundedAt f k β c.IsBoundedAt (SlashAction.map k g f) k - OnePoint.IsZeroAt.smul_iff π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
{c : OnePoint β} {f : UpperHalfPlane β β} {k : β€} {g : GL (Fin 2) β} : (g β’ c).IsZeroAt f k β c.IsZeroAt (SlashAction.map k g f) k - OnePoint.isBoundedAt_iff π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
{c : OnePoint β} {f : UpperHalfPlane β β} {k : β€} {g : GL (Fin 2) β} (hg : g β’ OnePoint.infty = c) : c.IsBoundedAt f k β UpperHalfPlane.IsBoundedAtImInfty (SlashAction.map k g f) - OnePoint.IsBoundedAt.eq_1 π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
(c : OnePoint β) (f : UpperHalfPlane β β) (k : β€) : c.IsBoundedAt f k = β (g : GL (Fin 2) β), g β’ OnePoint.infty = c β UpperHalfPlane.IsBoundedAtImInfty (SlashAction.map k g f) - OnePoint.isZeroAt_iff π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
{c : OnePoint β} {f : UpperHalfPlane β β} {k : β€} {g : GL (Fin 2) β} (hg : g β’ OnePoint.infty = c) : c.IsZeroAt f k β UpperHalfPlane.IsZeroAtImInfty (SlashAction.map k g f) - OnePoint.IsZeroAt.eq_1 π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
(c : OnePoint β) (f : UpperHalfPlane β β) (k : β€) : c.IsZeroAt f k = β (g : GL (Fin 2) β), g β’ OnePoint.infty = c β UpperHalfPlane.IsZeroAtImInfty (SlashAction.map k g f) - OnePoint.isBoundedAt_iff_forall_SL2Z π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
{c : OnePoint β} {f : UpperHalfPlane β β} {k : β€} (hc : IsCusp c (Matrix.SpecialLinearGroup.mapGL β).range) : c.IsBoundedAt f k β β (Ξ³ : Matrix.SpecialLinearGroup (Fin 2) β€), (Matrix.SpecialLinearGroup.mapGL β) Ξ³ β’ OnePoint.infty = c β UpperHalfPlane.IsBoundedAtImInfty (SlashAction.map k Ξ³ f) - OnePoint.isZeroAt_iff_forall_SL2Z π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
{c : OnePoint β} {f : UpperHalfPlane β β} {k : β€} (hc : IsCusp c (Matrix.SpecialLinearGroup.mapGL β).range) : c.IsZeroAt f k β β (Ξ³ : Matrix.SpecialLinearGroup (Fin 2) β€), (Matrix.SpecialLinearGroup.mapGL β) Ξ³ β’ OnePoint.infty = c β UpperHalfPlane.IsZeroAtImInfty (SlashAction.map k Ξ³ f) - OnePoint.isBoundedAt_iff_exists_SL2Z π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
{c : OnePoint β} {f : UpperHalfPlane β β} {k : β€} (hc : IsCusp c (Matrix.SpecialLinearGroup.mapGL β).range) : c.IsBoundedAt f k β β Ξ³, (Matrix.SpecialLinearGroup.mapGL β) Ξ³ β’ OnePoint.infty = c β§ UpperHalfPlane.IsBoundedAtImInfty (SlashAction.map k Ξ³ f) - OnePoint.isZeroAt_iff_exists_SL2Z π Mathlib.NumberTheory.ModularForms.BoundedAtCusp
{c : OnePoint β} {f : UpperHalfPlane β β} {k : β€} (hc : IsCusp c (Matrix.SpecialLinearGroup.mapGL β).range) : c.IsZeroAt f k β β Ξ³, (Matrix.SpecialLinearGroup.mapGL β) Ξ³ β’ OnePoint.infty = c β§ UpperHalfPlane.IsZeroAtImInfty (SlashAction.map k Ξ³ f) - SlashInvariantForm.mk π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{Ξ : outParam (Subgroup (GL (Fin 2) β))} {k : outParam β€} (toFun : UpperHalfPlane β β) (slash_action_eq' : β Ξ³ β Ξ, SlashAction.map k Ξ³ toFun = toFun) : SlashInvariantForm Ξ k - SlashInvariantForm.slash_action_eq' π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{Ξ : outParam (Subgroup (GL (Fin 2) β))} {k : outParam β€} (self : SlashInvariantForm Ξ k) (Ξ³ : GL (Fin 2) β) : Ξ³ β Ξ β SlashAction.map k Ξ³ self.toFun = self.toFun - SlashInvariantForm.slash_action_eqn π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{F : Type u_1} {Ξ : Subgroup (GL (Fin 2) β)} {k : β€} [FunLike F UpperHalfPlane β] [SlashInvariantFormClass F Ξ k] (f : F) (Ξ³ : GL (Fin 2) β) (hΞ³ : Ξ³ β Ξ) : SlashAction.map k Ξ³ βf = βf - SlashInvariantForm.coe_mk π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{Ξ : outParam (Subgroup (GL (Fin 2) β))} {k : outParam β€} (f : UpperHalfPlane β β) (hf : β Ξ³ β Ξ, SlashAction.map k Ξ³ f = f) : β{ toFun := f, slash_action_eq' := hf } = f - SlashInvariantForm.mk.sizeOf_spec π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{Ξ : outParam (Subgroup (GL (Fin 2) β))} {k : outParam β€} (toFun : UpperHalfPlane β β) (slash_action_eq' : β Ξ³ β Ξ, SlashAction.map k Ξ³ toFun = toFun) : sizeOf { toFun := toFun, slash_action_eq' := slash_action_eq' } = 1 - SlashInvariantFormClass.mk π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{F : Type u_1} {Ξ : outParam (Subgroup (GL (Fin 2) β))} {k : outParam β€} [FunLike F UpperHalfPlane β] (slash_action_eq : β (f : F), β Ξ³ β Ξ, SlashAction.map k Ξ³ βf = βf) : SlashInvariantFormClass F Ξ k - SlashInvariantFormClass.slash_action_eq π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{F : Type u_1} {Ξ : outParam (Subgroup (GL (Fin 2) β))} {k : outParam β€} {instβ : FunLike F UpperHalfPlane β} [self : SlashInvariantFormClass F Ξ k] (f : F) (Ξ³ : GL (Fin 2) β) : Ξ³ β Ξ β SlashAction.map k Ξ³ βf = βf - SlashInvariantForm.mk.injEq π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{Ξ : outParam (Subgroup (GL (Fin 2) β))} {k : outParam β€} (toFun : UpperHalfPlane β β) (slash_action_eq' : β Ξ³ β Ξ, SlashAction.map k Ξ³ toFun = toFun) (toFunβ : UpperHalfPlane β β) (slash_action_eq'β : β Ξ³ β Ξ, SlashAction.map k Ξ³ toFunβ = toFunβ) : ({ toFun := toFun, slash_action_eq' := slash_action_eq' } = { toFun := toFunβ, slash_action_eq' := slash_action_eq'β }) = (toFun = toFunβ) - SlashInvariantForm.coe_translate π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{F : Type u_1} {Ξ : Subgroup (GL (Fin 2) β)} {k : β€} [FunLike F UpperHalfPlane β] [SlashInvariantFormClass F Ξ k] (f : F) (g : GL (Fin 2) β) : β(SlashInvariantForm.translate f g) = SlashAction.map k g βf - SlashInvariantForm.coe_translateGL π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{F : Type u_1} {Ξ : Subgroup (GL (Fin 2) β)} {k : β€} [FunLike F UpperHalfPlane β] [SlashInvariantFormClass F Ξ k] (f : F) (g : GL (Fin 2) β) : β(SlashInvariantForm.translate f g) = SlashAction.map k g βf - SlashInvariantForm.coe_translateGLPos π Mathlib.NumberTheory.ModularForms.SlashInvariantForms
{F : Type u_1} {Ξ : Subgroup (GL (Fin 2) β)} {k : β€} [FunLike F UpperHalfPlane β] [SlashInvariantFormClass F Ξ k] (f : F) (g : GL (Fin 2) β) : β(SlashInvariantForm.translate f g) = SlashAction.map k g βf - ModularFormClass.bdd_at_infty_slash π Mathlib.NumberTheory.ModularForms.Basic
{k : β€} {F : Type u_1} {Ξ : Subgroup (GL (Fin 2) β)} [Ξ.IsArithmetic] [FunLike F UpperHalfPlane β] [ModularFormClass F Ξ k] (f : F) (g : Matrix.SpecialLinearGroup (Fin 2) β€) : UpperHalfPlane.IsBoundedAtImInfty (SlashAction.map k g βf) - CuspFormClass.zero_at_infty_slash π Mathlib.NumberTheory.ModularForms.Basic
{k : β€} {F : Type u_1} {Ξ : Subgroup (GL (Fin 2) β)} [Ξ.IsArithmetic] [FunLike F UpperHalfPlane β] [CuspFormClass F Ξ k] (f : F) (g : Matrix.SpecialLinearGroup (Fin 2) β€) : UpperHalfPlane.IsZeroAtImInfty (SlashAction.map k g βf) - MDifferentiable.slash π Mathlib.NumberTheory.ModularForms.Basic
{f : UpperHalfPlane β β} (hf : MDifferentiable (modelWithCornersSelf β β) (modelWithCornersSelf β β) f) (k : β€) (g : GL (Fin 2) β) : MDifferentiable (modelWithCornersSelf β β) (modelWithCornersSelf β β) (SlashAction.map k g f) - ModularForm.coe_translate π Mathlib.NumberTheory.ModularForms.Basic
{k : β€} {Ξ : Subgroup (GL (Fin 2) β)} {F : Type u_1} [FunLike F UpperHalfPlane β] (f : F) [ModularFormClass F Ξ k] (g : GL (Fin 2) β) : β(ModularForm.translate f g) = SlashAction.map k g βf - CuspForm.coe_translate π Mathlib.NumberTheory.ModularForms.Basic
{k : β€} {Ξ : Subgroup (GL (Fin 2) β)} {F : Type u_1} [FunLike F UpperHalfPlane β] (f : F) [CuspFormClass F Ξ k] (g : Matrix.SpecialLinearGroup (Fin 2) β€) : β(CuspForm.translate f (Matrix.SpecialLinearGroup.toGL ((Matrix.SpecialLinearGroup.map (Int.castRingHom β)) g))) = SlashAction.map k g βf - EisensteinSeries.eisensteinSeries_slash_apply π Mathlib.NumberTheory.ModularForms.EisensteinSeries.Defs
{N : β} (a : Fin 2 β ZMod N) (k : β€) (Ξ³ : Matrix.SpecialLinearGroup (Fin 2) β€) : SlashAction.map k Ξ³ (eisensteinSeries a k) = eisensteinSeries (Matrix.vecMul a β((Matrix.SpecialLinearGroup.map (Int.castRingHom (ZMod N))) Ξ³)) k - EisensteinSeries.isBoundedAtImInfty_eisensteinSeries_SIF π Mathlib.NumberTheory.ModularForms.EisensteinSeries.IsBoundedAtImInfty
{N : β} [NeZero N] (a : Fin 2 β ZMod N) {k : β€} (hk : 3 β€ k) (A : Matrix.SpecialLinearGroup (Fin 2) β€) : UpperHalfPlane.IsBoundedAtImInfty (SlashAction.map k A (EisensteinSeries.eisensteinSeries_SIF a k).toFun) - UpperHalfPlane.petersson_slash_SL π Mathlib.NumberTheory.ModularForms.Petersson
(k : β€) (f f' : UpperHalfPlane β β) (g : Matrix.SpecialLinearGroup (Fin 2) β€) (Ο : UpperHalfPlane) : UpperHalfPlane.petersson k (SlashAction.map k g f) (SlashAction.map k g f') Ο = UpperHalfPlane.petersson k f f' (g β’ Ο) - UpperHalfPlane.petersson_slash π Mathlib.NumberTheory.ModularForms.Petersson
(k : β€) (f f' : UpperHalfPlane β β) (g : GL (Fin 2) β) (Ο : UpperHalfPlane) : UpperHalfPlane.petersson k (SlashAction.map k g f) (SlashAction.map k g f') Ο = β|β(Matrix.GeneralLinearGroup.det g)| ^ (k - 2) * (UpperHalfPlane.Ο g) (UpperHalfPlane.petersson k f f' (g β’ Ο))
About
Loogle searches Lean and Mathlib definitions and theorems.
You can use Loogle from within the Lean4 VSCode language extension
using (by default) Ctrl-K Ctrl-S. You can also try the
#loogle
command from LeanSearchClient,
the CLI version, the Loogle
VS Code extension, the lean.nvim
integration or the Zulip bot.
Usage
Loogle finds definitions and lemmas in various ways:
By constant:
πReal.sin
finds all lemmas whose statement somehow mentions the sine function.By lemma name substring:
π"differ"
finds all lemmas that have"differ"
somewhere in their lemma name.By subexpression:
π_ * (_ ^ _)
finds all lemmas whose statements somewhere include a product where the second argument is raised to some power.The pattern can also be non-linear, as in
πReal.sqrt ?a * Real.sqrt ?a
If the pattern has parameters, they are matched in any order. Both of these will find
List.map
:
π(?a -> ?b) -> List ?a -> List ?b
πList ?a -> (?a -> ?b) -> List ?b
By main conclusion:
π|- tsum _ = _ * tsum _
finds all lemmas where the conclusion (the subexpression to the right of allβ
andβ
) has the given shape.As before, if the pattern has parameters, they are matched against the hypotheses of the lemma in any order; for example,
π|- _ < _ β tsum _ < tsum _
will findtsum_lt_tsum
even though the hypothesisf i < g i
is not the last.
If you pass more than one such search filter, separated by commas
Loogle will return lemmas which match all of them. The
search
π Real.sin, "two", tsum, _ * _, _ ^ _, |- _ < _ β _
would find all lemmas which mention the constants Real.sin
and tsum
, have "two"
as a substring of the
lemma name, include a product and a power somewhere in the type,
and have a hypothesis of the form _ < _
(if
there were any such lemmas). Metavariables (?a
) are
assigned independently in each filter.
The #lucky
button will directly send you to the
documentation of the first hit.
Source code
You can find the source code for this service at https://github.com/nomeata/loogle. The https://loogle.lean-lang.org/ service is provided by the Lean FRO.
This is Loogle revision ee8c038
serving mathlib revision 7a9e177