Loogle!
Result
Found 77 declarations mentioning Units and Subsingleton. Of these, 69 match your pattern(s).
- units_eq_one 📋 Mathlib.Algebra.Group.Units.Defs
{M : Type u_1} [Monoid M] [Subsingleton Mˣ] (u : Mˣ) : u = 1 - IsUnit.eq_one 📋 Mathlib.Algebra.Group.Units.Defs
{M : Type u_1} [Monoid M] {a : M} [Subsingleton Mˣ] : IsUnit a → a = 1 - Units.eq_one 📋 Mathlib.Algebra.Group.Units.Defs
{M : Type u_1} [Monoid M] [Subsingleton Mˣ] (u : Mˣ) : u = 1 - isUnit_iff_eq_one 📋 Mathlib.Algebra.Group.Units.Defs
{M : Type u_1} [Monoid M] {a : M} [Subsingleton Mˣ] : IsUnit a ↔ a = 1 - eq_one_of_mul_left 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [CommMonoid α] [Subsingleton αˣ] {a b : α} (h : a * b = 1) : b = 1 - eq_one_of_mul_right 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [CommMonoid α] [Subsingleton αˣ] {a b : α} (h : a * b = 1) : a = 1 - LeftCancelMonoid.eq_one_of_mul_left 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [LeftCancelMonoid α] [Subsingleton αˣ] {a b : α} (h : a * b = 1) : b = 1 - LeftCancelMonoid.eq_one_of_mul_right 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [LeftCancelMonoid α] [Subsingleton αˣ] {a b : α} (h : a * b = 1) : a = 1 - RightCancelMonoid.eq_one_of_mul_left 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [RightCancelMonoid α] [Subsingleton αˣ] {a b : α} (h : a * b = 1) : b = 1 - RightCancelMonoid.eq_one_of_mul_right 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [RightCancelMonoid α] [Subsingleton αˣ] {a b : α} (h : a * b = 1) : a = 1 - eq_one_of_mul_left' 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [CancelMonoid α] [Subsingleton αˣ] {a b : α} (h : a * b = 1) : b = 1 - eq_one_of_mul_right' 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [CancelMonoid α] [Subsingleton αˣ] {a b : α} (h : a * b = 1) : a = 1 - mul_eq_one 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [CommMonoid α] [Subsingleton αˣ] {a b : α} : a * b = 1 ↔ a = 1 ∧ b = 1 - mul_ne_one 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [CommMonoid α] [Subsingleton αˣ] {a b : α} : a * b ≠ 1 ↔ a ≠ 1 ∨ b ≠ 1 - LeftCancelMonoid.mul_eq_one 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [LeftCancelMonoid α] [Subsingleton αˣ] {a b : α} : a * b = 1 ↔ a = 1 ∧ b = 1 - LeftCancelMonoid.mul_ne_one 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [LeftCancelMonoid α] [Subsingleton αˣ] {a b : α} : a * b ≠ 1 ↔ a ≠ 1 ∨ b ≠ 1 - RightCancelMonoid.mul_eq_one 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [RightCancelMonoid α] [Subsingleton αˣ] {a b : α} : a * b = 1 ↔ a = 1 ∧ b = 1 - RightCancelMonoid.mul_ne_one 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [RightCancelMonoid α] [Subsingleton αˣ] {a b : α} : a * b ≠ 1 ↔ a ≠ 1 ∨ b ≠ 1 - mul_eq_one' 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [CancelMonoid α] [Subsingleton αˣ] {a b : α} : a * b = 1 ↔ a = 1 ∧ b = 1 - mul_ne_one' 📋 Mathlib.Algebra.Group.Units.Basic
{α : Type u} [CancelMonoid α] [Subsingleton αˣ] {a b : α} : a * b ≠ 1 ↔ a ≠ 1 ∨ b ≠ 1 - Prod.instSubsingletonUnits 📋 Mathlib.Algebra.Group.Prod
{M : Type u_3} {N : Type u_4} [Monoid M] [Monoid N] [Subsingleton Mˣ] [Subsingleton Nˣ] : Subsingleton (M × N)ˣ - dvd_antisymm 📋 Mathlib.Algebra.GroupWithZero.Divisibility
{α : Type u_1} [CancelCommMonoidWithZero α] {a b : α} [Subsingleton αˣ] : a ∣ b → b ∣ a → a = b - dvd_antisymm' 📋 Mathlib.Algebra.GroupWithZero.Divisibility
{α : Type u_1} [CancelCommMonoidWithZero α] {a b : α} [Subsingleton αˣ] : a ∣ b → b ∣ a → b = a - Dvd.dvd.antisymm 📋 Mathlib.Algebra.GroupWithZero.Divisibility
{α : Type u_1} [CancelCommMonoidWithZero α] {a b : α} [Subsingleton αˣ] : a ∣ b → b ∣ a → a = b - Dvd.dvd.antisymm' 📋 Mathlib.Algebra.GroupWithZero.Divisibility
{α : Type u_1} [CancelCommMonoidWithZero α] {a b : α} [Subsingleton αˣ] : a ∣ b → b ∣ a → b = a - eq_of_forall_dvd 📋 Mathlib.Algebra.GroupWithZero.Divisibility
{α : Type u_1} [CancelCommMonoidWithZero α] {a b : α} [Subsingleton αˣ] (h : ∀ (c : α), a ∣ c ↔ b ∣ c) : a = b - eq_of_forall_dvd' 📋 Mathlib.Algebra.GroupWithZero.Divisibility
{α : Type u_1} [CancelCommMonoidWithZero α] {a b : α} [Subsingleton αˣ] (h : ∀ (c : α), c ∣ a ↔ c ∣ b) : a = b - Finset.prod_eq_one_iff 📋 Mathlib.Algebra.BigOperators.Group.Finset.Basic
{ι : Type u_1} {M : Type u_4} {s : Finset ι} [CommMonoid M] {f : ι → M} [Subsingleton Mˣ] : ∏ i ∈ s, f i = 1 ↔ ∀ i ∈ s, f i = 1 - Irreducible.eq_one_or_eq_one 📋 Mathlib.Algebra.Group.Irreducible.Defs
{M : Type u_1} [Monoid M] {a b : M} [Subsingleton Mˣ] (hab : Irreducible (a * b)) : a = 1 ∨ b = 1 - Associates.mk_injective 📋 Mathlib.Algebra.GroupWithZero.Associated
{M : Type u_1} [Monoid M] [Subsingleton Mˣ] : Function.Injective Associates.mk - associated_eq_eq 📋 Mathlib.Algebra.GroupWithZero.Associated
{M : Type u_1} [Monoid M] [Subsingleton Mˣ] : Associated = Eq - associated_iff_eq 📋 Mathlib.Algebra.GroupWithZero.Associated
{M : Type u_1} [Monoid M] [Subsingleton Mˣ] {x y : M} : Associated x y ↔ x = y - prime_dvd_prime_iff_eq 📋 Mathlib.Algebra.GroupWithZero.Associated
{M : Type u_2} [CancelCommMonoidWithZero M] [Subsingleton Mˣ] {p q : M} (pp : Prime p) (qp : Prime q) : p ∣ q ↔ p = q - eq_of_prime_pow_eq 📋 Mathlib.Algebra.GroupWithZero.Associated
{R : Type u_2} [CancelCommMonoidWithZero R] [Subsingleton Rˣ] {p₁ p₂ : R} {k₁ k₂ : ℕ} (hp₁ : Prime p₁) (hp₂ : Prime p₂) (hk₁ : 0 < k₁) (h : p₁ ^ k₁ = p₂ ^ k₂) : p₁ = p₂ - eq_of_prime_pow_eq' 📋 Mathlib.Algebra.GroupWithZero.Associated
{R : Type u_2} [CancelCommMonoidWithZero R] [Subsingleton Rˣ] {p₁ p₂ : R} {k₁ k₂ : ℕ} (hp₁ : Prime p₁) (hp₂ : Prime p₂) (hk₁ : 0 < k₂) (h : p₁ ^ k₁ = p₂ ^ k₂) : p₁ = p₂ - NormalizationMonoid.ofUniqueUnits 📋 Mathlib.Algebra.GCDMonoid.Basic
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] : NormalizationMonoid α - subsingleton_gcdMonoid_of_unique_units 📋 Mathlib.Algebra.GCDMonoid.Basic
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] : Subsingleton (GCDMonoid α) - subsingleton_normalizedGCDMonoid_of_unique_units 📋 Mathlib.Algebra.GCDMonoid.Basic
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] : Subsingleton (NormalizedGCDMonoid α) - uniqueNormalizationMonoidOfUniqueUnits 📋 Mathlib.Algebra.GCDMonoid.Basic
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] : Unique (NormalizationMonoid α) - associatesEquivOfUniqueUnits 📋 Mathlib.Algebra.GCDMonoid.Basic
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] : Associates α ≃* α - normUnit_eq_one 📋 Mathlib.Algebra.GCDMonoid.Basic
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] (x : α) : normUnit x = 1 - normalize_eq 📋 Mathlib.Algebra.GCDMonoid.Basic
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] (x : α) : normalize x = x - exists_eq_pow_of_mul_eq_pow 📋 Mathlib.Algebra.GCDMonoid.Basic
{α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] [Subsingleton αˣ] {a b c : α} (hab : IsUnit (gcd a b)) {k : ℕ} (h : a * b = c ^ k) : ∃ d, a = d ^ k - associatesEquivOfUniqueUnits_apply 📋 Mathlib.Algebra.GCDMonoid.Basic
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] (a✝ : Associates α) : associatesEquivOfUniqueUnits a✝ = a✝.out - associatesEquivOfUniqueUnits_symm_apply 📋 Mathlib.Algebra.GCDMonoid.Basic
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] (a : α) : associatesEquivOfUniqueUnits.symm a = Associates.mk a - finite_irreducible 📋 Mathlib.Algebra.AffineMonoid.Irreducible
{M : Type u_1} [CommMonoid M] [Subsingleton Mˣ] [Monoid.FG M] : {p | Irreducible p}.Finite - irreducible_subset_of_submonoidClosure_eq_top 📋 Mathlib.Algebra.AffineMonoid.Irreducible
{M : Type u_1} [CommMonoid M] [Subsingleton Mˣ] {S : Set M} (hS : Submonoid.closure S = ⊤) : {p | Irreducible p} ⊆ S - irreducible_mem_submonoidClosure_subset 📋 Mathlib.Algebra.AffineMonoid.Irreducible
{M : Type u_1} [CommMonoid M] [Subsingleton Mˣ] {S : Set M} : {p | p ∈ Submonoid.closure S ∧ Irreducible p} ⊆ S - Submonoid.FG.finite_irreducible_mem_submonoidClosure 📋 Mathlib.Algebra.AffineMonoid.Irreducible
{M : Type u_1} [CommMonoid M] [Subsingleton Mˣ] {S : Submonoid M} (hS : S.FG) : {p | p ∈ S ∧ Irreducible p}.Finite - Submonoid.closure_irreducible 📋 Mathlib.Algebra.AffineMonoid.Irreducible
{M : Type u_1} [CancelCommMonoid M] [Subsingleton Mˣ] [Monoid.FG M] : Submonoid.closure {p | Irreducible p} = ⊤ - mem_list_primes_of_dvd_prod 📋 Mathlib.Data.List.Prime
{M : Type u_1} [CancelCommMonoidWithZero M] [Subsingleton Mˣ] {p : M} (hp : Prime p) {L : List M} (hL : ∀ q ∈ L, Prime q) (hpL : p ∣ L.prod) : p ∈ L - perm_of_prod_eq_prod 📋 Mathlib.Data.List.Prime
{M : Type u_1} [CancelCommMonoidWithZero M] [Subsingleton Mˣ] {l₁ l₂ : List M} : l₁.prod = l₂.prod → (∀ p ∈ l₁, Prime p) → (∀ p ∈ l₂, Prime p) → l₁.Perm l₂ - IsUnit.orderOf_eq_one 📋 Mathlib.GroupTheory.OrderOfElement
{G : Type u_1} [Monoid G] [Subsingleton Gˣ] {x : G} (h : IsUnit x) : orderOf x = 1 - Submonoid.instSubsingletonUnits 📋 Mathlib.Algebra.Group.Submonoid.Units
{M : Type u_1} [Monoid M] [Subsingleton Mˣ] {S : Submonoid M} : Subsingleton (↥S)ˣ - Finset.prod_primes_dvd 📋 Mathlib.Algebra.BigOperators.Associated
{M₀ : Type u_3} [CancelCommMonoidWithZero M₀] [Subsingleton M₀ˣ] {s : Finset M₀} (n : M₀) (h : ∀ a ∈ s, Prime a) (div : ∀ a ∈ s, a ∣ n) : ∏ p ∈ s, p ∣ n - UniqueFactorizationMonoid.factors_eq_normalizedFactors 📋 Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors
{M : Type u_2} [CancelCommMonoidWithZero M] [UniqueFactorizationMonoid M] [Subsingleton Mˣ] (x : M) : UniqueFactorizationMonoid.factors x = UniqueFactorizationMonoid.normalizedFactors x - UniqueFactorizationMonoid.normalizedFactors_prod_of_prime 📋 Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors
{α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] [UniqueFactorizationMonoid α] [Subsingleton αˣ] {m : Multiset α} (h : ∀ p ∈ m, Prime p) : UniqueFactorizationMonoid.normalizedFactors m.prod = m - UniqueFactorizationMonoid.mem_normalizedFactors_iff 📋 Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors
{α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] [UniqueFactorizationMonoid α] [Subsingleton αˣ] {p x : α} (hx : x ≠ 0) : p ∈ UniqueFactorizationMonoid.normalizedFactors x ↔ Prime p ∧ p ∣ x - exists_eq_pow_of_mul_eq_pow_of_coprime 📋 Mathlib.RingTheory.IntegralDomain
{R : Type u_2} [CommSemiring R] [IsDomain R] [GCDMonoid R] [Subsingleton Rˣ] {a b c : R} {n : ℕ} (cp : IsCoprime a b) (h : a * b = c ^ n) : ∃ d, a = d ^ n - Finset.exists_eq_pow_of_mul_eq_pow_of_coprime 📋 Mathlib.RingTheory.IntegralDomain
{ι : Type u_2} {R : Type u_3} [CommSemiring R] [IsDomain R] [GCDMonoid R] [Subsingleton Rˣ] {n : ℕ} {c : R} {s : Finset ι} {f : ι → R} (h : ∀ i ∈ s, ∀ j ∈ s, i ≠ j → IsCoprime (f i) (f j)) (hprod : ∏ i ∈ s, f i = c ^ n) (i : ι) : i ∈ s → ∃ d, f i = d ^ n - isMulIndecomposable_id_univ 📋 Mathlib.Algebra.Group.Irreducible.Indecomposable
{M : Type u_2} [Monoid M] [Subsingleton Mˣ] {x : M} (hx : x ≠ 1) : IsMulIndecomposable id Set.univ x ↔ Irreducible x - NormalizationMonoid.ofUniqueUnits.congr_simp 📋 Mathlib.RingTheory.ChainOfDivisors
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] : NormalizationMonoid.ofUniqueUnits = NormalizationMonoid.ofUniqueUnits - associatesEquivOfUniqueUnits.congr_simp 📋 Mathlib.RingTheory.ChainOfDivisors
{α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] : associatesEquivOfUniqueUnits = associatesEquivOfUniqueUnits - mkFactorOrderIsoOfFactorDvdEquiv 📋 Mathlib.RingTheory.ChainOfDivisors
{M : Type u_1} [CancelCommMonoidWithZero M] {N : Type u_2} [CancelCommMonoidWithZero N] [Subsingleton Mˣ] [Subsingleton Nˣ] {m : M} {n : N} {d : { l // l ∣ m } ≃ { l // l ∣ n }} (hd : ∀ (l l' : { l // l ∣ m }), ↑(d l) ∣ ↑(d l') ↔ ↑l ∣ ↑l') : ↑(Set.Iic (Associates.mk m)) ≃o ↑(Set.Iic (Associates.mk n)) - mem_normalizedFactors_factor_dvd_iso_of_mem_normalizedFactors 📋 Mathlib.RingTheory.ChainOfDivisors
{M : Type u_1} [CancelCommMonoidWithZero M] {N : Type u_2} [CancelCommMonoidWithZero N] [Subsingleton Mˣ] [Subsingleton Nˣ] [UniqueFactorizationMonoid M] [UniqueFactorizationMonoid N] {m p : M} {n : N} (hm : m ≠ 0) (hn : n ≠ 0) (hp : p ∈ UniqueFactorizationMonoid.normalizedFactors m) {d : { l // l ∣ m } ≃ { l // l ∣ n }} (hd : ∀ (l l' : { l // l ∣ m }), ↑(d l) ∣ ↑(d l') ↔ ↑l ∣ ↑l') : ↑(d ⟨p, ⋯⟩) ∈ UniqueFactorizationMonoid.normalizedFactors n - emultiplicity_factor_dvd_iso_eq_emultiplicity_of_mem_normalizedFactors 📋 Mathlib.RingTheory.ChainOfDivisors
{M : Type u_1} [CancelCommMonoidWithZero M] {N : Type u_2} [CancelCommMonoidWithZero N] [Subsingleton Mˣ] [Subsingleton Nˣ] [UniqueFactorizationMonoid M] [UniqueFactorizationMonoid N] {m p : M} {n : N} (hm : m ≠ 0) (hn : n ≠ 0) (hp : p ∈ UniqueFactorizationMonoid.normalizedFactors m) {d : { l // l ∣ m } ≃ { l // l ∣ n }} (hd : ∀ (l l' : { l // l ∣ m }), ↑(d l) ∣ ↑(d l') ↔ ↑l ∣ ↑l') : emultiplicity (↑(d ⟨p, ⋯⟩)) n = emultiplicity p m - mkFactorOrderIsoOfFactorDvdEquiv_apply_coe 📋 Mathlib.RingTheory.ChainOfDivisors
{M : Type u_1} [CancelCommMonoidWithZero M] {N : Type u_2} [CancelCommMonoidWithZero N] [Subsingleton Mˣ] [Subsingleton Nˣ] {m : M} {n : N} {d : { l // l ∣ m } ≃ { l // l ∣ n }} (hd : ∀ (l l' : { l // l ∣ m }), ↑(d l) ∣ ↑(d l') ↔ ↑l ∣ ↑l') (l : ↑(Set.Iic (Associates.mk m))) : ↑((mkFactorOrderIsoOfFactorDvdEquiv hd) l) = Associates.mk ↑(d ⟨associatesEquivOfUniqueUnits ↑l, ⋯⟩) - mkFactorOrderIsoOfFactorDvdEquiv_symm_apply_coe 📋 Mathlib.RingTheory.ChainOfDivisors
{M : Type u_1} [CancelCommMonoidWithZero M] {N : Type u_2} [CancelCommMonoidWithZero N] [Subsingleton Mˣ] [Subsingleton Nˣ] {m : M} {n : N} {d : { l // l ∣ m } ≃ { l // l ∣ n }} (hd : ∀ (l l' : { l // l ∣ m }), ↑(d l) ∣ ↑(d l') ↔ ↑l ∣ ↑l') (l : ↑(Set.Iic (Associates.mk n))) : ↑((RelIso.symm (mkFactorOrderIsoOfFactorDvdEquiv hd)) l) = Associates.mk ↑(d.symm ⟨associatesEquivOfUniqueUnits ↑l, ⋯⟩) - Valuation.Integers.bijective_algebraMap_of_subsingleton_units_mrange 📋 Mathlib.RingTheory.Valuation.Integers
{F : Type u} {Γ₀ : Type v} [Field F] [LinearOrderedCommGroupWithZero Γ₀] {v : Valuation F Γ₀} {O : Type w} [CommRing O] [Algebra O F] (hv : v.Integers O) [Subsingleton (↥(MonoidHom.mrange v))ˣ] : Function.Bijective ⇑(algebraMap O F)
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 ?aIf 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 ?bBy 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_tsumeven though the hypothesisf i < g iis 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 6ff4759 serving mathlib revision 519f454