Skip to content

Commit

Permalink
chore: substitute some . with · (#12137)
Browse files Browse the repository at this point in the history
Co-authored-by: Moritz Firsching <[email protected]>
  • Loading branch information
2 people authored and atarnoam committed Apr 16, 2024
1 parent 355146c commit 59ca71c
Show file tree
Hide file tree
Showing 13 changed files with 27 additions and 27 deletions.
2 changes: 1 addition & 1 deletion Mathlib/Algebra/GroupPower/Order.lean
Original file line number Diff line number Diff line change
Expand Up @@ -532,7 +532,7 @@ namespace Nat
variable {n : ℕ} {f : α → ℕ}

/-- See also `pow_left_strictMonoOn`. -/
protected lemma pow_left_strictMono (hn : n ≠ 0) : StrictMono (. ^ n : ℕ → ℕ) :=
protected lemma pow_left_strictMono (hn : n ≠ 0) : StrictMono (· ^ n : ℕ → ℕ) :=
fun _ _ h ↦ Nat.pow_lt_pow_left h hn
#align nat.pow_left_strict_mono Nat.pow_left_strictMono

Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Algebra/Polynomial/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -299,7 +299,7 @@ instance semiring : Semiring R[X] :=
toMul := Polynomial.mul'
toZero := Polynomial.zero
toOne := Polynomial.one
nsmul := (..)
nsmul := (··)
npow := fun n x => (x ^ n) }
#align polynomial.semiring Polynomial.semiring

Expand Down Expand Up @@ -1190,7 +1190,7 @@ instance ring : Ring R[X] :=
toIntCast := Polynomial.intCast
toNeg := Polynomial.neg'
toSub := Polynomial.sub
zsmul := ((..) : ℤ → R[X] → R[X]) }
zsmul := ((··) : ℤ → R[X] → R[X]) }
#align polynomial.ring Polynomial.ring

@[simp]
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/Fin/Tuple/Curry.lean
Original file line number Diff line number Diff line change
Expand Up @@ -73,7 +73,7 @@ variable {n : ℕ} {p : Fin n → Type u} {τ : Type u}
theorem curry_uncurry (f : Function.FromTypes p τ) : curry (uncurry f) = f := by
induction n with
| zero => rfl
| succ n ih => exact funext (ih $ f .)
| succ n ih => exact funext (ih $ f ·)

@[simp]
theorem uncurry_curry (f : ((i : Fin n) → p i) → τ) :
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Data/Pi/Lex.lean
Original file line number Diff line number Diff line change
Expand Up @@ -119,15 +119,15 @@ open Function

theorem toLex_monotone : Monotone (@toLex (∀ i, β i)) := fun a b h =>
or_iff_not_imp_left.2 fun hne =>
let ⟨i, hi, hl⟩ := IsWellFounded.wf.has_min (r := (· < · )) { i | a i ≠ b i }
let ⟨i, hi, hl⟩ := IsWellFounded.wf.has_min (r := (· < ·)) { i | a i ≠ b i }
(Function.ne_iff.1 hne)
⟨i, fun j hj => by
contrapose! hl
exact ⟨j, hl, hj⟩, (h i).lt_of_ne hi⟩
#align pi.to_lex_monotone Pi.toLex_monotone

