diff --git a/Mathlib/Geometry/Manifold/ContMDiff.lean b/Mathlib/Geometry/Manifold/ContMDiff.lean index 281a580891e08a..cafd6bc3300462 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff.lean @@ -1790,6 +1790,8 @@ theorem ContinuousLinearMap.contMDiff (L : E β†’L[π•œ] F) : ContMDiff π“˜(π•œ L.contDiff.contMDiff #align continuous_linear_map.cont_mdiff ContinuousLinearMap.contMDiff +theorem ContinuousLinearMap.smooth (L : E β†’L[π•œ] F) : Smooth π“˜(π•œ, E) π“˜(π•œ, F) L := L.contMDiff + theorem ContMDiffWithinAt.clm_comp {g : M β†’ F₁ β†’L[π•œ] F₃} {f : M β†’ Fβ‚‚ β†’L[π•œ] F₁} {s : Set M} {x : M} (hg : ContMDiffWithinAt I π“˜(π•œ, F₁ β†’L[π•œ] F₃) n g s x) (hf : ContMDiffWithinAt I π“˜(π•œ, Fβ‚‚ β†’L[π•œ] F₁) n f s x) : diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean index c01907c3472154..de6cc195248319 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean @@ -52,11 +52,12 @@ theorem smoothOn_continuousLinearMapCoordChange [SmoothManifoldWithCorners IB B] SmoothOn IB π“˜(π•œ, (F₁ β†’L[π•œ] Fβ‚‚) β†’L[π•œ] F₁ β†’L[π•œ] Fβ‚‚) (continuousLinearMapCoordChange (RingHom.id π•œ) e₁ e₁' eβ‚‚ eβ‚‚') (e₁.baseSet ∩ eβ‚‚.baseSet ∩ (e₁'.baseSet ∩ eβ‚‚'.baseSet)) := by - let L₁ := compL π•œ F₁ Fβ‚‚ Fβ‚‚ - have h₁ : Smooth _ _ _ := L₁.cont_mdiff - have hβ‚‚ : Smooth _ _ _ := (ContinuousLinearMap.flip (compL π•œ F₁ F₁ Fβ‚‚)).ContMDiff - have h₃ : SmoothOn IB _ _ _ := smooth_on_coord_change e₁' e₁ - have hβ‚„ : SmoothOn IB _ _ _ := smooth_on_coord_change eβ‚‚ eβ‚‚' + set L₁ := compL π•œ F₁ Fβ‚‚ Fβ‚‚ + have h₁ : Smooth π“˜(π•œ, Fβ‚‚ β†’L[π•œ] Fβ‚‚) π“˜(π•œ, (F₁ β†’L[π•œ] Fβ‚‚) β†’L[π•œ] (F₁ β†’L[π•œ] Fβ‚‚)) L₁ := + L₁.smooth + have hβ‚‚ : Smooth _ _ _ := (ContinuousLinearMap.flip L₁).contMDiff + have h₃ : SmoothOn IB _ _ _ := smoothOn_coordChange e₁' e₁ + have hβ‚„ : SmoothOn IB _ _ _ := smoothOn_coordChange eβ‚‚ eβ‚‚' refine' ((h₁.comp_smooth_on (hβ‚„.mono _)).clm_comp (hβ‚‚.comp_smooth_on (h₃.mono _))).congr _ Β· mfld_set_tac Β· mfld_set_tac @@ -77,15 +78,15 @@ theorem hom_chart (yβ‚€ y : LE₁Eβ‚‚) : variable {IB} theorem contMDiffAt_hom_bundle (f : M β†’ LE₁Eβ‚‚) {xβ‚€ : M} {n : β„•βˆž} : - ContMDiffAt IM (IB.Prod π“˜(π•œ, F₁ β†’L[π•œ] Fβ‚‚)) n f xβ‚€ ↔ + ContMDiffAt IM (IB.prod π“˜(π•œ, F₁ β†’L[π•œ] Fβ‚‚)) n f xβ‚€ ↔ ContMDiffAt IM IB n (fun x => (f x).1) xβ‚€ ∧ ContMDiffAt IM π“˜(π•œ, F₁ β†’L[π•œ] Fβ‚‚) n (fun x => inCoordinates F₁ E₁ Fβ‚‚ Eβ‚‚ (f xβ‚€).1 (f x).1 (f xβ‚€).1 (f x).1 (f x).2) xβ‚€ := - by apply cont_mdiff_at_total_space + contMDiffAt_totalSpace .. #align cont_mdiff_at_hom_bundle contMDiffAt_hom_bundle theorem smoothAt_hom_bundle (f : M β†’ LE₁Eβ‚‚) {xβ‚€ : M} : - SmoothAt IM (IB.Prod π“˜(π•œ, F₁ β†’L[π•œ] Fβ‚‚)) f xβ‚€ ↔ + SmoothAt IM (IB.prod π“˜(π•œ, F₁ β†’L[π•œ] Fβ‚‚)) f xβ‚€ ↔ SmoothAt IM IB (fun x => (f x).1) xβ‚€ ∧ SmoothAt IM π“˜(π•œ, F₁ β†’L[π•œ] Fβ‚‚) (fun x => inCoordinates F₁ E₁ Fβ‚‚ Eβ‚‚ (f xβ‚€).1 (f x).1 (f xβ‚€).1 (f x).1 (f x).2) xβ‚€ := @@ -96,18 +97,16 @@ variable [SmoothManifoldWithCorners IB B] [SmoothVectorBundle F₁ E₁ IB] [SmoothVectorBundle Fβ‚‚ Eβ‚‚ IB] instance Bundle.ContinuousLinearMap.vectorPrebundle.isSmooth : - (Bundle.ContinuousLinearMap.vectorPrebundle (RingHom.id π•œ) F₁ E₁ Fβ‚‚ Eβ‚‚).IsSmooth IB - where exists_smooth_coord_change := by + (Bundle.ContinuousLinearMap.vectorPrebundle (RingHom.id π•œ) F₁ E₁ Fβ‚‚ Eβ‚‚).IsSmooth IB where + exists_smoothCoordChange := by rintro _ ⟨e₁, eβ‚‚, he₁, heβ‚‚, rfl⟩ _ ⟨e₁', eβ‚‚', he₁', heβ‚‚', rfl⟩ - skip - refine' - ⟨continuous_linear_map_coord_change (RingHom.id π•œ) e₁ e₁' eβ‚‚ eβ‚‚', - smoothOn_continuousLinearMapCoordChange IB, - continuous_linear_map_coord_change_apply (RingHom.id π•œ) e₁ e₁' eβ‚‚ eβ‚‚'⟩ + exact ⟨continuousLinearMapCoordChange (RingHom.id π•œ) e₁ e₁' eβ‚‚ eβ‚‚', + smoothOn_continuousLinearMapCoordChange IB, + continuousLinearMapCoordChange_apply (RingHom.id π•œ) e₁ e₁' eβ‚‚ eβ‚‚'⟩ #align bundle.continuous_linear_map.vector_prebundle.is_smooth Bundle.ContinuousLinearMap.vectorPrebundle.isSmooth instance SmoothVectorBundle.continuousLinearMap : SmoothVectorBundle (F₁ β†’L[π•œ] Fβ‚‚) (Bundle.ContinuousLinearMap (RingHom.id π•œ) E₁ Eβ‚‚) IB := - (Bundle.ContinuousLinearMap.vectorPrebundle (RingHom.id π•œ) F₁ E₁ Fβ‚‚ Eβ‚‚).SmoothVectorBundle IB + (Bundle.ContinuousLinearMap.vectorPrebundle (RingHom.id π•œ) F₁ E₁ Fβ‚‚ Eβ‚‚).smoothVectorBundle IB #align smooth_vector_bundle.continuous_linear_map SmoothVectorBundle.continuousLinearMap