working…
This commit is contained in:
		| @@ -219,6 +219,7 @@ theorem laplace_compCLMAt | |||||||
|     obtain ⟨u, hu₁, hu₂⟩ : ∃ u ∈ nhds x, ContDiffOn ℝ 2 f u := by |     obtain ⟨u, hu₁, hu₂⟩ : ∃ u ∈ nhds x, ContDiffOn ℝ 2 f u := by | ||||||
|       apply ContDiffAt.contDiffOn h |       apply ContDiffAt.contDiffOn h | ||||||
|       rfl |       rfl | ||||||
|  |       simp | ||||||
|     obtain ⟨v, hv₁, hv₂, hv₃⟩ := mem_nhds_iff.1 hu₁ |     obtain ⟨v, hv₁, hv₂, hv₃⟩ := mem_nhds_iff.1 hu₁ | ||||||
|     use v |     use v | ||||||
|     constructor |     constructor | ||||||
|   | |||||||
| @@ -281,7 +281,11 @@ lemma partialDeriv_fderiv {f : E → F} (hf : ContDiff 𝕜 2 f) (z a b : E) : | |||||||
|   unfold partialDeriv |   unfold partialDeriv | ||||||
|   rw [fderiv_clm_apply] |   rw [fderiv_clm_apply] | ||||||
|   · simp |   · simp | ||||||
|   · exact (contDiff_succ_iff_fderiv.1 hf).2.differentiable le_rfl z |   · apply Differentiable.differentiableAt | ||||||
|  |     rw [← one_add_one_eq_two] at hf | ||||||
|  |     rw [contDiff_succ_iff_fderiv] at hf | ||||||
|  |     apply hf.2.2.differentiable | ||||||
|  |     simp | ||||||
|   · simp |   · simp | ||||||
|  |  | ||||||
|  |  | ||||||
| @@ -298,7 +302,9 @@ lemma partialDeriv_fderivOn | |||||||
|   · simp |   · simp | ||||||
|   · convert DifferentiableOn.differentiableAt _ (IsOpen.mem_nhds hs hz) |   · convert DifferentiableOn.differentiableAt _ (IsOpen.mem_nhds hs hz) | ||||||
|     apply ContDiffOn.differentiableOn _ (Preorder.le_refl 1) |     apply ContDiffOn.differentiableOn _ (Preorder.le_refl 1) | ||||||
|     exact ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 hf).2 |     rw [← one_add_one_eq_two] at hf | ||||||
|  |     rw [contDiffOn_succ_iff_fderiv_of_isOpen hs] at hf | ||||||
|  |     exact hf.2.2 | ||||||
|   · simp |   · simp | ||||||
|  |  | ||||||
|  |  | ||||||
| @@ -407,7 +413,8 @@ theorem partialDeriv_comm | |||||||
|     let f'' := (fderiv ℝ f' z) |     let f'' := (fderiv ℝ f' z) | ||||||
|     have h₁ : HasFDerivAt f' f'' z := by |     have h₁ : HasFDerivAt f' f'' z := by | ||||||
|       apply DifferentiableAt.hasFDerivAt |       apply DifferentiableAt.hasFDerivAt | ||||||
|       apply (contDiff_succ_iff_fderiv.1 h).right.differentiable (Preorder.le_refl 1) |       rw [← one_add_one_eq_two] at h | ||||||
|  |       apply (contDiff_succ_iff_fderiv.1 h).2.2.differentiable (Preorder.le_refl 1) | ||||||
|  |  | ||||||
|     apply second_derivative_symmetric h₀ h₁ v₁ v₂ |     apply second_derivative_symmetric h₀ h₁ v₁ v₂ | ||||||
|  |  | ||||||
| @@ -442,7 +449,8 @@ theorem partialDeriv_commOn | |||||||
|       apply DifferentiableAt.hasFDerivAt |       apply DifferentiableAt.hasFDerivAt | ||||||
|       apply DifferentiableOn.differentiableAt _ (IsOpen.mem_nhds hs hz) |       apply DifferentiableOn.differentiableAt _ (IsOpen.mem_nhds hs hz) | ||||||
|       apply ContDiffOn.differentiableOn _ (Preorder.le_refl 1) |       apply ContDiffOn.differentiableOn _ (Preorder.le_refl 1) | ||||||
|       exact ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 h).2 |       rw [← one_add_one_eq_two] at h | ||||||
|  |       exact ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 h).2.2 | ||||||
|  |  | ||||||
|     have h₀' : ∀ᶠ (y : E) in nhds z, HasFDerivAt f (f' y) y := by |     have h₀' : ∀ᶠ (y : E) in nhds z, HasFDerivAt f (f' y) y := by | ||||||
|       apply eventually_nhds_iff.mpr |       apply eventually_nhds_iff.mpr | ||||||
| @@ -463,7 +471,9 @@ theorem partialDeriv_commAt | |||||||
|   (h : ContDiffAt ℝ 2 f z) : |   (h : ContDiffAt ℝ 2 f z) : | ||||||
|   ∀ v₁ v₂ : E, partialDeriv ℝ v₁ (partialDeriv ℝ v₂ f) z = partialDeriv ℝ v₂ (partialDeriv ℝ v₁ f) z := by |   ∀ v₁ v₂ : E, partialDeriv ℝ v₁ (partialDeriv ℝ v₂ f) z = partialDeriv ℝ v₂ (partialDeriv ℝ v₁ f) z := by | ||||||
|  |  | ||||||
|   obtain ⟨u, hu₁, hu₂⟩ := h.contDiffOn le_rfl |   let A := h.contDiffOn le_rfl | ||||||
|  |   simp at A | ||||||
|  |   obtain ⟨u, hu₁, hu₂⟩ := A | ||||||
|   obtain ⟨v, hv₁, hv₂, hv₃⟩ := mem_nhds_iff.1 hu₁ |   obtain ⟨v, hv₁, hv₂, hv₃⟩ := mem_nhds_iff.1 hu₁ | ||||||
|  |  | ||||||
|   intro v₁ v₂ |   intro v₁ v₂ | ||||||
|   | |||||||
| @@ -1,64 +1,24 @@ | |||||||
| {"version": "1.1.0", | {"version": "1.1.0", | ||||||
|  "packagesDir": ".lake/packages", |  "packagesDir": ".lake/packages", | ||||||
|  "packages": |  "packages": | ||||||
|  [{"url": "https://github.com/leanprover-community/batteries", |  [{"url": "https://github.com/leanprover-community/mathlib4.git", | ||||||
|    "type": "git", |    "type": "git", | ||||||
|    "subDir": null, |    "subDir": null, | ||||||
|    "scope": "leanprover-community", |    "scope": "", | ||||||
|    "rev": "c933dd9b00271d869e22b802a015092d1e8e454a", |    "rev": "7cf807751deab8d4943d867898dc1b31d61b746a", | ||||||
|    "name": "batteries", |    "name": "mathlib", | ||||||
|    "manifestFile": "lake-manifest.json", |    "manifestFile": "lake-manifest.json", | ||||||
|    "inputRev": "main", |    "inputRev": null, | ||||||
|    "inherited": true, |    "inherited": false, | ||||||
|    "configFile": "lakefile.toml"}, |  | ||||||
|   {"url": "https://github.com/leanprover-community/quote4", |  | ||||||
|    "type": "git", |  | ||||||
|    "subDir": null, |  | ||||||
|    "scope": "leanprover-community", |  | ||||||
|    "rev": "303b23fbcea94ac4f96e590c1cad6618fd4f5f41", |  | ||||||
|    "name": "Qq", |  | ||||||
|    "manifestFile": "lake-manifest.json", |  | ||||||
|    "inputRev": "master", |  | ||||||
|    "inherited": true, |  | ||||||
|    "configFile": "lakefile.lean"}, |    "configFile": "lakefile.lean"}, | ||||||
|   {"url": "https://github.com/leanprover-community/aesop", |   {"url": "https://github.com/leanprover-community/plausible", | ||||||
|    "type": "git", |    "type": "git", | ||||||
|    "subDir": null, |    "subDir": null, | ||||||
|    "scope": "leanprover-community", |    "scope": "leanprover-community", | ||||||
|    "rev": "de91b59101763419997026c35a41432ac8691f15", |    "rev": "8e5cb8d424df462f84997dd68af6f40e347c3e35", | ||||||
|    "name": "aesop", |    "name": "plausible", | ||||||
|    "manifestFile": "lake-manifest.json", |    "manifestFile": "lake-manifest.json", | ||||||
|    "inputRev": "master", |    "inputRev": "v4.15.0-rc1", | ||||||
|    "inherited": true, |  | ||||||
|    "configFile": "lakefile.toml"}, |  | ||||||
|   {"url": "https://github.com/leanprover-community/ProofWidgets4", |  | ||||||
|    "type": "git", |  | ||||||
|    "subDir": null, |  | ||||||
|    "scope": "leanprover-community", |  | ||||||
|    "rev": "1383e72b40dd62a566896a6e348ffe868801b172", |  | ||||||
|    "name": "proofwidgets", |  | ||||||
|    "manifestFile": "lake-manifest.json", |  | ||||||
|    "inputRev": "v0.0.46", |  | ||||||
|    "inherited": true, |  | ||||||
|    "configFile": "lakefile.lean"}, |  | ||||||
|   {"url": "https://github.com/leanprover/lean4-cli", |  | ||||||
|    "type": "git", |  | ||||||
|    "subDir": null, |  | ||||||
|    "scope": "leanprover", |  | ||||||
|    "rev": "726b3c9ad13acca724d4651f14afc4804a7b0e4d", |  | ||||||
|    "name": "Cli", |  | ||||||
|    "manifestFile": "lake-manifest.json", |  | ||||||
|    "inputRev": "main", |  | ||||||
|    "inherited": true, |  | ||||||
|    "configFile": "lakefile.toml"}, |  | ||||||
|   {"url": "https://github.com/leanprover-community/import-graph", |  | ||||||
|    "type": "git", |  | ||||||
|    "subDir": null, |  | ||||||
|    "scope": "leanprover-community", |  | ||||||
|    "rev": "119b022b3ea88ec810a677888528e50f8144a26e", |  | ||||||
|    "name": "importGraph", |  | ||||||
|    "manifestFile": "lake-manifest.json", |  | ||||||
|    "inputRev": "main", |  | ||||||
|    "inherited": true, |    "inherited": true, | ||||||
|    "configFile": "lakefile.toml"}, |    "configFile": "lakefile.toml"}, | ||||||
|   {"url": "https://github.com/leanprover-community/LeanSearchClient", |   {"url": "https://github.com/leanprover-community/LeanSearchClient", | ||||||
| @@ -71,25 +31,65 @@ | |||||||
|    "inputRev": "main", |    "inputRev": "main", | ||||||
|    "inherited": true, |    "inherited": true, | ||||||
|    "configFile": "lakefile.toml"}, |    "configFile": "lakefile.toml"}, | ||||||
|   {"url": "https://github.com/leanprover-community/plausible", |   {"url": "https://github.com/leanprover-community/import-graph", | ||||||
|    "type": "git", |    "type": "git", | ||||||
|    "subDir": null, |    "subDir": null, | ||||||
|    "scope": "leanprover-community", |    "scope": "leanprover-community", | ||||||
|    "rev": "42dc02bdbc5d0c2f395718462a76c3d87318f7fa", |    "rev": "ed3b856bd8893ade75cafe13e8544d4c2660f377", | ||||||
|    "name": "plausible", |    "name": "importGraph", | ||||||
|  |    "manifestFile": "lake-manifest.json", | ||||||
|  |    "inputRev": "v4.15.0-rc1", | ||||||
|  |    "inherited": true, | ||||||
|  |    "configFile": "lakefile.toml"}, | ||||||
|  |   {"url": "https://github.com/leanprover-community/ProofWidgets4", | ||||||
|  |    "type": "git", | ||||||
|  |    "subDir": null, | ||||||
|  |    "scope": "leanprover-community", | ||||||
|  |    "rev": "2b000e02d50394af68cfb4770a291113d94801b5", | ||||||
|  |    "name": "proofwidgets", | ||||||
|  |    "manifestFile": "lake-manifest.json", | ||||||
|  |    "inputRev": "v0.0.48", | ||||||
|  |    "inherited": true, | ||||||
|  |    "configFile": "lakefile.lean"}, | ||||||
|  |   {"url": "https://github.com/leanprover-community/aesop", | ||||||
|  |    "type": "git", | ||||||
|  |    "subDir": null, | ||||||
|  |    "scope": "leanprover-community", | ||||||
|  |    "rev": "8b6048aa0a4a4b6bcf83597802d8dee734e64b7e", | ||||||
|  |    "name": "aesop", | ||||||
|  |    "manifestFile": "lake-manifest.json", | ||||||
|  |    "inputRev": "v4.15.0-rc1", | ||||||
|  |    "inherited": true, | ||||||
|  |    "configFile": "lakefile.toml"}, | ||||||
|  |   {"url": "https://github.com/leanprover-community/quote4", | ||||||
|  |    "type": "git", | ||||||
|  |    "subDir": null, | ||||||
|  |    "scope": "leanprover-community", | ||||||
|  |    "rev": "ad942fdf0b15c38bface6acbb01d63855a2519ac", | ||||||
|  |    "name": "Qq", | ||||||
|  |    "manifestFile": "lake-manifest.json", | ||||||
|  |    "inputRev": "v4.14.0", | ||||||
|  |    "inherited": true, | ||||||
|  |    "configFile": "lakefile.lean"}, | ||||||
|  |   {"url": "https://github.com/leanprover-community/batteries", | ||||||
|  |    "type": "git", | ||||||
|  |    "subDir": null, | ||||||
|  |    "scope": "leanprover-community", | ||||||
|  |    "rev": "7805acf1864ba1a2e359f86a8f092ccf1438ad83", | ||||||
|  |    "name": "batteries", | ||||||
|    "manifestFile": "lake-manifest.json", |    "manifestFile": "lake-manifest.json", | ||||||
|    "inputRev": "main", |    "inputRev": "main", | ||||||
|    "inherited": true, |    "inherited": true, | ||||||
|    "configFile": "lakefile.toml"}, |    "configFile": "lakefile.toml"}, | ||||||
|   {"url": "https://github.com/leanprover-community/mathlib4.git", |   {"url": "https://github.com/leanprover/lean4-cli", | ||||||
|    "type": "git", |    "type": "git", | ||||||
|    "subDir": null, |    "subDir": null, | ||||||
|    "scope": "", |    "scope": "leanprover", | ||||||
|    "rev": "ab780650a6c8db2942d137b91ede7c94c47a4950", |    "rev": "0c8ea32a15a4f74143e4e1e107ba2c412adb90fd", | ||||||
|    "name": "mathlib", |    "name": "Cli", | ||||||
|    "manifestFile": "lake-manifest.json", |    "manifestFile": "lake-manifest.json", | ||||||
|    "inputRev": null, |    "inputRev": "main", | ||||||
|    "inherited": false, |    "inherited": true, | ||||||
|    "configFile": "lakefile.lean"}], |    "configFile": "lakefile.toml"}], | ||||||
|  "name": "nevanlinna", |  "name": "nevanlinna", | ||||||
|  "lakeDir": ".lake"} |  "lakeDir": ".lake"} | ||||||
|   | |||||||
| @@ -1 +1 @@ | |||||||
| leanprover/lean4:v4.14.0-rc3 | leanprover/lean4:v4.15.0-rc1 | ||||||
|   | |||||||
		Reference in New Issue
	
	Block a user
	 Stefan Kebekus
					Stefan Kebekus