theorem toLex_strictMono : StrictMono (@toLex (∀ i, β i)) := fun a b h =>
let ⟨i, hi, hl⟩ := IsWellFounded.wf.has_min (r := (· < · )) { i | a i ≠ b i }
let ⟨i, hi, hl⟩ := IsWellFounded.wf.has_min (r := (· < ·)) { i | a i ≠ b i }
(Function.ne_iff.1 h.ne)
⟨i, fun j hj => by
contrapose! hl
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/Set/Subset.lean
Original file line number Diff line number Diff line change
Expand Up @@ -74,7 +74,7 @@ lemma preimage_val_sInter : A ↓∩ (⋂₀ S) = ⋂₀ { (A ↓∩ B) | B ∈
erw [sInter_image]
simp_rw [sInter_eq_biInter, preimage_iInter]

lemma preimage_val_sInter_eq_sInter : A ↓∩ (⋂₀ S) = ⋂₀ ((A ↓∩ .) '' S) := by
lemma preimage_val_sInter_eq_sInter : A ↓∩ (⋂₀ S) = ⋂₀ ((A ↓∩ ·) '' S) := by
simp only [preimage_sInter, sInter_image]

lemma eq_of_preimage_val_eq_of_subset (hB : B ⊆ A) (hC : C ⊆ A) (h : A ↓∩ B = A ↓∩ C) : B = C := by
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/GroupTheory/Exponent.lean
Original file line number Diff line number Diff line change
Expand Up @@ -93,7 +93,7 @@ open MulOpposite in
theorem _root_.MulOpposite.exponent : exponent (MulOpposite G) = exponent G := by
simp only [Monoid.exponent, ExponentExists]
congr!
all_goals exact ⟨(op_injective <| · <| op ·), (unop_injective <| . <| unop .)⟩
all_goals exact ⟨(op_injective <| · <| op ·), (unop_injective <| · <| unop ·)⟩

@[to_additive]
theorem ExponentExists.isOfFinOrder (h : ExponentExists G) {g : G} : IsOfFinOrder g :=
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/GroupTheory/HNNExtension.lean
Original file line number Diff line number Diff line change
Expand Up @@ -591,7 +591,7 @@ open NormalWord
theorem of_injective : Function.Injective (of : G → HNNExtension G A B φ) := by
rcases TransversalPair.nonempty G A B with ⟨d⟩
refine Function.Injective.of_comp
(f := ((..) : HNNExtension G A B φ → NormalWord d → NormalWord d)) ?_
(f := ((··) : HNNExtension G A B φ → NormalWord d → NormalWord d)) ?_
intros _ _ h
exact eq_of_smul_eq_smul (fun w : NormalWord d =>
by simp_all [Function.funext_iff, of_smul_eq_smul])
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/GroupTheory/PushoutI.lean
Original file line number Diff line number Diff line change
Expand Up @@ -596,7 +596,7 @@ theorem of_injective (hφ : ∀ i, Function.Injective (φ i)) (i : ι) :
let _ := Classical.decEq ι
let _ := fun i => Classical.decEq (G i)
refine Function.Injective.of_comp
(f := ((..) : PushoutI φ → NormalWord d → NormalWord d)) ?_
(f := ((··) : PushoutI φ → NormalWord d → NormalWord d)) ?_
intros _ _ h
exact eq_of_smul_eq_smul (fun w : NormalWord d =>
by simp_all [Function.funext_iff, of_smul_eq_smul])
Expand All @@ -607,7 +607,7 @@ theorem base_injective (hφ : ∀ i, Function.Injective (φ i)) :
let _ := Classical.decEq ι
let _ := fun i => Classical.decEq (G i)
refine Function.Injective.of_comp
(f := ((..) : PushoutI φ → NormalWord d → NormalWord d)) ?_
(f := ((··) : PushoutI φ → NormalWord d → NormalWord d)) ?_
intros _ _ h
exact eq_of_smul_eq_smul (fun w : NormalWord d =>
by simp_all [Function.funext_iff, base_smul_eq_smul])
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/ModelTheory/Definability.lean
Original file line number Diff line number Diff line change
Expand Up @@ -60,7 +60,7 @@ theorem Definable.map_expansion {L' : FirstOrder.Language} [L'.Structure M] (h :
theorem definable_iff_exists_formula_sum :
A.Definable L s ↔ ∃ φ : L.Formula (A ⊕ α), s = {v | φ.Realize (Sum.elim (↑) v)} := by
rw [Definable, Equiv.exists_congr_left (BoundedFormula.constantsVarsEquiv)]
refine exists_congr (fun φ => iff_iff_eq.2 (congr_arg (s = .) ?_))
refine exists_congr (fun φ => iff_iff_eq.2 (congr_arg (s = ·) ?_))
ext
simp only [Formula.Realize, BoundedFormula.constantsVarsEquiv, constantsOn, mk₂_Relations,
BoundedFormula.mapTermRelEquiv_symm_apply, mem_setOf_eq]
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/ModelTheory/Syntax.lean
Original file line number Diff line number Diff line change
Expand Up @@ -667,11 +667,11 @@ def toFormula : ∀ {n : ℕ}, L.BoundedFormula α n → L.Formula (Sum α (Fin

/-- take the disjunction of a finite set of formulas -/
noncomputable def iSup (s : Finset β) (f : β → L.BoundedFormula α n) : L.BoundedFormula α n :=
(s.toList.map f).foldr (..) ⊥
(s.toList.map f).foldr (··) ⊥

/-- take the conjunction of a finite set of formulas -/
noncomputable def iInf (s : Finset β) (f : β → L.BoundedFormula α n) : L.BoundedFormula α n :=
(s.toList.map f).foldr (..) ⊤
(s.toList.map f).foldr (··) ⊤


variable {l : ℕ} {φ ψ : L.BoundedFormula α l} {θ : L.BoundedFormula α l.succ}
Expand Down
14 changes: 7 additions & 7 deletions Mathlib/Order/CompletePartialOrder.lean
Original file line number Diff line number Diff line change
Expand Up @@ -39,31 +39,31 @@ Complete partial orders are partial orders where every directed set has a least
-/
class CompletePartialOrder (α : Type*) extends PartialOrder α, SupSet α where
/-- For each directed set `d`, `sSup d` is the least upper bound of `d`. -/
lubOfDirected : ∀ d, DirectedOn (..) d → IsLUB d (sSup d)
lubOfDirected : ∀ d, DirectedOn (··) d → IsLUB d (sSup d)

variable [CompletePartialOrder α] [Preorder β] {f : ι → α} {d : Set α} {a : α}

protected lemma DirectedOn.isLUB_sSup : DirectedOn (..) d → IsLUB d (sSup d) :=
protected lemma DirectedOn.isLUB_sSup : DirectedOn (··) d → IsLUB d (sSup d) :=
CompletePartialOrder.lubOfDirected _

protected lemma DirectedOn.le_sSup (hd : DirectedOn (..) d) (ha : a ∈ d) : a ≤ sSup d :=
protected lemma DirectedOn.le_sSup (hd : DirectedOn (··) d) (ha : a ∈ d) : a ≤ sSup d :=
hd.isLUB_sSup.1 ha

protected lemma DirectedOn.sSup_le (hd : DirectedOn (..) d) (ha : ∀ b ∈ d, b ≤ a) : sSup d ≤ a :=
protected lemma DirectedOn.sSup_le (hd : DirectedOn (··) d) (ha : ∀ b ∈ d, b ≤ a) : sSup d ≤ a :=
hd.isLUB_sSup.2 ha

protected lemma Directed.le_iSup (hf : Directed (..) f) (i : ι) : f i ≤ ⨆ j, f j :=
protected lemma Directed.le_iSup (hf : Directed (··) f) (i : ι) : f i ≤ ⨆ j, f j :=
hf.directedOn_range.le_sSup <| Set.mem_range_self _

protected lemma Directed.iSup_le (hf : Directed (..) f) (ha : ∀ i, f i ≤ a) : ⨆ i, f i ≤ a :=
protected lemma Directed.iSup_le (hf : Directed (··) f) (ha : ∀ i, f i ≤ a) : ⨆ i, f i ≤ a :=
hf.directedOn_range.sSup_le <| Set.forall_mem_range.2 ha

--TODO: We could mimic more `sSup`/`iSup` lemmas

/-- Scott-continuity takes on a simpler form in complete partial orders. -/
lemma CompletePartialOrder.scottContinuous {f : α → β} :
ScottContinuous f ↔
∀ ⦃d : Set α⦄, d.Nonempty → DirectedOn (..) d → IsLUB (f '' d) (f (sSup d)) := by
∀ ⦃d : Set α⦄, d.Nonempty → DirectedOn (··) d → IsLUB (f '' d) (f (sSup d)) := by
refine' ⟨fun h d hd₁ hd₂ ↦ h hd₁ hd₂ hd₂.isLUB_sSup, fun h d hne hd a hda ↦ ?_⟩
rw [hda.unique hd.isLUB_sSup]
exact h hne hd
Expand Down
8 changes: 4 additions & 4 deletions Mathlib/Order/RelSeries.lean
Original file line number Diff line number Diff line change
Expand Up @@ -256,7 +256,7 @@ For two types `α, β` and relation on them `r, s`, if `f : α → β` preserves
def map (p : RelSeries r) (f : r →r s) : RelSeries s where
length := p.length
toFun := f.1.comp p
step := (f.2 <| p.step .)
step := (f.2 <| p.step ·)

@[simp] lemma map_apply (p : RelSeries r) (f : r →r s) (i : Fin (p.length + 1)) :
p.map f i = f (p i) := rfl
Expand Down Expand Up @@ -479,7 +479,7 @@ end RelSeries

/-- A type is finite dimensional if its `LTSeries` has bounded length. -/
abbrev FiniteDimensionalOrder (γ : Type*) [Preorder γ] :=
Rel.FiniteDimensional ((. < .) : γ → γ → Prop)
Rel.FiniteDimensional ((· < ·) : γ → γ → Prop)

instance FiniteDimensionalOrder.ofUnique (γ : Type*) [Preorder γ] [Unique γ] :
FiniteDimensionalOrder γ where
Expand All @@ -489,15 +489,15 @@ instance FiniteDimensionalOrder.ofUnique (γ : Type*) [Preorder γ] [Unique γ]

/-- A type is infinite dimensional if it has `LTSeries` of at least arbitrary length -/
abbrev InfiniteDimensionalOrder (γ : Type*) [Preorder γ] :=
Rel.InfiniteDimensional ((. < .) : γ → γ → Prop)
Rel.InfiniteDimensional ((· < ·) : γ → γ → Prop)

section LTSeries

variable (α) [Preorder α] [Preorder β]
/--
If `α` is a preorder, a LTSeries is a relation series of the less than relation.
-/
abbrev LTSeries := RelSeries ((. < .) : Rel α α)
abbrev LTSeries := RelSeries ((· < ·) : Rel α α)

namespace LTSeries

Expand Down
4 changes: 2 additions & 2 deletions Mathlib/SetTheory/Ordinal/Arithmetic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2009,8 +2009,8 @@ def blsub₂ (o₁ o₂ : Ordinal) (op : {a : Ordinal} → (a < o₁) → {b : O
theorem lt_blsub₂ {o₁ o₂ : Ordinal}
(op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) {a b : Ordinal}
(ha : a < o₁) (hb : b < o₂) : op ha hb < blsub₂ o₁ o₂ op := by
convert lt_lsub _ (Prod.mk (enum (· < · ) a (by rwa [type_lt]))
(enum (· < · ) b (by rwa [type_lt])))
convert lt_lsub _ (Prod.mk (enum (· < ·) a (by rwa [type_lt]))
(enum (· < ·) b (by rwa [type_lt])))
simp only [typein_enum]
#align ordinal.lt_blsub₂ Ordinal.lt_blsub₂

Expand Down

0 comments on commit 59ca71c

Please sign in to comment.