Skip to content

Commit

Permalink
fix some errors
Browse files Browse the repository at this point in the history
  • Loading branch information
JovanGerb committed May 18, 2024
1 parent f82b58a commit 4c9406c
Show file tree
Hide file tree
Showing 4 changed files with 9 additions and 6 deletions.
2 changes: 1 addition & 1 deletion Mathlib/Analysis/BoxIntegral/Partition/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -622,7 +622,7 @@ theorem filter_true : (π.filter fun _ => True) = π :=
theorem iUnion_filter_not (π : Prepartition I) (p : Box ι → Prop) :
(π.filter fun J => ¬p J).iUnion = π.iUnion \ (π.filter p).iUnion := by
simp only [Prepartition.iUnion]
convert (@Set.biUnion_diff_biUnion_eq _ (Box ι) π.boxes (π.filter p).boxes (↑) _).symm
convert (@Set.biUnion_diff_biUnion_eq _ (Box ι) π.boxes (π.filter p).boxes Box.toSet _).symm
· simp (config := { contextual := true })
· rw [Set.PairwiseDisjoint]
convert π.pairwiseDisjoint
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Analysis/Convex/Segment.lean
Original file line number Diff line number Diff line change
Expand Up @@ -337,12 +337,12 @@ theorem midpoint_mem_segment [Invertible (2 : 𝕜)] (x y : E) : midpoint 𝕜 x
#align midpoint_mem_segment midpoint_mem_segment

theorem mem_segment_sub_add [Invertible (2 : 𝕜)] (x y : E) : x ∈ [x - y -[𝕜] x + y] := by
convert @midpoint_mem_segment 𝕜 _ _ _ _ _ _ _
convert @midpoint_mem_segment 𝕜 E _ _ _ _ _ _
rw [midpoint_sub_add]
#align mem_segment_sub_add mem_segment_sub_add

theorem mem_segment_add_sub [Invertible (2 : 𝕜)] (x y : E) : x ∈ [x + y -[𝕜] x - y] := by
convert @midpoint_mem_segment 𝕜 _ _ _ _ _ _ _
convert @midpoint_mem_segment 𝕜 E _ _ _ _ _ _
rw [midpoint_add_sub]
#align mem_segment_add_sub mem_segment_add_sub

Expand Down
7 changes: 5 additions & 2 deletions Mathlib/CategoryTheory/Category/PartialFun.lean
Original file line number Diff line number Diff line change
Expand Up @@ -184,6 +184,9 @@ noncomputable def typeToPartialFunIsoPartialFunToPointed :
fun f =>
Pointed.Hom.ext _ _ <|
funext fun a => Option.recOn a rfl fun a => by
classical
convert Part.some_toOption _
simp

Check failure on line 187 in Mathlib/CategoryTheory/Category/PartialFun.lean

View workflow job for this annotation

GitHub Actions / Lint style

Mathlib/CategoryTheory/Category/PartialFun.lean:187 ERR_NSP: Non-terminal simp. Replace with `simp?` and use the suggested output
rw [← Part.get_eq_iff_mem]
· rfl
· trivial

#align Type_to_PartialFun_iso_PartialFun_to_Pointed typeToPartialFunIsoPartialFunToPointed
2 changes: 1 addition & 1 deletion Mathlib/NumberTheory/Padics/PadicVal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -768,7 +768,7 @@ digits of `k` plus the sum of the digits of `n - k` minus the sum of digits of `
theorem sub_one_mul_padicValNat_choose_eq_sub_sum_digits {k n : ℕ} [hp : Fact p.Prime]
(h : k ≤ n) : (p - 1) * padicValNat p (choose n k) =
(p.digits k).sum + (p.digits (n - k)).sum - (p.digits n).sum := by
convert @sub_one_mul_padicValNat_choose_eq_sub_sum_digits' _ _ _ _
convert @sub_one_mul_padicValNat_choose_eq_sub_sum_digits' p _ _ _
all_goals exact Nat.eq_add_of_sub_eq h rfl

end padicValNat
Expand Down

0 comments on commit 4c9406c

Please sign in to comment.