Skip to content

Commit

Permalink
chore: replace remaining lambda syntax (#11405)
Browse files Browse the repository at this point in the history
Includes some doc comments and real code: this is exhaustive, with two exceptions:
- some files are handled in #11409 instead
- I left FunProp/{ToStd,RefinedDiscTree}.lean, Tactic/NormNum and Tactic/Simps alone, as these seem likely enough to end up in std.

Follow-up to #11301, much shorter this time.
  • Loading branch information
grunweg authored and utensil committed Mar 26, 2024
1 parent 476ea3e commit 9c12d8b
Show file tree
Hide file tree
Showing 10 changed files with 20 additions and 21 deletions.
4 changes: 2 additions & 2 deletions Mathlib/CategoryTheory/Bicategory/Functor.lean
Original file line number Diff line number Diff line change
Expand Up @@ -33,8 +33,8 @@ pseudofunctors can be defined by using the composition of oplax functors as foll
```lean
def comp (F : Pseudofunctor B C) (G : Pseudofunctor C D) : Pseudofunctor B D :=
mkOfOplax ((F : OplaxFunctor B C).comp G)
{ mapIdIso := λ a => (G.mapFunctor _ _).mapIso (F.mapId a) ≪≫ G.mapId (F.obj a),
mapCompIso := λ f g =>
{ mapIdIso := fun a (G.mapFunctor _ _).mapIso (F.mapId a) ≪≫ G.mapId (F.obj a),
mapCompIso := fun f g
(G.mapFunctor _ _).mapIso (F.mapComp f g) ≪≫ G.mapComp (F.map f) (F.map g) }
```
although the composition of pseudofunctors in this file is defined by using the default constructor
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/CategoryTheory/Limits/Shapes/WideEqualizers.lean
Original file line number Diff line number Diff line change
Expand Up @@ -429,7 +429,7 @@ theorem Cotrident.IsColimit.homIso_natural [Nonempty J] {t : Cotrident f} {Z Z'

/-- This is a helper construction that can be useful when verifying that a category has certain wide
equalizers. Given `F : WalkingParallelFamily ⥤ C`, which is really the same as
`parallelFamily (λ j, F.map (line j))`, and a trident on `fun j ↦ F.map (line j)`,
`parallelFamily (fun j ↦ F.map (line j))`, and a trident on `fun j ↦ F.map (line j)`,
we get a cone on `F`.
If you're thinking about using this, have a look at
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/CategoryTheory/Limits/VanKampen.lean
Original file line number Diff line number Diff line change
Expand Up @@ -339,7 +339,7 @@ theorem IsUniversalColimit.map_reflective
let c'' : Cocone (F' ⋙ Gr) := by
refine
{ pt := pullback (Gr.map f) (adj.unit.app _)
ι := { app := λ j ↦ pullback.lift (Gr.map <| c'.ι.app j) (Gr.map (α'.app j) ≫ c.ι.app j) ?_
ι := { app := fun j ↦ pullback.lift (Gr.map <| c'.ι.app j) (Gr.map (α'.app j) ≫ c.ι.app j) ?_
naturality := ?_ } }
· rw [← Gr.map_comp, ← hc'']
erw [← adj.unit_naturality]
Expand Down
8 changes: 4 additions & 4 deletions Mathlib/CategoryTheory/Products/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -384,11 +384,11 @@ open Opposite
@[simps]
def prodOpEquiv : (C × D)ᵒᵖ ≌ Cᵒᵖ × Dᵒᵖ where
functor :=
{ obj := λ X => ⟨op X.unop.1, op X.unop.2⟩,
map := λ f => ⟨f.unop.1.op, f.unop.2.op⟩ }
{ obj := fun X ⟨op X.unop.1, op X.unop.2⟩,
map := fun f ⟨f.unop.1.op, f.unop.2.op⟩ }
inverse :=
{ obj := λ ⟨X,Y⟩ => op ⟨X.unop, Y.unop⟩,
map := λ ⟨f,g⟩ => op ⟨f.unop, g.unop⟩ }
{ obj := fun ⟨X,Y⟩ op ⟨X.unop, Y.unop⟩,
map := fun ⟨f,g⟩ op ⟨f.unop, g.unop⟩ }
unitIso := Iso.refl _
counitIso := Iso.refl _
functor_unitIso_comp := fun ⟨X, Y⟩ => by
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/CategoryTheory/Sites/CoverLifting.lean
Original file line number Diff line number Diff line change
Expand Up @@ -258,7 +258,7 @@ theorem gluedSection_isAmalgamation : x.IsAmalgamation (gluedSection ℱ hS hx)
-- Porting note: next line was `ext W`
-- Now `ext` can't see that `ran` is defined as a limit.
-- See https://github.com/leanprover-community/mathlib4/issues/5229
refine limit.hom_ext (λ (W : StructuredArrow (op V) G.op) => ?_)
refine limit.hom_ext (fun (W : StructuredArrow (op V) G.op) ?_)
simp only [Functor.comp_map, limit.lift_pre, coyoneda_obj_map, ran_obj_map, gluedSection]
erw [limit.lift_π]
symm
Expand All @@ -275,7 +275,7 @@ theorem gluedSection_is_unique (y) (hy : x.IsAmalgamation y) : y = gluedSection
-- Porting note: next line was `ext W`
-- Now `ext` can't see that `ran` is defined as a limit.
-- See https://github.com/leanprover-community/mathlib4/issues/5229
refine limit.hom_ext (λ (W : StructuredArrow (op U) G.op) => ?_)
refine limit.hom_ext (fun (W : StructuredArrow (op U) G.op) ?_)
erw [limit.lift_π]
convert helper ℱ hS hx (𝟙 _) y W _
· simp only [op_id, StructuredArrow.map_id]
Expand Down
7 changes: 4 additions & 3 deletions Mathlib/Control/Fold.lean
Original file line number Diff line number Diff line change
Expand Up @@ -88,10 +88,11 @@ We can view the above as a composition of functions:
We can use traverse and const to construct this composition:
```
calc const.run (traverse (λ y, const.mk' (flip f y)) [y₀,y₁]) x
= const.run ((::) <$> const.mk' (flip f y₀) <*> traverse (λ y, const.mk' (flip f y)) [y₁]) x
calc const.run (traverse (fun y ↦ const.mk' (flip f y)) [y₀,y₁]) x
= const.run ((::) <$> const.mk' (flip f y₀) <*>
traverse (fun y ↦ const.mk' (flip f y)) [y₁]) x
... = const.run ((::) <$> const.mk' (flip f y₀) <*>
( (::) <$> const.mk' (flip f y₁) <*> traverse (λ y, const.mk' (flip f y)) [] )) x
( (::) <$> const.mk' (flip f y₁) <*> traverse (fun y ↦ const.mk' (flip f y)) [] )) x
... = const.run ((::) <$> const.mk' (flip f y₀) <*>
( (::) <$> const.mk' (flip f y₁) <*> pure [] )) x
... = const.run ( ((::) <$> const.mk' (flip f y₁) <*> pure []) ∘
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Tactic/Linarith/Verification.lean
Original file line number Diff line number Diff line change
Expand Up @@ -109,7 +109,7 @@ def mkLTZeroProof : List (Expr × ℕ) → MetaM Expr
return t
| ((h, c)::t) => do
let (iq, h') ← mkSingleCompZeroOf c h
let (_, t) ← t.foldlM (λ pr ce => step pr.1 pr.2 ce.1 ce.2) (iq, h')
let (_, t) ← t.foldlM (fun pr ce step pr.1 pr.2 ce.1 ce.2) (iq, h')
return t
where
/--
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Testing/SlimCheck/Testable.lean
Original file line number Diff line number Diff line change
Expand Up @@ -47,9 +47,9 @@ of `Shrinkable MyType` and `SampleableExt MyType`. We can define one as follows:
```lean
instance : Shrinkable MyType where
shrink := λ ⟨x,y,h⟩ =>
shrink := fun ⟨x,y,h⟩
let proxy := Shrinkable.shrink (x, y - x)
proxy.map (λ ⟨⟨fst, snd⟩, ha⟩ => ⟨⟨fst, fst + snd, sorry⟩, sorry⟩)
proxy.map (fun ⟨⟨fst, snd⟩, ha⟩ ⟨⟨fst, fst + snd, sorry⟩, sorry⟩)
instance : SampleableExt MyType :=
SampleableExt.mkSelfContained do
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Util/Tactic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -33,7 +33,7 @@ If `mvarId` does not refer to a declared metavariable, nothing happens.
-/
def modifyMetavarDecl [MonadMCtx m] (mvarId : MVarId)
(f : MetavarDecl → MetavarDecl) : m Unit := do
modifyMCtx λ mctx =>
modifyMCtx fun mctx
match mctx.decls.find? mvarId with
| none => mctx
| some mdecl => { mctx with decls := mctx.decls.insert mvarId (f mdecl) }
Expand Down
6 changes: 2 additions & 4 deletions test/apply_fun.lean
Original file line number Diff line number Diff line change
Expand Up @@ -142,11 +142,9 @@ example (a b : ℕ) (h : a ≠ b) : a + 1 ≠ b + 1 := by

-- TODO
-- -- monotonicity will be proved by `mono` in the next example
-- example (a b : ℕ) (h : a ≤ b) : a + 1 ≤ b + 1 :=
-- begin
-- apply_fun (λ n, n+1) at h,
-- example (a b : ℕ) (h : a ≤ b) : a + 1 ≤ b + 1 := by
-- apply_fun (fun n ↦ n+1) at h
-- exact h
-- end

example {n : Type} [Fintype n] {X : Type} [Semiring X]
(f : Matrix n n X → Matrix n n X) (A B : Matrix n n X) (h : A * B = 0) : f (A * B) = f 0 := by
Expand Down

0 comments on commit 9c12d8b

Please sign in to comment.