Loogle!
Result
Found 54 declarations mentioning Subalgebra.map.
- Subalgebra.map π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A ββ[R] B) (S : Subalgebra R A) : Subalgebra R B - Subalgebra.map_id π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (S : Subalgebra R A) : Subalgebra.map (AlgHom.id R A) S = S - Subalgebra.map_injective π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {f : A ββ[R] B} (hf : Function.Injective βf) : Function.Injective (Subalgebra.map f) - Subalgebra.comap_map_eq_self_of_injective π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {f : A ββ[R] B} (hf : Function.Injective βf) (S : Subalgebra R A) : Subalgebra.comap f (Subalgebra.map f S) = S - AlgHom.range_comp π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} {C : Type w'} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Semiring C] [Algebra R C] (f : A ββ[R] B) (g : B ββ[R] C) : (g.comp f).range = Subalgebra.map g f.range - Subalgebra.coe_map π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A ββ[R] B) (S : Subalgebra R A) : β(Subalgebra.map f S) = βf '' βS - Subalgebra.gc_map_comap π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A ββ[R] B) : GaloisConnection (Subalgebra.map f) (Subalgebra.comap f) - Subalgebra.map_map π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} {C : Type w'} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Semiring C] [Algebra R C] (S : Subalgebra R A) (g : B ββ[R] C) (f : A ββ[R] B) : Subalgebra.map g (Subalgebra.map f S) = Subalgebra.map (g.comp f) S - Subalgebra.map_toSubsemiring π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A ββ[R] B) (S : Subalgebra R A) : (Subalgebra.map f S).toSubsemiring = Subsemiring.map (βf) S.toSubsemiring - Subalgebra.mem_map π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {S : Subalgebra R A} {f : A ββ[R] B} {y : B} : y β Subalgebra.map f S β β x β S, f x = y - Subalgebra.map_mono π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {Sβ Sβ : Subalgebra R A} {f : A ββ[R] B} : Sβ β€ Sβ β Subalgebra.map f Sβ β€ Subalgebra.map f Sβ - Subalgebra.map_le π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {S : Subalgebra R A} {f : A ββ[R] B} {U : Subalgebra R B} : Subalgebra.map f S β€ U β S β€ Subalgebra.comap f U - AlgHom.subalgebraMap π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (S : Subalgebra R A) (f : A ββ[R] B) : β₯S ββ[R] β₯(Subalgebra.map f S) - Subalgebra.range_comp_val π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (S : Subalgebra R A) (f : A ββ[R] B) : (f.comp S.val).range = Subalgebra.map f S - Subalgebra.equivMapOfInjective π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (S : Subalgebra R A) (f : A ββ[R] B) (hf : Function.Injective βf) : β₯S ββ[R] β₯(Subalgebra.map f S) - AlgEquiv.subalgebraMap π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (e : A ββ[R] B) (S : Subalgebra R A) : β₯S ββ[R] β₯(Subalgebra.map (βe) S) - AlgHom.subalgebraMap_surjective π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (S : Subalgebra R A) (f : A ββ[R] B) : Function.Surjective β(AlgHom.subalgebraMap S f) - AlgHom.subalgebraMap_coe_apply π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {S : Subalgebra R A} (f : A ββ[R] B) (x : β₯S) : β((AlgHom.subalgebraMap S f) x) = f βx - Subalgebra.coe_equivMapOfInjective_apply π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (S : Subalgebra R A) (f : A ββ[R] B) (hf : Function.Injective βf) (x : β₯S) : β((S.equivMapOfInjective f hf) x) = f βx - AlgEquiv.subalgebraMap_apply_coe π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (e : A ββ[R] B) (S : Subalgebra R A) (x : ββS.toAddSubmonoid) : β((e.subalgebraMap S) x) = e βx - Subalgebra.map_toSubmodule π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {S : Subalgebra R A} {f : A ββ[R] B} : Subalgebra.toSubmodule (Subalgebra.map f S) = Submodule.map f.toLinearMap (Subalgebra.toSubmodule S) - AlgEquiv.subalgebraMap_symm_apply_coe π Mathlib.Algebra.Algebra.Subalgebra.Basic
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (e : A ββ[R] B) (S : Subalgebra R A) (y : β(ββe.toRingEquiv.toAddEquiv '' βS.toAddSubmonoid)) : β((e.subalgebraMap S).symm y) = (ββe).symm βy - AlgHom.map_adjoin π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (Ο : A ββ[R] B) (s : Set A) : Subalgebra.map Ο (Algebra.adjoin R s) = Algebra.adjoin R (βΟ '' s) - Algebra.adjoin_image π Mathlib.Algebra.Algebra.Subalgebra.Lattice
(R : Type uR) {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (f : A ββ[R] B) (s : Set A) : Algebra.adjoin R (βf '' s) = Subalgebra.map f (Algebra.adjoin R s) - Subalgebra.map_comap_eq_self_of_surjective π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {f : A ββ[R] B} (hf : Function.Surjective βf) (S : Subalgebra R B) : Subalgebra.map f (Subalgebra.comap f S) = S - AlgHom.map_adjoin_singleton π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (e : A ββ[R] B) (x : A) : Subalgebra.map e (Algebra.adjoin R {x}) = Algebra.adjoin R {e x} - Subalgebra.map_comap_eq π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A ββ[R] B) (S : Subalgebra R B) : Subalgebra.map f (Subalgebra.comap f S) = S β f.range - Subalgebra.map_comap_eq_self π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {f : A ββ[R] B} {S : Subalgebra R B} (h : S β€ f.range) : Subalgebra.map f (Subalgebra.comap f S) = S - Algebra.map_iInf π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {ΞΉ : Sort u_1} [Nonempty ΞΉ] (f : A ββ[R] B) (hf : Function.Injective βf) (s : ΞΉ β Subalgebra R A) : Subalgebra.map f (iInf s) = β¨ i, Subalgebra.map f (s i) - Algebra.map_sup π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A ββ[R] B) (S T : Subalgebra R A) : Subalgebra.map f (S β T) = Subalgebra.map f S β Subalgebra.map f T - Subalgebra.comap_map_eq_self π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Ring A] [Algebra R A] [Ring B] [Algebra R B] {f : A ββ[R] B} {S : Subalgebra R A} (h : βf β»ΒΉ' {0} β βS) : Subalgebra.comap f (Subalgebra.map f S) = S - Algebra.map_top π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A ββ[R] B) : Subalgebra.map f β€ = f.range - Algebra.map_inf π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A ββ[R] B) (hf : Function.Injective βf) (S T : Subalgebra R A) : Subalgebra.map f (S β T) = Subalgebra.map f S β Subalgebra.map f T - Subalgebra.comap_map_eq π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Ring A] [Algebra R A] [Ring B] [Algebra R B] (f : A ββ[R] B) (S : Subalgebra R A) : Subalgebra.comap f (Subalgebra.map f S) = S β Algebra.adjoin R (βf β»ΒΉ' {0}) - Algebra.map_bot π Mathlib.Algebra.Algebra.Subalgebra.Lattice
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A ββ[R] B) : Subalgebra.map f β₯ = β₯ - Subalgebra.centralizer_coe_map_includeRight_eq_center_tensorProduct π Mathlib.Algebra.Algebra.Subalgebra.Centralizer
(R : Type u_1) [CommSemiring R] (A : Type u_2) [Semiring A] [Algebra R A] (B : Type u_3) [Semiring B] [Algebra R B] (S : Subalgebra R B) [Module.Free R A] : Subalgebra.centralizer R β(Subalgebra.map Algebra.TensorProduct.includeRight S) = (Algebra.TensorProduct.map (AlgHom.id R A) (Subalgebra.centralizer R βS).val).range - Subalgebra.centralizer_coe_map_includeLeft_eq_center_tensorProduct π Mathlib.Algebra.Algebra.Subalgebra.Centralizer
(R : Type u_1) [CommSemiring R] (A : Type u_2) [Semiring A] [Algebra R A] (B : Type u_3) [Semiring B] [Algebra R B] (S : Subalgebra R A) [Module.Free R B] : Subalgebra.centralizer R β(Subalgebra.map Algebra.TensorProduct.includeLeft S) = (Algebra.TensorProduct.map (Subalgebra.centralizer R βS).val (AlgHom.id R B)).range - Algebra.adjoin_algebraMap π Mathlib.RingTheory.Adjoin.Basic
(R : Type uR) {S : Type uS} (A : Type uA) [CommSemiring R] [CommSemiring S] [Semiring A] [Algebra R S] [Algebra R A] [Algebra S A] [IsScalarTower R S A] (s : Set S) : Algebra.adjoin R (β(algebraMap S A) '' s) = Subalgebra.map (IsScalarTower.toAlgHom R S A) (Algebra.adjoin R s) - Subalgebra.FG.map π Mathlib.RingTheory.Adjoin.FG
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {S : Subalgebra R A} (f : A ββ[R] B) (hs : S.FG) : (Subalgebra.map f S).FG - Subalgebra.fg_of_fg_map π Mathlib.RingTheory.Adjoin.FG
{R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (S : Subalgebra R A) (f : A ββ[R] B) (hf : Function.Injective βf) (hs : (Subalgebra.map f S).FG) : S.FG - Algebra.adjoin_restrictScalars π Mathlib.RingTheory.Adjoin.Tower
(C : Type u_1) (D : Type u_2) (E : Type u_3) [CommSemiring C] [CommSemiring D] [CommSemiring E] [Algebra C D] [Algebra C E] [Algebra D E] [IsScalarTower C D E] (S : Set E) : Subalgebra.restrictScalars C (Algebra.adjoin D S) = Subalgebra.restrictScalars C (Algebra.adjoin (β₯(Subalgebra.map (IsScalarTower.toAlgHom C D E) β€)) S) - StarSubalgebra.map_toSubalgebra π Mathlib.Algebra.Star.Subalgebra
{R : Type u_2} {A : Type u_3} {B : Type u_4} [CommSemiring R] [StarRing R] [Semiring A] [StarRing A] [Algebra R A] [StarModule R A] [Semiring B] [StarRing B] [Algebra R B] [StarModule R B] {S : StarSubalgebra R A} {f : A βββ[R] B} : (StarSubalgebra.map f S).toSubalgebra = Subalgebra.map f.toAlgHom S.toSubalgebra - Algebra.EssFiniteType.aux π Mathlib.RingTheory.EssentialFiniteness
(R : Type u_1) (S : Type u_2) (T : Type u_3) [CommRing R] [CommRing S] [CommRing T] [Algebra R S] [Algebra R T] [Algebra S T] [IsScalarTower R S T] (Ο : Subalgebra R S) (hΟ : β (s : S), β t β Ο, IsUnit t β§ s * t β Ο) (Ο : Set T) (t : T) (ht : t β Algebra.adjoin S Ο) : β s β Ο, IsUnit s β§ s β’ t β Subalgebra.map (IsScalarTower.toAlgHom R S T) Ο β Algebra.adjoin R Ο - integralClosure_map_algEquiv π Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Basic
{R : Type u_1} {A : Type u_2} {S : Type u_4} [CommRing R] [CommRing A] [CommRing S] [Algebra R A] [Algebra R S] (f : A ββ[R] S) : Subalgebra.map (βf) (integralClosure R A) = integralClosure R S - Algebra.TensorProduct.includeRight_map_center_le π Mathlib.Algebra.Central.TensorProduct
(K : Type u_1) (B : Type u_2) (C : Type u_3) [CommSemiring K] [Semiring B] [Semiring C] [Algebra K B] [Algebra K C] : Subalgebra.map Algebra.TensorProduct.includeRight (Subalgebra.center K C) β€ Subalgebra.center K (TensorProduct K B C) - Algebra.TensorProduct.includeLeft_map_center_le π Mathlib.Algebra.Central.TensorProduct
(K : Type u_1) (B : Type u_2) (C : Type u_3) [CommSemiring K] [Semiring B] [Semiring C] [Algebra K B] [Algebra K C] : Subalgebra.map Algebra.TensorProduct.includeLeft (Subalgebra.center K B) β€ Subalgebra.center K (TensorProduct K B C) - IntermediateField.toSubalgebra_map π Mathlib.FieldTheory.IntermediateField.Basic
{K : Type u_1} {L : Type u_2} {L' : Type u_3} [Field K] [Field L] [Field L'] [Algebra K L] [Algebra K L'] (S : IntermediateField K L) (f : L ββ[K] L') : (IntermediateField.map f S).toSubalgebra = Subalgebra.map f S.toSubalgebra - Subalgebra.topologicalClosure_map π Mathlib.Topology.Algebra.Algebra
{R : Type u_1} [CommSemiring R] {A : Type u_2} [Semiring A] [TopologicalSpace A] {B : Type u_3} [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] [IsTopologicalSemiring A] [IsTopologicalSemiring B] (f : A βA[R] B) (s : Subalgebra R A) : Subalgebra.map (βf) s.topologicalClosure β€ (Subalgebra.map f.toAlgHom s).topologicalClosure - DenseRange.topologicalClosure_map_subalgebra π Mathlib.Topology.Algebra.Algebra
{R : Type u_1} [CommSemiring R] {A : Type u_2} [Semiring A] [TopologicalSpace A] {B : Type u_3} [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] [IsTopologicalSemiring A] [IsTopologicalSemiring B] {f : A βA[R] B} (hf' : DenseRange βf) {s : Subalgebra R A} (hs : s.topologicalClosure = β€) : (Subalgebra.map (βf) s).topologicalClosure = β€ - polynomialFunctions.eq_1 π Mathlib.Topology.ContinuousMap.Polynomial
{R : Type u_1} [CommSemiring R] [TopologicalSpace R] [IsTopologicalSemiring R] (X : Set R) : polynomialFunctions X = Subalgebra.map (Polynomial.toContinuousMapOnAlgHom X) β€ - RCLike.restrict_toContinuousMap_eq_toContinuousMapStar_restrict π Mathlib.Analysis.RCLike.BoundedContinuous
(π : Type u_1) (E : Type u_2) [RCLike π] [PseudoEMetricSpace E] {A : StarSubalgebra π (BoundedContinuousFunction E π)} : Subalgebra.map (BoundedContinuousFunction.toContinuousMapβ β) (Subalgebra.comap (BoundedContinuousFunction.AlgHom.compLeftContinuousBounded β RCLike.ofRealAm β―) (Subalgebra.restrictScalars β A.toSubalgebra)) = Subalgebra.comap (AlgHom.compLeftContinuous β RCLike.ofRealAm β―) (Subalgebra.restrictScalars β (StarSubalgebra.map (BoundedContinuousFunction.toContinuousMapStarβ π) A).toSubalgebra) - dist_integral_mulExpNegMulSq_comp_le π Mathlib.Analysis.SpecialFunctions.MulExpNegMulSqIntegral
{Ξ΅ : β} {E : Type u_2} [MeasurableSpace E] [PseudoEMetricSpace E] [BorelSpace E] [CompleteSpace E] [SecondCountableTopology E] {P P' : MeasureTheory.Measure E} [MeasureTheory.IsFiniteMeasure P] [MeasureTheory.IsFiniteMeasure P'] (f : BoundedContinuousFunction E β) {A : Subalgebra β (BoundedContinuousFunction E β)} (hA : (Subalgebra.map (BoundedContinuousFunction.toContinuousMapβ β) A).SeparatesPoints) (heq : β g β A, β« (x : E), g x βP = β« (x : E), g x βP') (hΞ΅ : 0 < Ξ΅) : |β« (x : E), Ξ΅.mulExpNegMulSq (f x) βP - β« (x : E), Ξ΅.mulExpNegMulSq (f x) βP'| β€ 6 * βΞ΅ - Subalgebra.mulMap_map_comp_eq π Mathlib.LinearAlgebra.TensorProduct.Subalgebra
{R : Type u_1} {S : Type u_2} {T : Type u_3} [CommSemiring R] [CommSemiring S] [Algebra R S] [CommSemiring T] [Algebra R T] (A B : Subalgebra R S) (f : S ββ[R] T) : ((Subalgebra.map f A).mulMap (Subalgebra.map f B)).comp (Algebra.TensorProduct.map (AlgHom.subalgebraMap A f) (AlgHom.subalgebraMap B f)) = f.comp (A.mulMap B) - Subalgebra.LinearDisjoint.map π Mathlib.RingTheory.LinearDisjoint
{R : Type u} {S : Type v} [CommSemiring R] [Semiring S] [Algebra R S] {A B : Subalgebra R S} (H : A.LinearDisjoint B) {T : Type w} [Semiring T] [Algebra R T] (f : S ββ[R] T) (hf : Function.Injective βf) : (Subalgebra.map f A).LinearDisjoint (Subalgebra.map f B)
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 ff04530
serving mathlib revision 8623f65