chore: Bump to lean v.4.12.0

This commit is contained in:
jstoobysmith 2024-10-03 13:50:18 +00:00
parent 47d9649c44
commit 987bbf6013
23 changed files with 92 additions and 58 deletions

View file

@ -51,8 +51,8 @@ variable {d : } {X X' Y Y' Z W : Type} [Fintype X] [DecidableEq X] [Fintype Y
/-- A relation on colors which is true if the two colors are equal or are duals. -/
def colorRel (μ ν : 𝓒.Color) : Prop := μ = ν μ = 𝓒ν
instance : Decidable (colorRel 𝓒 μ ν) :=
Or.decidable
instance : Decidable (colorRel 𝓒 μ ν) := instDecidableOr
omit [Fintype 𝓒.Color] [DecidableEq 𝓒.Color]
/-- An equivalence relation on colors which is true if the two colors are equal or are duals. -/
lemma colorRel_equivalence : Equivalence 𝓒.colorRel where
@ -90,8 +90,7 @@ instance colorSetoid : Setoid 𝓒.Color := ⟨𝓒.colorRel, 𝓒.colorRel_equi
def colorQuot (μ : 𝓒.Color) : Quotient 𝓒.colorSetoid :=
Quotient.mk 𝓒.colorSetoid μ
instance (μ ν : 𝓒.Color) : Decidable (μ ≈ ν) :=
Or.decidable
instance (μ ν : 𝓒.Color) : Decidable (μ ≈ ν) := instDecidableOr
instance : DecidableEq (Quotient 𝓒.colorSetoid) :=
instDecidableEqQuotientOfDecidableEquiv

View file

@ -75,7 +75,7 @@ lemma orderEmbOfFin_univ (n m : ) (h : n = m) :
intro x
exact Finset.mem_univ ((Fin.castOrderIso (Eq.symm h)).toFun x)
exact fun ⦃a b⦄ a => a
exact Eq.symm (Fin.orderEmbedding_eq (congrArg Set.range (id (Eq.symm h1))))
exact Eq.symm (OrderEmbedding.range_inj.mp (congrArg Set.range (id (Eq.symm h1))))
/-!

View file

@ -170,7 +170,7 @@ lemma symm (h : ContrPerm l l') : ContrPerm l' l := by
rw [ContrPerm] at h ⊢
apply And.intro h.1.symm
apply And.intro (Subperm.symm h.2.1 h.1)
rw [← Function.comp.assoc, ← h.2.2, Function.comp.assoc, Function.comp.assoc]
rw [← Function.comp_assoc, ← h.2.2, Function.comp_assoc, Function.comp_assoc]
rw [show (l.contr.getDualInOtherEquiv l'.contr) =
(l'.contr.getDualInOtherEquiv l.contr).symm from rfl]
simp only [Equiv.symm_comp_self, CompTriple.comp_eq]

View file

@ -309,27 +309,27 @@ lemma filter_sort_comm {n : } (s : Finset (Fin n)) (p : Fin n → Prop) [Deci
intro m
simp only [Multiset.quot_mk_to_coe'', Multiset.coe_sort, Multiset.filter_coe]
have h1 : List.Sorted (fun i j => i ≤ j) (List.filter (fun b => decide (p b))
(List.mergeSort (fun i j => i ≤ j) m)) := by
(List.mergeSort' (fun i j => i ≤ j) m)) := by
simp only [List.Sorted]
rw [List.pairwise_filter, List.pairwise_iff_get]
intro i j h1 _ _
have hs : List.Sorted (fun i j => i ≤ j) (List.mergeSort (fun i j => i ≤ j) m) := by
exact List.sorted_mergeSort (fun i j => i ≤ j) m
have hs : List.Sorted (fun i j => i ≤ j) (List.mergeSort' (fun i j => i ≤ j) m) := by
exact List.sorted_mergeSort' (fun i j => i ≤ j) m
simp only [List.Sorted] at hs
rw [List.pairwise_iff_get] at hs
exact hs i j h1
have hp1 : (List.mergeSort (fun i j => i ≤ j) m).Perm m := by
exact List.perm_mergeSort (fun i j => i ≤ j) m
have hp2 : (List.filter (fun b => decide (p b)) ((List.mergeSort (fun i j => i ≤ j) m))).Perm
have hp1 : (List.mergeSort' (fun i j => i ≤ j) m).Perm m := by
exact List.perm_mergeSort' (fun i j => i ≤ j) m
have hp2 : (List.filter (fun b => decide (p b)) ((List.mergeSort' (fun i j => i ≤ j) m))).Perm
(List.filter (fun b => decide (p b)) m) := by
exact List.Perm.filter (fun b => decide (p b)) hp1
have hp3 : (List.filter (fun b => decide (p b)) m).Perm
(List.mergeSort (fun i j => i ≤ j) (List.filter (fun b => decide (p b)) m)) := by
exact List.Perm.symm (List.perm_mergeSort (fun i j => i ≤ j)
(List.mergeSort' (fun i j => i ≤ j) (List.filter (fun b => decide (p b)) m)) := by
exact List.Perm.symm (List.perm_mergeSort' (fun i j => i ≤ j)
(List.filter (fun b => decide (p b)) m))
have hp4 := hp2.trans hp3
refine List.eq_of_perm_of_sorted hp4 h1 ?_
exact List.sorted_mergeSort (fun i j => i ≤ j) (List.filter (fun b => decide (p b)) m)
exact List.sorted_mergeSort' (fun i j => i ≤ j) (List.filter (fun b => decide (p b)) m)
exact this s.val
omit [IndexNotation X] [Fintype X] [DecidableEq X] in

View file

@ -331,7 +331,8 @@ lemma countSelf_eq_countDual_iff_of_isSome (hl : l.OnlyUniqueDuals)
exact hn hn''
erw [hm'']
rfl
· exact true_iff_iff.mpr hm
· rw [true_iff]
exact hm
· simp only [hm, iff_false, ne_eq]
simp only [colorMap, List.get_eq_getElem] at hm
have hm' : decide (l.val[↑i].toColor = 𝓒.τ l.val[↑i].toColor) = false := by simpa using hm

View file

@ -129,7 +129,9 @@ lemma mem_withDualInOther_of_withUniqueDualInOther (i : l.withUniqueDualInOther
@[simp]
lemma withDual_isSome (i : l.withDual) : (l.getDual? i).isSome := by
simpa [withDual] using i.2
have hx := i.2
simp only [Finset.mem_filter, withDual] at hx
exact hx.2
@[simp]
lemma mem_withDual_iff_isSome (i : Fin l.length) : i ∈ l.withDual ↔ (l.getDual? i).isSome := by

View file

@ -172,7 +172,7 @@ lemma contr_of_withDual_empty (T : 𝓣.TensorIndex) (h : T.withDual = ∅) :
Function.comp_apply, LinearEquiv.coe_mk, Equiv.cast_apply, LinearEquiv.coe_symm_mk, cast_cast]
apply cast_eq_iff_heq.mpr
let hl := i.contrEquiv_on_withDual_empty l h
exact let_value_heq f hl
exact congr_arg_heq f hl
omit [DecidableEq 𝓣.Color] in
lemma contr_tensor_of_withDual_empty (T : 𝓣.TensorIndex) (h : T.withDual = ∅) :