Skip to content

Commit

Permalink
Remove unused tactics and some tidying around these changes.
Browse files Browse the repository at this point in the history
With incremental compilation, save is probably unnecessary.
  • Loading branch information
grunweg committed Jul 7, 2024
1 parent 0092b9e commit f469c7a
Show file tree
Hide file tree
Showing 5 changed files with 7 additions and 17 deletions.
8 changes: 2 additions & 6 deletions SphereEversion/Global/Gromov.lean
Original file line number Diff line number Diff line change
Expand Up @@ -40,9 +40,7 @@ theorem RelMfld.Ample.satisfiesHPrinciple' (hRample : R.Ample) (hRopen : IsOpen
clear! A
intro A hA 𝓕₀ h𝓕₀
cases' isEmpty_or_nonempty M with hM hM
· refine ⟨emptyHtpyFormalSol R, ?_, ?_, ?_, ?_⟩
all_goals try apply eventually_of_forall _
all_goals try intro
· refine ⟨emptyHtpyFormalSol R, ?_, ?_, ?_, ?_⟩ <;> intro
all_goals try intro
all_goals
first
Expand Down Expand Up @@ -194,9 +192,7 @@ theorem RelMfld.Ample.satisfiesHPrinciple (hRample : R.Ample) (hRopen : IsOpen R
clear! A
intro A hA 𝓕₀ h𝓕₀
cases' isEmpty_or_nonempty M with hM hM
· refine ⟨emptyHtpyFormalSol R, ?_, ?_, ?_, ?_⟩
all_goals try apply eventually_of_forall _
all_goals try intro
· refine ⟨emptyHtpyFormalSol R, ?_, ?_, ?_, ?_⟩ <;> intro
all_goals try intro
all_goals
first
Expand Down
1 change: 0 additions & 1 deletion SphereEversion/Global/Immersion.lean
Original file line number Diff line number Diff line change
Expand Up @@ -325,7 +325,6 @@ theorem sphere_eversion :
(formalEversion E ω) (formalEversion_hol_near_zero_one E ω) with
⟨f, h₁, h₂, -, h₅⟩
have := h₂.forall_mem principal_le_nhdsSet
save
refine ⟨f, h₁, ?_, ?_, ?_/-h₅-/
· ext x
rw [this (0, x) (by simp)]
Expand Down
5 changes: 2 additions & 3 deletions SphereEversion/Global/SmoothEmbedding.lean
Original file line number Diff line number Diff line change
Expand Up @@ -424,9 +424,8 @@ theorem smooth_update (f : M' → M → N) (g : M' → X → Y) {k : M' → M} {
have h₃ : V ∪ U = univ := by
rw [← compl_subset_iff_union, compl_compl]
exact image_subset_range φ K
have h₄ : ∀ x, k x ∈ U → update φ ψ (f x) (g x) (k x) = (ψ ∘ g x ∘ φ.invFun) (k x) := fun m hm ↦ by
intros
exact if_pos hm
have h₄ (x) : k x ∈ U → update φ ψ (f x) (g x) (k x) = (ψ ∘ g x ∘ φ.invFun) (k x) :=
fun hm ↦ if_pos hm
by_cases hx : k x ∈ U
· exact ⟨k ⁻¹' U, φ.isOpen_range.preimage hk.continuous, hx,
(contMDiffOn_congr h₄).mpr <| ψ.smooth_to.comp_contMDiffOn <| hg.comp_contMDiffOn
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -221,12 +221,9 @@ theorem hasFDerivAt_parametric_primitive_of_contDiff' {F : H → ℝ → E} (hF
rw [mapsTo']
rintro ⟨x, s⟩ ⟨x', hx, h⟩; cases h
exact ⟨ball_subset_closedBall hx, mem_Icc_of_Ioo t_in⟩
have cont_x : ∀ x, Continuous (F x) := fun x ↦ hF.continuous.comp (Continuous.Prod.mk x)
have int_Icc : ∀ x, IntegrableOn (F x) (Icc a₀ b₀) := fun x ↦
(cont_x x).continuousOn.integrableOn_Icc
have int_Ioo : ∀ x, IntegrableOn (F x) (Ioo a₀ b₀) := by
exact fun x ↦ (int_Icc x).mono_set Ioo_subset_Icc_self
save
have cont_x (x) : Continuous (F x) := hF.continuous.comp (Continuous.Prod.mk x)
have int_Icc (x) : IntegrableOn (F x) (Icc a₀ b₀) := (cont_x x).continuousOn.integrableOn_Icc
have int_Ioo (x) : IntegrableOn (F x) (Ioo a₀ b₀) := (int_Icc x).mono_set Ioo_subset_Icc_self
apply
hasFDerivAt_parametric_primitive_of_lip' _ _ zero_lt_one ha ht₀
(fun x _hx ↦ (cont_x x).aestronglyMeasurable) (int_Ioo x₀) (cont_x x₀).continuousAt _ _ _
Expand Down
1 change: 0 additions & 1 deletion SphereEversion/ToMathlib/SmoothBarycentric.lean
Original file line number Diff line number Diff line change
Expand Up @@ -139,7 +139,6 @@ theorem smooth_barycentric [DecidablePred (· ∈ affineBases ι 𝕜 F)] [Finit
refine contDiff_pi.mpr fun j ↦ contDiff_pi.mpr fun j' ↦ ?_
simp only [Matrix.updateRow_apply]
simp only [AffineBasis.toMatrix_apply, AffineBasis.coords_apply]
save
by_cases hij : j = i
· simp only [hij, if_true, eq_self_iff_true]
exact (smooth_barycentric_coord b j').fst'
Expand Down

0 comments on commit f469c7a

Please sign in to comment.