Skip to content

Commit

Permalink
bump mathlib
Browse files Browse the repository at this point in the history
  • Loading branch information
YaelDillies committed Nov 12, 2023
1 parent 80240ea commit 206cf42
Show file tree
Hide file tree
Showing 16 changed files with 32 additions and 35 deletions.
6 changes: 3 additions & 3 deletions LeanAPAP.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,13 +8,13 @@ import LeanAPAP.Mathlib.Algebra.BigOperators.Pi
import LeanAPAP.Mathlib.Algebra.BigOperators.Ring
import LeanAPAP.Mathlib.Algebra.DirectSum.Basic
import LeanAPAP.Mathlib.Algebra.Group.Basic
import LeanAPAP.Mathlib.Algebra.Group.Equiv.Basic
import LeanAPAP.Mathlib.Algebra.Group.Hom.Defs
import LeanAPAP.Mathlib.Algebra.Group.Hom.Instances
import LeanAPAP.Mathlib.Algebra.GroupPower.Basic
import LeanAPAP.Mathlib.Algebra.GroupPower.Hom
import LeanAPAP.Mathlib.Algebra.GroupPower.Order
import LeanAPAP.Mathlib.Algebra.Group.TypeTags
import LeanAPAP.Mathlib.Algebra.Hom.Equiv.Basic
import LeanAPAP.Mathlib.Algebra.Hom.Group
import LeanAPAP.Mathlib.Algebra.Hom.GroupInstances
import LeanAPAP.Mathlib.Algebra.Module.Basic
import LeanAPAP.Mathlib.Algebra.Order.LatticeGroup
import LeanAPAP.Mathlib.Algebra.Order.Ring.Canonical
Expand Down
2 changes: 1 addition & 1 deletion LeanAPAP/Mathlib/Algebra/BigOperators/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -59,7 +59,7 @@ def processBigOpBinder (processed : (Array (Term × Term)))
| `(bigOpBinder| $x:term) =>
match x with
| `(($a + $b = $n)) => -- Maybe this is too cute.
return processed |>.push (← `(⟨$a, $b⟩), ← `(Finset.Nat.antidiagonal $n))
return processed |>.push (← `(⟨$a, $b⟩), ← `(Finset.HasAntidiagonal.antidiagonal $n))
| _ => return processed |>.push (x, ← ``(Finset.univ))
| `(bigOpBinder| $x : $t) => return processed |>.push (x, ← ``((Finset.univ : Finset $t)))
| `(bigOpBinder| $x ∈ $s) => return processed |>.push (x, ← `(finset% $s))
Expand Down
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
import Mathlib.Algebra.Hom.Equiv.Basic
import Mathlib.Algebra.Group.Equiv.Basic

open Function

Expand Down
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
import Mathlib.Algebra.Hom.Group.Defs
import Mathlib.Algebra.Group.Hom.Defs
import Mathlib.Tactic.Coe

namespace AddMonoid.End
Expand Down
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
import Mathlib.Algebra.Hom.GroupInstances
import Mathlib.Algebra.Group.Hom.Instances
import Mathlib.Data.Pi.Algebra

variable {α β : Type*}
Expand Down
4 changes: 2 additions & 2 deletions LeanAPAP/Mathlib/Algebra/GroupPower/Hom.lean
Original file line number Diff line number Diff line change
@@ -1,5 +1,5 @@
import LeanAPAP.Mathlib.Algebra.Hom.Group
import LeanAPAP.Mathlib.Algebra.Hom.GroupInstances
import LeanAPAP.Mathlib.Algebra.Group.Hom.Defs
import LeanAPAP.Mathlib.Algebra.Group.Hom.Instances
import LeanAPAP.Mathlib.GroupTheory.OrderOfElement

open Finset Function
Expand Down
4 changes: 2 additions & 2 deletions LeanAPAP/Mathlib/Data/Finset/NatAntidiagonal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -17,10 +17,10 @@ lemma antidiagonal_congr' (hp : p ∈ antidiagonal n) (hq : q ∈ antidiagonal n
exact antidiagonal_congr (swap_mem_antidiagonal hp) (swap_mem_antidiagonal hq)

lemma antidiagonal_eq_map :
Nat.antidiagonal n =
antidiagonal n =
(range (n + 1)).map ⟨fun i ↦ (i, n - i), fun _i _j h ↦ (Prod.ext_iff.1 h).1⟩ := rfl

lemma antidiagonal_eq_image : Nat.antidiagonal n = (range (n + 1)).image fun i ↦ (i, n - i) := by
lemma antidiagonal_eq_image : antidiagonal n = (range (n + 1)).image fun i ↦ (i, n - i) := by
simp only [antidiagonal_eq_map, map_eq_image, Function.Embedding.coeFn_mk]

end Nat
Expand Down
1 change: 0 additions & 1 deletion LeanAPAP/Mathlib/Data/Finset/Pointwise.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,7 +3,6 @@ import LeanAPAP.Mathlib.Data.Fintype.Pi
import Mathlib.GroupTheory.GroupAction.Pi

open Finset

open scoped Pointwise

namespace Finset
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -101,7 +101,7 @@ lemma nsmul_apply (n : ℕ) (ψ : AddChar G R) (a : G) : (ψ ^ n) a = ψ a ^ n :
variable {ι : Type*}

@[simp, norm_cast]
lemma coe_sum (s : Finset ι) (ψ : ι → AddChar G R) : ⇑(∑ i in s, ψ i) = ∏ i in s, ⇑ψ i := by
lemma coe_sum (s : Finset ι) (ψ : ι → AddChar G R) : ∑ i in s, ψ i = ∏ i in s, ⇑(ψ i) := by
induction s using Finset.cons_induction <;> simp [*]

lemma sum_apply (s : Finset ι) (ψ : ι → AddChar G R) (a : G) :
Expand Down
2 changes: 0 additions & 2 deletions LeanAPAP/Physics/Unbalancing.lean
Original file line number Diff line number Diff line change
Expand Up @@ -152,8 +152,6 @@ private lemma unbalancing' (p : ℕ) (hp : 5 ≤ p) (hp₁ : Odd p) (hε₀ : 0
· exact rpow_nonneg (sum_nonneg fun i _ ↦ by sorry) -- positivity
· positivity
· exact sum_nonneg fun i _ ↦ by sorry -- positivity
· sorry -- positivity
· sorry -- positivity
set p' := 24 / ε * log (3 / ε) * p
have hp' : 0 < p' := p'_pos hp hε₀ hε₁
have : 1 - 8⁻¹ * ε ≤ (∑ i in T, ↑(ν i)) ^ p'⁻¹ := by
Expand Down
10 changes: 5 additions & 5 deletions LeanAPAP/Prereqs/Cut.lean
Original file line number Diff line number Diff line change
Expand Up @@ -39,7 +39,7 @@ lemma mem_cut (s : Finset ι) (n : ℕ) (f : ι → ℕ) :
exact hf _

lemma cut_equiv_antidiag (n : ℕ) :
Equiv.finsetCongr (Equiv.boolArrowEquivProd _) (cut univ n) = Nat.antidiagonal n := by
Equiv.finsetCongr (Equiv.boolArrowEquivProd _) (cut univ n) = antidiagonal n := by
ext ⟨x₁, x₂⟩
simp_rw [Equiv.finsetCongr_apply, mem_map, Equiv.toEmbedding, Function.Embedding.coeFn_mk, ←
Equiv.eq_symm_apply]
Expand All @@ -63,13 +63,13 @@ lemma cut_univ_fin_eq_antidiagonalTuple (n k : ℕ) : cut univ n = Nat.antidiago
eq_empty_of_forall_not_mem fun x ↦ by simp [mem_cut, @eq_comm _ 0]

lemma cut_insert [DecidableEq (ι → ℕ)] (n : ℕ) (a : ι) (s : Finset ι) (h : a ∉ s) :
cut (insert a s) n = (Nat.antidiagonal n).biUnion fun p : ℕ × ℕ ↦ (cut s p.snd).map
cut (insert a s) n = (antidiagonal n).biUnion fun p : ℕ × ℕ ↦ (cut s p.snd).map
(addLeftEmbedding fun t ↦ if t = a then p.fst else 0) := by
ext f
rw [mem_cut, mem_biUnion, sum_insert h]
constructor
· rintro ⟨rfl, h₁⟩
simp only [exists_prop, Function.Embedding.coeFn_mk, mem_map, Nat.mem_antidiagonal, Prod.exists]
simp only [exists_prop, Function.Embedding.coeFn_mk, mem_map, mem_antidiagonal, Prod.exists]
refine' ⟨f a, s.sum f, rfl, fun i ↦ if i = a then 0 else f i, _, _⟩
· rw [mem_cut]
refine' ⟨_, _⟩
Expand All @@ -85,7 +85,7 @@ lemma cut_insert [DecidableEq (ι → ℕ)] (n : ℕ) (a : ι) (s : Finset ι) (
obtain rfl | h := eq_or_ne x a
· simp
· simp [if_neg h]
· simp only [mem_insert, Function.Embedding.coeFn_mk, mem_map, Nat.mem_antidiagonal, Prod.exists,
· simp only [mem_insert, Function.Embedding.coeFn_mk, mem_map, mem_antidiagonal, Prod.exists,
exists_prop, mem_cut, not_or]
rintro ⟨p, q, rfl, g, ⟨rfl, hg₂⟩, rfl⟩
refine' ⟨_, _⟩
Expand All @@ -94,7 +94,7 @@ lemma cut_insert [DecidableEq (ι → ℕ)] (n : ℕ) (a : ι) (s : Finset ι) (
simp [if_neg h₁, hg₂ _ h₂]

lemma cut_insert_disjoint_bUnion (n : ℕ) (a : ι) (s : Finset ι) (h : a ∉ s) :
(Nat.antidiagonal n : Set (ℕ × ℕ)).PairwiseDisjoint fun p : ℕ × ℕ ↦
(antidiagonal n : Set (ℕ × ℕ)).PairwiseDisjoint fun p : ℕ × ℕ ↦
(cut s p.snd).map (addLeftEmbedding fun t ↦ if t = a then p.fst else 0) := by
simp only [Set.PairwiseDisjoint, Function.onFun_apply, Finset.disjoint_iff_ne, mem_map,
Function.Embedding.coeFn_mk, Ne.def, forall_exists_index, Set.Pairwise, mem_coe, mem_cut,
Expand Down
2 changes: 1 addition & 1 deletion LeanAPAP/Prereqs/LpNorm.lean
Original file line number Diff line number Diff line change
Expand Up @@ -57,7 +57,7 @@ lemma l2norm_eq_sum (f : ∀ i, α i) : ‖f‖_[2] = sqrt (∑ i, ‖f i‖ ^ 2
lemma L1norm_eq_sum (f : ∀ i, α i) : ‖f‖_[1] = ∑ i, ‖f i‖ := by simp [lpNorm_eq_sum']

lemma L0norm_eq_card (f : ∀ i, α i) : ‖f‖_[0] = {i | f i ≠ 0}.toFinite.toFinset.card :=
PiLp.norm_eq_card _
(PiLp.norm_eq_card _).trans $ by simp

lemma Linftynorm_eq_csupr (f : ∀ i, α i) : ‖f‖_[∞] = ⨆ i, ‖f i‖ := PiLp.norm_eq_ciSup _

Expand Down
4 changes: 2 additions & 2 deletions LeanAPAP/Prereqs/Multinomial.lean
Original file line number Diff line number Diff line change
Expand Up @@ -31,7 +31,7 @@ lemma multinomial_expansion' {α β : Type*} [DecidableEq α] [CommSemiring β]
Pi.add_apply, eq_self_iff_true, if_true, Nat.cast_mul, prod_insert has, eq_self_iff_true,
if_true, sum_add_distrib, sum_ite_eq', has, if_false, add_zero,
addLeftEmbedding_eq_addRightEmbedding, addRightEmbedding_apply]
suffices : ∀ p : ℕ × ℕ, p ∈ Nat.antidiagonal n →
suffices : ∀ p : ℕ × ℕ, p ∈ antidiagonal n →
∑ f in cut s p.snd, ((f a + p.fst + s.sum f).choose (f a + p.fst) : β) *
multinomial s (f + fun t ↦ ite (t = a) p.fst 0) *
(x a ^ (f a + p.fst) * ∏ t : α in s, x t ^ (f t + ite (t = a) p.fst 0)) =
Expand All @@ -46,7 +46,7 @@ lemma multinomial_expansion' {α β : Type*} [DecidableEq α] [CommSemiring β]
rw [mem_cut] at hf
rw [hf.2 _ has, zero_add, hf.1]
congr 2
· rw [Nat.mem_antidiagonal.1 hp]
· rw [mem_antidiagonal.1 hp]
· rw [multinomial_congr]
intro t ht
rw [Pi.add_apply, if_neg, add_zero]
Expand Down
22 changes: 11 additions & 11 deletions lake-manifest.json
Original file line number Diff line number Diff line change
Expand Up @@ -4,58 +4,58 @@
[{"git":
{"url": "https://github.com/leanprover-community/mathlib4.git",
"subDir?": null,
"rev": "3ce43c18f614b76e161f911b75a3e1ef641620ff",
"rev": "adefdc2a051225c3e96946f16d99bf93cd74ba12",
"opts": {},
"name": "mathlib",
"inputRev?": null,
"inherited": false}},
{"git":
{"url": "https://github.com/leanprover/doc-gen4",
"subDir?": null,
"rev": "8bccb92b531248af1b6692d65486e8640c8bcd10",
"rev": "162de4ad98c7c79bc1fc0b1110816247679b3ce0",
"opts": {},
"name": "«doc-gen4»",
"inputRev?": "main",
"inherited": false}},
{"git":
{"url": "https://github.com/leanprover/std4",
"subDir?": null,
"rev": "727fa6aa1113c376ea1873812d1ab5c17a24f1d2",
"rev": "fb07d160aff0e8bdf403a78a5167fc7acf9c8227",
"opts": {},
"name": "std",
"inputRev?": "main",
"inherited": true}},
{"git":
{"url": "https://github.com/leanprover-community/quote4",
"subDir?": null,
"rev": "a387c0eb611857e2460cf97a8e861c944286e6b2",
"rev": "511eb96eca98a7474683b8ae55193a7e7c51bc39",
"opts": {},
"name": "Qq",
"inputRev?": "master",
"inherited": true}},
{"git":
{"url": "https://github.com/leanprover-community/aesop",
"subDir?": null,
"rev": "b601328752091a1cfcaebdd6b6b7c30dc5ffd946",
"rev": "cb87803274405db79ec578fc07c4730c093efb90",
"opts": {},
"name": "aesop",
"inputRev?": "master",
"inherited": true}},
{"git":
{"url": "https://github.com/leanprover/lean4-cli",
"subDir?": null,
"rev": "39229f3630d734af7d9cfb5937ddc6b41d3aa6aa",
"rev": "a751d21d4b68c999accb6fc5d960538af26ad5ec",
"opts": {},
"name": "Cli",
"inputRev?": "nightly",
"inputRev?": "main",
"inherited": true}},
{"git":
{"url": "https://github.com/leanprover-community/ProofWidgets4",
"subDir?": null,
"rev": "27715d1daf32b9657dc38cd52172d77b19bde4ba",
"rev": "f1a5c7808b001305ba07d8626f45ee054282f589",
"opts": {},
"name": "proofwidgets",
"inputRev?": "v0.0.18",
"inputRev?": "v0.0.21",
"inherited": true}},
{"git":
{"url": "https://github.com/xubaiw/CMark.lean",
Expand All @@ -68,9 +68,9 @@
{"git":
{"url": "https://github.com/fgdorais/lean4-unicode-basic",
"subDir?": null,
"rev": "f09250282cea3ed8c010f430264d9e8e50d7bc32",
"rev": "4ecf4f1f98d14d03a9e84fa1c082630fa69df88b",
"opts": {},
"name": "«lean4-unicode-basic»",
"name": "UnicodeBasic",
"inputRev?": "main",
"inherited": true}},
{"git":
Expand Down
2 changes: 1 addition & 1 deletion lean-toolchain
Original file line number Diff line number Diff line change
@@ -1 +1 @@
leanprover/lean4:v4.2.0-rc4
leanprover/lean4:v4.3.0-rc1
Empty file modified scripts/update_mathlib.sh
100644 → 100755
Empty file.

0 comments on commit 206cf42

Please sign in to comment.