From 7d1f15e18acb48d2ba942446dd322fb368fb6a12 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Fri, 3 Jan 2025 15:13:16 +0000 Subject: [PATCH 1/8] feat: Contractions and involutions --- HepLean/Mathematics/List.lean | 22 +- .../Contractions/Basic.lean | 627 +++++++++++++++++- .../Wick/CreateAnnihilateSection.lean | 36 +- .../Wick/Signs/KoszulSignInsert.lean | 56 +- .../Wick/StaticTheorem.lean | 20 +- 5 files changed, 662 insertions(+), 99 deletions(-) diff --git a/HepLean/Mathematics/List.lean b/HepLean/Mathematics/List.lean index c03d5855..28a06500 100644 --- a/HepLean/Mathematics/List.lean +++ b/HepLean/Mathematics/List.lean @@ -662,14 +662,6 @@ def optionErase {I : Type} (l : List I) (i : Option (Fin l.length)) : List I := | none => l | some i => List.eraseIdx l i -/-- Optional erase of an element in a list, with addition for `none`. For `none` adds `a` to the - front of the list, for `some i` removes the `i`th element of the list (does not add `a`). - E.g. `optionEraseZ [0, 1, 2] 4 none = [4, 0, 1, 2]` and - `optionEraseZ [0, 1, 2] 4 (some 1) = [0, 2]`. -/ -def optionEraseZ {I : Type} (l : List I) (a : I) (i : Option (Fin l.length)) : List I := - match i with - | none => a :: l - | some i => List.eraseIdx l i lemma eraseIdx_length {I : Type} (l : List I) (i : Fin l.length) : (List.eraseIdx l i).length + 1 = l.length := by @@ -735,4 +727,18 @@ lemma eraseIdx_insertionSort_fin {I : Type} (le1 : I → I → Prop) [DecidableR = List.insertionSort le1 (r.eraseIdx n) := eraseIdx_insertionSort le1 n.val r (Fin.prop n) +/-- Optional erase of an element in a list, with addition for `none`. For `none` adds `a` to the + front of the list, for `some i` removes the `i`th element of the list (does not add `a`). + E.g. `optionEraseZ [0, 1, 2] 4 none = [4, 0, 1, 2]` and + `optionEraseZ [0, 1, 2] 4 (some 1) = [0, 2]`. -/ +def optionEraseZ {I : Type} (l : List I) (a : I) (i : Option (Fin l.length)) : List I := + match i with + | none => a :: l + | some i => List.eraseIdx l i + +@[simp] +lemma optionEraseZ_some_length {I : Type} (l : List I) (a : I) (i : (Fin l.length)) : + (optionEraseZ l a (some i)).length = l.length - 1 := by + simp [optionEraseZ, List.length_eraseIdx] + end HepLean.List diff --git a/HepLean/PerturbationTheory/Contractions/Basic.lean b/HepLean/PerturbationTheory/Contractions/Basic.lean index e57f787e..f59a5cf0 100644 --- a/HepLean/PerturbationTheory/Contractions/Basic.lean +++ b/HepLean/PerturbationTheory/Contractions/Basic.lean @@ -17,15 +17,15 @@ open FieldStatistic variable {𝓕 : Type} -/-- Given a list of fields `l`, the type of pairwise-contractions associated with `l` - which have the list `aux` uncontracted. -/ -inductive ContractionsAux : (l : List 𝓕) → (aux : List 𝓕) → Type +/-- Given a list of fields `φs`, the type of pairwise-contractions associated with `φs` + which have the list `φsᵤₙ` uncontracted. -/ +inductive ContractionsAux : (φs : List 𝓕) → (φsᵤₙ : List 𝓕) → Type | nil : ContractionsAux [] [] - | cons {l : List 𝓕} {aux : List 𝓕} {a : 𝓕} (i : Option (Fin aux.length)) : - ContractionsAux l aux → ContractionsAux (a :: l) (optionEraseZ aux a i) + | cons {φs : List 𝓕} {φsᵤₙ : List 𝓕} {φ : 𝓕} (i : Option (Fin φsᵤₙ.length)) : + ContractionsAux φs φsᵤₙ → ContractionsAux (φ :: φs) (optionEraseZ φsᵤₙ φ i) /-- Given a list of fields `l`, the type of pairwise-contractions associated with `l`. -/ -def Contractions (l : List 𝓕) : Type := Σ aux, ContractionsAux l aux +def Contractions (φs : List 𝓕) : Type := Σ aux, ContractionsAux φs aux namespace Contractions @@ -40,6 +40,47 @@ lemma contractions_nil (a : Contractions ([] : List 𝓕)) : a = ⟨[], Contract cases c rfl +def embedUncontracted {l : List 𝓕} (c : Contractions l) : Fin c.normalize.length → Fin l.length := + match l, c with + | [], ⟨[], ContractionsAux.nil⟩ => Fin.elim0 + | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => + Fin.cons ⟨0, Nat.zero_lt_succ φs.length⟩ (Fin.succ ∘ (embedUncontracted ⟨aux, c⟩)) + | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => by + let lc := embedUncontracted ⟨aux, c⟩ + refine Fin.succ ∘ lc ∘ Fin.cast ?_ ∘ Fin.succAbove ⟨n, by + rw [normalize, optionEraseZ_some_length] + omega⟩ + simp only [normalize, optionEraseZ_some_length] + have hq : aux.length ≠ 0 := by + by_contra hn + rw [hn] at n + exact Fin.elim0 n + omega + +lemma embedUncontracted_injective {l : List 𝓕} (c : Contractions l) : + Function.Injective c.embedUncontracted := by + match l, c with + | [], ⟨[], ContractionsAux.nil⟩ => + dsimp [embedUncontracted] + intro i + exact Fin.elim0 i + | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => + dsimp [embedUncontracted] + refine Fin.cons_injective_iff.mpr ?_ + apply And.intro + · simp only [Set.mem_range, Function.comp_apply, not_exists] + exact fun x => Fin.succ_ne_zero (embedUncontracted ⟨aux, c⟩ x) + · exact Function.Injective.comp (Fin.succ_injective φs.length) + (embedUncontracted_injective ⟨aux, c⟩) + | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some i) c⟩ => + dsimp [embedUncontracted] + refine Function.Injective.comp (Fin.succ_injective φs.length) ?hf + refine Function.Injective.comp (embedUncontracted_injective ⟨aux, c⟩) ?hf.hf + refine Function.Injective.comp (Fin.cast_injective (embedUncontracted.proof_5 φ φs aux i c)) + Fin.succAbove_right_injective + + + /-- Establishes uniqueness of contractions for a single field, showing that any contraction of a single field must be equivalent to the trivial contraction with no pairing. -/ lemma contractions_single {i : 𝓕} (a : Contractions [i]) : a = @@ -67,46 +108,46 @@ def nilEquiv : Contractions ([] : List 𝓕) ≃ Unit where /-- The equivalence between contractions of `a :: l` and contractions of `Contractions l` paired with an optional element of `Fin (c.normalize).length` specifying what `a` contracts with if any. -/ -def consEquiv {a : 𝓕} {l : List 𝓕} : - Contractions (a :: l) ≃ (c : Contractions l) × Option (Fin (c.normalize).length) where +def consEquiv {φ : 𝓕} {φs : List 𝓕} : + Contractions (φ :: φs) ≃ (c : Contractions φs) × Option (Fin c.normalize.length) where toFun c := match c with | ⟨aux, c⟩ => match c with - | ContractionsAux.cons (aux := aux') i c => ⟨⟨aux', c⟩, i⟩ + | ContractionsAux.cons (φsᵤₙ := aux') i c => ⟨⟨aux', c⟩, i⟩ invFun ci := - ⟨(optionEraseZ (ci.fst.normalize) a ci.2), ContractionsAux.cons (a := a) ci.2 ci.1.2⟩ + ⟨(optionEraseZ (ci.fst.normalize) φ ci.2), ContractionsAux.cons (φ := φ) ci.2 ci.1.2⟩ left_inv c := by match c with | ⟨aux, c⟩ => match c with - | ContractionsAux.cons (aux := aux') i c => rfl + | ContractionsAux.cons (φsᵤₙ := aux') i c => rfl right_inv ci := by rfl /-- The type of contractions is decidable. -/ -instance decidable : (l : List 𝓕) → DecidableEq (Contractions l) +instance decidable : (φs : List 𝓕) → DecidableEq (Contractions φs) | [] => fun a b => match a, b with | ⟨_, a⟩, ⟨_, b⟩ => match a, b with | ContractionsAux.nil, ContractionsAux.nil => isTrue rfl - | _ :: l => - haveI : DecidableEq (Contractions l) := decidable l - haveI : DecidableEq ((c : Contractions l) × Option (Fin (c.normalize).length)) := + | _ :: φs => + haveI : DecidableEq (Contractions φs) := decidable φs + haveI : DecidableEq ((c : Contractions φs) × Option (Fin (c.normalize).length)) := Sigma.instDecidableEqSigma Equiv.decidableEq consEquiv /-- The type of contractions is finite. -/ -instance fintype : (l : List 𝓕) → Fintype (Contractions l) +instance fintype : (φs : List 𝓕) → Fintype (Contractions φs) | [] => { elems := {⟨[], ContractionsAux.nil⟩} complete := by intro a rw [Finset.mem_singleton] exact contractions_nil a} - | a :: l => - haveI : Fintype (Contractions l) := fintype l - haveI : Fintype ((c : Contractions l) × Option (Fin (c.normalize).length)) := + | φ :: φs => + haveI : Fintype (Contractions φs) := fintype φs + haveI : Fintype ((c : Contractions φs) × Option (Fin (c.normalize).length)) := Sigma.instFintype Fintype.ofEquiv _ consEquiv.symm @@ -120,10 +161,525 @@ instance isFull_decidable : Decidable c.IsFull := by simp [IsFull] apply decidable_of_decidable_of_iff hn.symm +def toOptionList : {l : List 𝓕} → (c : Contractions l) → List (Option (Fin l.length)) + | [], ⟨[], ContractionsAux.nil⟩ => [] + | _ :: _, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => none :: + List.map (Option.map Fin.succ) (toOptionList ⟨aux, c⟩) + | _ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => + some (Fin.succ (embedUncontracted ⟨aux, c⟩ n)) :: + List.set ((List.map (Option.map Fin.succ) (toOptionList ⟨aux, c⟩))) + (embedUncontracted ⟨aux, c⟩ n) (some ⟨0, Nat.zero_lt_succ φs.length⟩) + +lemma toOptionList_length {l : List 𝓕} (c : Contractions l) : c.toOptionList.length = l.length := by + match l, c with + | [], ⟨[], ContractionsAux.nil⟩ => + dsimp [toOptionList] + | _ :: _, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => + dsimp [toOptionList] + rw [List.length_map, toOptionList_length ⟨aux, c⟩] + | _ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => + dsimp [toOptionList] + rw [List.length_set, List.length_map, toOptionList_length ⟨aux, c⟩] + +def toFinOptionMap {l : List 𝓕} (c : Contractions l) : Fin l.length → Option (Fin l.length) := + c.toOptionList.get ∘ Fin.cast (toOptionList_length c).symm + +lemma toFinOptionMap_neq_self {l : List 𝓕} (c : Contractions l) (i : Fin l.length) : + c.toFinOptionMap i ≠ some i := by + match l, c with + | [], ⟨[], ContractionsAux.nil⟩ => + exact Fin.elim0 i + | _ :: _, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => + dsimp [toFinOptionMap, toOptionList] + match i with + | ⟨0, _⟩ => + exact Option.noConfusion + | ⟨n + 1, h⟩ => + simp only [List.getElem_cons_succ, List.getElem_map, List.length_cons] + have hn := toFinOptionMap_neq_self ⟨aux, c⟩ ⟨n, Nat.succ_lt_succ_iff.mp h⟩ + simp only [Option.map_eq_some', not_exists, not_and] + intro x hx + simp only [toFinOptionMap, Function.comp_apply, Fin.cast_mk, List.get_eq_getElem, hx, ne_eq, + Option.some.injEq] at hn + rw [Fin.ext_iff] at hn ⊢ + simp_all + | _ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => + dsimp [toFinOptionMap, toOptionList] + match i with + | ⟨0, _⟩ => + simp only [List.getElem_cons_zero, List.length_cons, Fin.zero_eta, Option.some.injEq, ne_eq] + exact Fin.succ_ne_zero (embedUncontracted ⟨aux, c⟩ n) + | ⟨n + 1, h⟩ => + simp only [List.getElem_cons_succ, List.length_cons, ne_eq] + rw [List.getElem_set] + split + · exact ne_of_beq_false rfl + · simp only [List.getElem_map, Option.map_eq_some', not_exists, not_and] + intro x hx + have hn := toFinOptionMap_neq_self ⟨aux, c⟩ ⟨n, Nat.succ_lt_succ_iff.mp h⟩ + simp only [toFinOptionMap, Function.comp_apply, Fin.cast_mk, List.get_eq_getElem, hx, ne_eq, + Option.some.injEq] at hn + rw [Fin.ext_iff] at hn ⊢ + simp_all + +@[simp] +lemma toFinOptionMap_embedUncontracted {l : List 𝓕} (c : Contractions l) + (i : Fin c.normalize.length) : c.toFinOptionMap (embedUncontracted c i) = none := by + match l, c with + | [], ⟨[], ContractionsAux.nil⟩ => + exact Fin.elim0 i + | _ :: _, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => + dsimp [toFinOptionMap, toOptionList, embedUncontracted] + match i with + | ⟨0, _⟩ => + rfl + | ⟨n + 1, h⟩ => + rw [show ⟨n + 1, h⟩ = Fin.succ ⟨n, Nat.succ_lt_succ_iff.mp h⟩ by rfl] + rw [Fin.cons_succ] + simp only [Function.comp_apply, Fin.val_succ, List.getElem_cons_succ, List.getElem_map, + Option.map_eq_none'] + exact toFinOptionMap_embedUncontracted ⟨aux, c⟩ ⟨n, Nat.succ_lt_succ_iff.mp h⟩ + | _ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => + dsimp [toFinOptionMap, toOptionList, embedUncontracted] + rw [List.getElem_set] + split + · rename_i hs + have hx := embedUncontracted_injective ⟨aux, c⟩ (Fin.val_injective hs) + refine False.elim ?_ + have hn := Fin.succAbove_ne ⟨n, embedUncontracted.proof_6 _ φs aux n c⟩ i + simp [Fin.ext_iff] at hx + simp [Fin.ext_iff] at hn + exact hn (id (Eq.symm hx)) + · simp + exact toFinOptionMap_embedUncontracted ⟨aux, c⟩ _ + +lemma toFinOptionMap_involutive {l : List 𝓕} (c : Contractions l) (i j : Fin l.length) : + c.toFinOptionMap i = some j → c.toFinOptionMap j = some i := by + match l, c with + | [], ⟨[], ContractionsAux.nil⟩ => + exact Fin.elim0 i + | _ :: _, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => + dsimp [toFinOptionMap, toOptionList] + match i, j with + | ⟨0, _⟩, j => + exact Option.noConfusion + | ⟨n + 1, h⟩, ⟨0, _⟩ => + simp + intro x hx + exact Fin.succ_ne_zero x + | ⟨n + 1, hn⟩, ⟨m + 1, hm⟩ => + intro h1 + have hnm := toFinOptionMap_involutive ⟨aux, c⟩ ⟨n, Nat.succ_lt_succ_iff.mp hn⟩ + ⟨m, Nat.succ_lt_succ_iff.mp hm⟩ + simp + simp [toFinOptionMap] at hnm + have hnm' := hnm (by + simp at h1 + obtain ⟨a, ha, ha2⟩ := h1 + rw [ha] + simp only [Option.some.injEq] + rw [Fin.ext_iff] at ha2 ⊢ + simp_all) + use ⟨n, Nat.succ_lt_succ_iff.mp hn⟩ + simpa using hnm' + | _ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => + dsimp [toFinOptionMap, toOptionList] + match i, j with + | ⟨0, _⟩, j => + intro hj + simp only [List.getElem_cons_zero, Option.some.injEq] at hj + subst hj + simp + | ⟨n' + 1, h⟩, ⟨0, _⟩ => + intro hj + simp at hj + simp + rw [List.getElem_set] at hj + simp_all only [List.length_cons, lt_add_iff_pos_left, add_pos_iff, zero_lt_one, or_true, List.getElem_map, + ite_eq_left_iff, Option.map_eq_some'] + simp [Fin.ext_iff] + by_contra hn + have hn' := hj hn + obtain ⟨a, ha, ha2⟩ := hn' + exact Fin.succ_ne_zero a ha2 + | ⟨n' + 1, hn⟩, ⟨m + 1, hm⟩ => + simp + rw [List.getElem_set, List.getElem_set] + simp + split + · intro h + simp [Fin.ext_iff] at h + rename_i hn' + intro h1 + split + · rename_i hnx + have hnm := toFinOptionMap_involutive ⟨aux, c⟩ ⟨n', Nat.succ_lt_succ_iff.mp hn⟩ + ⟨m, Nat.succ_lt_succ_iff.mp hm⟩ + subst hnx + simp at hnm + refine False.elim (hnm ?_) + simp at h1 + obtain ⟨a, ha, ha2⟩ := h1 + have ha' : (embedUncontracted ⟨aux, c⟩ n) = a := by + rw [Fin.ext_iff] at ha2 ⊢ + simp_all + rw [ha'] + rw [← ha] + rfl + · rename_i hnx + have hnm := toFinOptionMap_involutive ⟨aux, c⟩ ⟨n', Nat.succ_lt_succ_iff.mp hn⟩ + ⟨m, Nat.succ_lt_succ_iff.mp hm⟩ (by + dsimp [toFinOptionMap] + simp at h1 + obtain ⟨a, ha, ha2⟩ := h1 + have ha' : a = ⟨m, by exact Nat.succ_lt_succ_iff.mp hm⟩ := by + rw [Fin.ext_iff] at ha2 ⊢ + simp_all + rw [← ha', ← ha]) + change Option.map Fin.succ (toFinOptionMap ⟨aux, c⟩ ⟨m, Nat.succ_lt_succ_iff.mp hm⟩) = _ + rw [hnm] + rfl + +def toInvolution {l : List 𝓕} (c : Contractions l) : Fin l.length → Fin l.length := + fun i => + if h : Option.isSome (c.toFinOptionMap i) then Option.get (c.toFinOptionMap i) h else i + +lemma toInvolution_of_isSome {l : List 𝓕} {c : Contractions l} {i : Fin l.length} + (hs : Option.isSome (c.toFinOptionMap i)) : + c.toInvolution i = Option.get (c.toFinOptionMap i) hs := by + dsimp [toInvolution] + rw [dif_pos hs] + +lemma toInvolution_of_eq_none {l : List 𝓕} {c : Contractions l} {i : Fin l.length} + (hs : (c.toFinOptionMap i) = none) : + c.toInvolution i = i := by + dsimp [toInvolution] + rw [dif_neg] + simp_all + +lemma toInvolution_image_isSome {l : List 𝓕} {c : Contractions l} {i : Fin l.length} + (hs : Option.isSome (c.toFinOptionMap i)) : + Option.isSome (c.toFinOptionMap (c.toInvolution i)) := by + dsimp [toInvolution] + rw [dif_pos hs] + have hx := toFinOptionMap_involutive c i ((c.toFinOptionMap i).get hs) + simp at hx + rw [hx] + rfl + +lemma toInvolution_involutive {l : List 𝓕} (c : Contractions l) : + Function.Involutive c.toInvolution := by + intro i + by_cases h : Option.isSome (c.toFinOptionMap i) + · have hx := toFinOptionMap_involutive c i ((c.toFinOptionMap i).get h) + rw [toInvolution_of_isSome (toInvolution_image_isSome h)] + simp only [Option.some_get, forall_const] at hx + simp only [toInvolution_of_isSome h, hx, Option.get_some] + · simp at h + rw [toInvolution_of_eq_none h] + rw [toInvolution_of_eq_none h] + +def involutionEquiv1 {l : List 𝓕} : + {f : Fin l.length → Fin l.length // Function.Involutive f } ≃ + {f : Fin l.length → Option (Fin l.length) // (∀ i, f i ≠ some i) ∧ + ∀ i j, f i = some j → f j = some i} where + toFun f := ⟨fun i => if h : f.1 i = i then none else f.1 i, + fun i => by + simp, + fun i j => by + simp + intro hi hij + subst hij + rw [f.2] + simp + exact fun a => hi (id (Eq.symm a))⟩ + invFun f := ⟨fun i => if h : (f.1 i).isSome then Option.get (f.1 i) h else i, + fun i => by + by_cases h : Option.isSome (f.1 i) + · simp [h] + have hf := f.2.2 i (Option.get (f.1 i) h) (by simp) + simp [hf] + · simp + rw [dif_neg] + · rw [dif_neg h] + · rw [dif_neg h] + exact h⟩ + left_inv f := by + simp + ext i + simp + by_cases hf : f.1 i = i + · simp [hf] + · simp [hf] + right_inv f := by + simp + ext1 + simp + funext i + obtain ⟨val, property⟩ := f + obtain ⟨left, right⟩ := property + simp_all only [ne_eq] + split + next h => + ext a : 1 + simp_all only [Option.mem_def, reduceCtorEq, false_iff] + apply Aesop.BuiltinRules.not_intro + intro a_1 + simp_all only [Option.isSome_some, Option.get_some, forall_const] + next h => + simp_all only [not_forall] + obtain ⟨w, h⟩ := h + simp_all only [↓reduceDIte, Option.some_get] + +def involutionCons (n : ℕ): + {f : Fin n.succ → Fin n.succ // Function.Involutive f } ≃ + (f : {f : Fin n → Fin n // Function.Involutive f}) × {i : Option (Fin n) // + ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} where + toFun f := ⟨⟨ + fun i => + if h : f.1 i.succ = 0 then i + else Fin.pred (f.1 i.succ) h , by + intro i + by_cases h : f.1 i.succ = 0 + · simp [h] + · simp [h] + simp [f.2 i.succ] + intro h + exact False.elim (Fin.succ_ne_zero i h)⟩, + ⟨if h : f.1 0 = 0 then none else Fin.pred (f.1 0) h , by + by_cases h0 : f.1 0 = 0 + · simp [h0] + · simp [h0] + refine fun h => False.elim (h (f.2 0))⟩⟩ + invFun f := ⟨ + if h : (f.2.1).isSome then + Fin.cons (f.2.1.get h).succ (Function.update (Fin.succ ∘ f.1.1) (f.2.1.get h) 0) + else + Fin.cons 0 (Fin.succ ∘ f.1.1) + , by + by_cases hs : (f.2.1).isSome + · simp only [Nat.succ_eq_add_one, hs, ↓reduceDIte, Fin.coe_eq_castSucc] + let a := f.2.1.get hs + change Function.Involutive (Fin.cons a.succ (Function.update (Fin.succ ∘ ↑f.fst) a 0)) + intro i + rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ + · subst hi + rw [Fin.cons_zero, Fin.cons_succ] + simp + · subst hj + rw [Fin.cons_succ] + by_cases hja : j = a + · subst hja + simp + · rw [Function.update_apply ] + rw [if_neg hja] + simp + have hf2 := f.2.2 hs + change f.1.1 a = a at hf2 + have hjf1 : f.1.1 j ≠ a := by + by_contra hn + have haj : j = f.1.1 a := by + rw [← hn] + rw [f.1.2] + rw [hf2] at haj + exact hja haj + rw [Function.update_apply, if_neg hjf1] + simp + rw [f.1.2] + · simp [hs] + intro i + rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ + · subst hi + simp + · subst hj + simp + rw [f.1.2]⟩ + left_inv f := by + match f with + | ⟨f, hf⟩ => + simp + ext i + by_cases h0 : f 0 = 0 + · simp [h0] + rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ + · subst hi + simp [h0] + · subst hj + simp [h0] + by_cases hj : f j.succ =0 + · rw [← h0] at hj + have hn := Function.Involutive.injective hf hj + exact False.elim (Fin.succ_ne_zero j hn) + · simp [hj] + rw [Fin.ext_iff] at hj + simp at hj + omega + · rw [if_neg h0] + by_cases hf' : i = f 0 + · subst hf' + simp + rw [hf] + simp + · rw [Function.update_apply, if_neg hf'] + rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ + · subst hi + simp + · subst hj + simp + by_cases hj : f j.succ =0 + · rw [← hj] at hf' + rw [hf] at hf' + simp at hf' + · simp [hj] + rw [Fin.ext_iff] at hj + simp at hj + omega + right_inv f := by + match f with + | ⟨⟨f, hf⟩, ⟨f0, hf0⟩⟩ => + ext i + · simp + by_cases hs : f0.isSome + · simp [hs] + by_cases hi : i = f0.get hs + · simp [hi, Function.update_apply] + exact Eq.symm (Fin.val_eq_of_eq (hf0 hs)) + · simp [hi] + split + · rename_i h + exact False.elim (Fin.succ_ne_zero (f i) h) + · rfl + · simp [hs] + split + · rename_i h + exact False.elim (Fin.succ_ne_zero (f i) h) + · rfl + · simp only [Nat.succ_eq_add_one, Option.mem_def, + Option.dite_none_left_eq_some, Option.some.injEq] + by_cases hs : f0.isSome + · simp only [hs, ↓reduceDIte] + simp [Fin.cons_zero] + have hx : ¬ (f0.get hs).succ = 0 := (Fin.succ_ne_zero (f0.get hs)) + simp [hx] + refine Iff.intro (fun hi => ?_) (fun hi => ?_) + · rw [← hi] + exact + Option.eq_some_of_isSome + (Eq.mpr_prop (Eq.refl (f0.isSome = true)) + (of_eq_true (Eq.trans (congrArg (fun x => x = true) hs) (eq_self true)))) + · subst hi + exact rfl + · simp [hs] + simp at hs + subst hs + exact ne_of_beq_false rfl + +def uncontractedFromInduction : {l : List 𝓕} → (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) → + List 𝓕 + | [], _ => [] + | φ :: φs, f => + let f' := involutionCons φs.length f + let c' := uncontractedFromInduction f'.1 + if f'.2.1.isSome then + c' + else + φ :: c' + +lemma uncontractedFromInduction_length : {l : List 𝓕} → (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) → + (uncontractedFromInduction f).length = ∑ i, if f.1 i = i then 1 else 0 + | [] => by + intro f + rfl + | φ :: φs => by + intro f + let f' := involutionCons φs.length f + let c' := uncontractedFromInduction f'.1 + by_cases h : f'.2.1.isSome + · dsimp [uncontractedFromInduction] + rw [if_pos h] + rw [uncontractedFromInduction_length f'.1] + rw [Fin.sum_univ_succ] + simp [f', involutionCons] at h + rw [if_neg h] + simp + sorry + · sorry + + + + +def uncontractedEquiv {l : List 𝓕} (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) : + {i : Option (Fin l.length) // + ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} ≃ + Option (Fin (uncontractedFromInduction f).length) := by + let e1 : {i : Option (Fin l.length) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} + ≃ Option {i : Fin l.length // f.1 i = i} := + { toFun := fun i => match i with + | ⟨some i, h⟩ => some ⟨i, by simpa using h⟩ + | ⟨none, h⟩ => none + invFun := fun i => match i with + | some ⟨i, h⟩ => ⟨some i, by simpa using h⟩ + | none => ⟨none, by simp⟩ + left_inv := by + intro a + cases a + aesop + right_inv := by + intro a + cases a + rfl + simp_all only [Subtype.coe_eta] } + let s : Finset (Fin l.length) := Finset.univ.filter fun i => f.1 i = i + let e2' : { i : Fin l.length // f.1 i = i} ≃ {i // i ∈ s} := by + refine Equiv.subtypeEquivProp ?h + funext i + simp [s] + let e2 : {i // i ∈ s} ≃ Fin (Finset.card s) := by + refine (Finset.orderIsoOfFin _ ?_).symm.toEquiv + simp [s] + have hcard : (uncontractedFromInduction f).length = Finset.card s := by + simp [s] + rw [Finset.card_filter] + rw [uncontractedFromInduction_length] + sorry + + + + +def involutionEquiv : (l : List 𝓕) → Contractions l ≃ + {f : Fin l.length → Fin l.length // Function.Involutive f} + | [] => { + toFun := fun c => ⟨fun i => i, fun i => rfl⟩, + invFun := fun f => ⟨[], ContractionsAux.nil⟩, + left_inv := by + intro a + cases a + rename_i aux c + cases c + rfl + right_inv := by + intro f + ext i + exact Fin.elim0 i} + | φ :: φs => by + refine Equiv.trans consEquiv ?_ + refine Equiv.trans ?_ (involutionCons φs.length).symm + refine Equiv.sigmaCongr (involutionEquiv φs) (fun c => ?_) + exact { + toFun := fun i? => ⟨Option.map c.embedUncontracted i?, by + intro h + + sorry⟩ + invFun := fun i => sorry + left_inv := by + sorry + right_inv := by sorry + } + /-- A structure specifying when a type `I` and a map `f :I → Type` corresponds to the splitting of a fields `φ` into a creation `n` and annihlation part `p`. -/ structure Splitting (f : 𝓕 → Type) [∀ i, Fintype (f i)] - (le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1] where + (le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le] where /-- The creation part of the fields. The label `n` corresponds to the fact that in normal ordering these feilds get pushed to the negative (left) direction. -/ 𝓑n : 𝓕 → (Σ i, f i) @@ -135,22 +691,23 @@ structure Splitting (f : 𝓕 → Type) [∀ i, Fintype (f i)] /-- The complex coefficent of annhilation part of a field `i`. This is usually `0` or `1`. -/ 𝓧p : 𝓕 → ℂ h𝓑 : ∀ i, ofListLift f [i] 1 = ofList [𝓑n i] (𝓧n i) + ofList [𝓑p i] (𝓧p i) - h𝓑n : ∀ i j, le1 (𝓑n i) j - h𝓑p : ∀ i j, le1 j (𝓑p i) + h𝓑n : ∀ i j, le (𝓑n i) j + h𝓑p : ∀ i j, le j (𝓑p i) /-- In the static wick's theorem, the term associated with a contraction `c` containing the contractions. -/ noncomputable def toCenterTerm (f : 𝓕 → Type) [∀ i, Fintype (f i)] (q : 𝓕 → FieldStatistic) - (le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1] + (le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le] {A : Type} [Semiring A] [Algebra ℂ A] (F : FreeAlgebra ℂ (Σ i, f i) →ₐ[ℂ] A) : - {r : List 𝓕} → (c : Contractions r) → (S : Splitting f le1) → A + {φs : List 𝓕} → (c : Contractions φs) → (S : Splitting f le) → A | [], ⟨[], .nil⟩, _ => 1 - | _ :: _, ⟨_, .cons (aux := aux') none c⟩, S => toCenterTerm f q le1 F ⟨aux', c⟩ S - | a :: _, ⟨_, .cons (aux := aux') (some n) c⟩, S => toCenterTerm f q le1 F ⟨aux', c⟩ S * + | _ :: _, ⟨_, .cons (φsᵤₙ := aux') none c⟩, S => toCenterTerm f q le F ⟨aux', c⟩ S + | a :: _, ⟨_, .cons (φsᵤₙ := aux') (some n) c⟩, S => toCenterTerm f q le F ⟨aux', c⟩ S * superCommuteCoef q [aux'.get n] (List.take (↑n) aux') • - F (((superCommute fun i => q i.fst) (ofList [S.𝓑p a] (S.𝓧p a))) (ofListLift f [aux'.get n] 1)) + F (((superCommute fun i => q i.fst) (ofList [S.𝓑p a] (S.𝓧p a))) + (ofListLift f [aux'.get n] 1)) /-- Shows that adding a field with no contractions (none) to an existing set of contractions results in the same center term as the original contractions. @@ -158,13 +715,13 @@ noncomputable def toCenterTerm (f : 𝓕 → Type) [∀ i, Fintype (f i)] Physically, this represents that an uncontracted (free) field does not affect the contraction structure of other fields in Wick's theorem. -/ lemma toCenterTerm_none (f : 𝓕 → Type) [∀ i, Fintype (f i)] - (q : 𝓕 → FieldStatistic) {r : List 𝓕} - (le1 : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le1] + (q : 𝓕 → FieldStatistic) {φs : List 𝓕} + (le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le] {A : Type} [Semiring A] [Algebra ℂ A] (F : FreeAlgebra ℂ (Σ i, f i) →ₐ A) - (S : Splitting f le1) (a : 𝓕) (c : Contractions r) : - toCenterTerm (r := a :: r) f q le1 F (Contractions.consEquiv.symm ⟨c, none⟩) S = - toCenterTerm f q le1 F c S := by + (S : Splitting f le) (φ : 𝓕) (c : Contractions φs) : + toCenterTerm (φs := φ :: φs) f q le F (Contractions.consEquiv.symm ⟨c, none⟩) S = + toCenterTerm f q le F c S := by rw [consEquiv] simp only [Equiv.coe_fn_symm_mk] dsimp [toCenterTerm] @@ -177,15 +734,15 @@ lemma toCenterTerm_center (f : 𝓕 → Type) [∀ i, Fintype (f i)] (le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le] {A : Type} [Semiring A] [Algebra ℂ A] (F : FreeAlgebra ℂ (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F] : - {r : List 𝓕} → (c : Contractions r) → (S : Splitting f le) → + {φs : List 𝓕} → (c : Contractions φs) → (S : Splitting f le) → (c.toCenterTerm f q le F S) ∈ Subalgebra.center ℂ A | [], ⟨[], .nil⟩, _ => by dsimp [toCenterTerm] exact Subalgebra.one_mem (Subalgebra.center ℂ A) - | _ :: _, ⟨_, .cons (aux := aux') none c⟩, S => by + | _ :: _, ⟨_, .cons (φsᵤₙ := aux') none c⟩, S => by dsimp [toCenterTerm] exact toCenterTerm_center f q le F ⟨aux', c⟩ S - | a :: _, ⟨_, .cons (aux := aux') (some n) c⟩, S => by + | a :: _, ⟨_, .cons (φsᵤₙ := aux') (some n) c⟩, S => by dsimp [toCenterTerm] refine Subalgebra.mul_mem (Subalgebra.center ℂ A) ?hx ?hy exact toCenterTerm_center f q le F ⟨aux', c⟩ S diff --git a/HepLean/PerturbationTheory/Wick/CreateAnnihilateSection.lean b/HepLean/PerturbationTheory/Wick/CreateAnnihilateSection.lean index f23315c8..1d43b95b 100644 --- a/HepLean/PerturbationTheory/Wick/CreateAnnihilateSection.lean +++ b/HepLean/PerturbationTheory/Wick/CreateAnnihilateSection.lean @@ -14,37 +14,37 @@ namespace Wick open HepLean.List open FieldStatistic -/-- The sections of `Σ i, f i` over a list `l : List 𝓕`. +/-- The sections of `Σ i, f i` over a list `φs : List 𝓕`. In terms of physics, given some fields `φ₁...φₙ`, the different ways one can associate each field as a `creation` or an `annilation` operator. E.g. the number of terms `φ₁⁰φ₂¹...φₙ⁰` `φ₁¹φ₂¹...φₙ⁰` etc. If some fields are exclusively creation or annhilation operators at this point (e.g. ansymptotic states) this is accounted for. -/ -def CreateAnnihilateSect {𝓕 : Type} (f : 𝓕 → Type) (l : List 𝓕) : Type := - Π i, f (l.get i) +def CreateAnnihilateSect {𝓕 : Type} (f : 𝓕 → Type) (φs : List 𝓕) : Type := + Π i, f (φs.get i) namespace CreateAnnihilateSect section basic_defs -variable {𝓕 : Type} {f : 𝓕 → Type} [∀ i, Fintype (f i)] {l : List 𝓕} (a : CreateAnnihilateSect f l) +variable {𝓕 : Type} {f : 𝓕 → Type} [∀ i, Fintype (f i)] {φs : List 𝓕} (a : CreateAnnihilateSect f φs) -/-- The type `CreateAnnihilateSect f l` is finite. -/ -instance fintype : Fintype (CreateAnnihilateSect f l) := Pi.fintype +/-- The type `CreateAnnihilateSect f φs` is finite. -/ +instance fintype : Fintype (CreateAnnihilateSect f φs) := Pi.fintype -/-- The section got by dropping the first element of `l` if it exists. -/ -def tail : {l : List 𝓕} → (a : CreateAnnihilateSect f l) → CreateAnnihilateSect f l.tail +/-- The section got by dropping the first element of `φs` if it exists. -/ +def tail : {φs : List 𝓕} → (a : CreateAnnihilateSect f φs) → CreateAnnihilateSect f φs.tail | [], a => a | _ :: _, a => fun i => a (Fin.succ i) /-- For a list of fields `i :: l` the value of the section at the head `i`. -/ -def head {i : 𝓕} (a : CreateAnnihilateSect f (i :: l)) : f i := a ⟨0, Nat.zero_lt_succ l.length⟩ +def head {φ : 𝓕} (a : CreateAnnihilateSect f (φ :: φs)) : f φ := a ⟨0, Nat.zero_lt_succ φs.length⟩ end basic_defs section toList_basic variable {𝓕 : Type} {f : 𝓕 → Type} (q : 𝓕 → FieldStatistic) - {l : List 𝓕} (a : CreateAnnihilateSect f l) + {φs : List 𝓕} (a : CreateAnnihilateSect f φs) /-- The list `List (Σ i, f i)` defined by `a`. -/ def toList : {l : List 𝓕} → (a : CreateAnnihilateSect f l) → List (Σ i, f i) @@ -52,8 +52,8 @@ def toList : {l : List 𝓕} → (a : CreateAnnihilateSect f l) → List (Σ i, | i :: _, a => ⟨i, a.head⟩ :: toList a.tail @[simp] -lemma toList_length : (toList a).length = l.length := by - induction l with +lemma toList_length : (toList a).length = φs.length := by + induction φs with | nil => rfl | cons i l ih => simp only [toList, List.length_cons, Fin.zero_eta] @@ -64,13 +64,13 @@ lemma toList_tail : {l : List 𝓕} → (a : CreateAnnihilateSect f l) → toLis | i :: l, a => by simp [toList] -lemma toList_cons {i : 𝓕} (a : CreateAnnihilateSect f (i :: l)) : +lemma toList_cons {i : 𝓕} (a : CreateAnnihilateSect f (i :: φs)) : (toList a) = ⟨i, a.head⟩ :: toList a.tail := by rfl -lemma toList_get (a : CreateAnnihilateSect f l) : - (toList a).get = (fun i => ⟨l.get i, a i⟩) ∘ Fin.cast (by simp) := by - induction l with +lemma toList_get (a : CreateAnnihilateSect f φs) : + (toList a).get = (fun i => ⟨φs.get i, a i⟩) ∘ Fin.cast (by simp) := by + induction φs with | nil => funext i exact Fin.elim0 i @@ -89,8 +89,8 @@ lemma toList_get (a : CreateAnnihilateSect f l) : @[simp] lemma toList_grade : FieldStatistic.ofList (fun i => q i.fst) a.toList = fermionic ↔ - FieldStatistic.ofList q l = fermionic := by - induction l with + FieldStatistic.ofList q φs = fermionic := by + induction φs with | nil => simp [toList] | cons i r ih => diff --git a/HepLean/PerturbationTheory/Wick/Signs/KoszulSignInsert.lean b/HepLean/PerturbationTheory/Wick/Signs/KoszulSignInsert.lean index 5406909d..1e8cfe85 100644 --- a/HepLean/PerturbationTheory/Wick/Signs/KoszulSignInsert.lean +++ b/HepLean/PerturbationTheory/Wick/Signs/KoszulSignInsert.lean @@ -20,57 +20,57 @@ variable {𝓕 : Type} (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Pro /-- Gives a factor of `-1` when inserting `a` into a list `List I` in the ordered position for each fermion-fermion cross. -/ def koszulSignInsert {𝓕 : Type} (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop) - [DecidableRel le] (a : 𝓕) : List 𝓕 → ℂ + [DecidableRel le] (φ : 𝓕) : List 𝓕 → ℂ | [] => 1 - | b :: l => if le a b then koszulSignInsert q le a l else - if q a = fermionic ∧ q b = fermionic then - koszulSignInsert q le a l else - koszulSignInsert q le a l + | φ' :: φs => if le φ φ' then koszulSignInsert q le φ φs else + if q φ = fermionic ∧ q φ' = fermionic then - koszulSignInsert q le φ φs else + koszulSignInsert q le φ φs /-- When inserting a boson the `koszulSignInsert` is always `1`. -/ lemma koszulSignInsert_boson (q : 𝓕 → FieldStatistic) (le : 𝓕 → 𝓕 → Prop) [DecidableRel le] - (a : 𝓕) (ha : q a = bosonic) : (φs : List 𝓕) → koszulSignInsert q le a φs = 1 + (φ : 𝓕) (ha : q φ = bosonic) : (φs : List 𝓕) → koszulSignInsert q le φ φs = 1 | [] => by simp [koszulSignInsert] - | b :: l => by + | φ' :: φs => by simp only [koszulSignInsert, Fin.isValue, ite_eq_left_iff] - rw [koszulSignInsert_boson q le a ha l, ha] + rw [koszulSignInsert_boson q le φ ha φs, ha] simp only [reduceCtorEq, false_and, ↓reduceIte, ite_self] @[simp] -lemma koszulSignInsert_mul_self (a : 𝓕) : - (φs : List 𝓕) → koszulSignInsert q le a φs * koszulSignInsert q le a φs = 1 +lemma koszulSignInsert_mul_self (φ : 𝓕) : + (φs : List 𝓕) → koszulSignInsert q le φ φs * koszulSignInsert q le φ φs = 1 | [] => by simp [koszulSignInsert] - | b :: l => by + | φ' :: φs => by simp only [koszulSignInsert, Fin.isValue, mul_ite, ite_mul, neg_mul, mul_neg] - by_cases hr : le a b + by_cases hr : le φ φ' · simp only [hr, ↓reduceIte] rw [koszulSignInsert_mul_self] · simp only [hr, ↓reduceIte] - by_cases hq : q a = fermionic ∧ q b = fermionic + by_cases hq : q φ = fermionic ∧ q φ' = fermionic · simp only [hq, and_self, ↓reduceIte, neg_neg] rw [koszulSignInsert_mul_self] · simp only [hq, ↓reduceIte] rw [koszulSignInsert_mul_self] -lemma koszulSignInsert_le_forall (a : 𝓕) (l : List 𝓕) (hi : ∀ b, le a b) : - koszulSignInsert q le a l = 1 := by - induction l with +lemma koszulSignInsert_le_forall (φ : 𝓕) (φs : List 𝓕) (hi : ∀ φ', le φ φ') : + koszulSignInsert q le φ φs = 1 := by + induction φs with | nil => rfl - | cons j l ih => + | cons φ' φs ih => simp only [koszulSignInsert, Fin.isValue, ite_eq_left_iff] rw [ih] simp only [Fin.isValue, ite_eq_left_iff, ite_eq_right_iff, and_imp] intro h - exact False.elim (h (hi j)) + exact False.elim (h (hi φ')) -lemma koszulSignInsert_ge_forall_append (φs : List 𝓕) (j i : 𝓕) (hi : ∀ j, le j i) : - koszulSignInsert q le j φs = koszulSignInsert q le j (φs ++ [i]) := by +lemma koszulSignInsert_ge_forall_append (φs : List 𝓕) (φ' φ : 𝓕) (hi : ∀ φ'', le φ'' φ) : + koszulSignInsert q le φ' φs = koszulSignInsert q le φ' (φs ++ [φ]) := by induction φs with | nil => simp [koszulSignInsert, hi] - | cons b l ih => + | cons φ'' φs ih => simp only [koszulSignInsert, Fin.isValue, List.append_eq] - by_cases hr : le j b + by_cases hr : le φ' φ'' · rw [if_pos hr, if_pos hr, ih] · rw [if_neg hr, if_neg hr, ih] @@ -136,20 +136,20 @@ lemma koszulSignInsert_eq_grade (φ : 𝓕) (φs : List 𝓕) : simpa [ofList] using ih simpa using hr1 -lemma koszulSignInsert_eq_perm (φs φs' : List 𝓕) (a : 𝓕) (h : φs.Perm φs') : - koszulSignInsert q le a φs = koszulSignInsert q le a φs' := by +lemma koszulSignInsert_eq_perm (φs φs' : List 𝓕) (φ : 𝓕) (h : φs.Perm φs') : + koszulSignInsert q le φ φs = koszulSignInsert q le φ φs' := by rw [koszulSignInsert_eq_grade, koszulSignInsert_eq_grade] congr 1 simp only [Fin.isValue, decide_not, eq_iff_iff, and_congr_right_iff] intro h' - have hg : ofList q (List.filter (fun i => !decide (le a i)) φs) = - ofList q (List.filter (fun i => !decide (le a i)) φs') := by + have hg : ofList q (List.filter (fun i => !decide (le φ i)) φs) = + ofList q (List.filter (fun i => !decide (le φ i)) φs') := by apply ofList_perm - exact List.Perm.filter (fun i => !decide (le a i)) h + exact List.Perm.filter (fun i => !decide (le φ i)) h rw [hg] -lemma koszulSignInsert_eq_sort (φs : List 𝓕) (a : 𝓕) : - koszulSignInsert q le a φs = koszulSignInsert q le a (List.insertionSort le φs) := by +lemma koszulSignInsert_eq_sort (φs : List 𝓕) (φ : 𝓕) : + koszulSignInsert q le φ φs = koszulSignInsert q le φ (List.insertionSort le φs) := by apply koszulSignInsert_eq_perm exact List.Perm.symm (List.perm_insertionSort le φs) diff --git a/HepLean/PerturbationTheory/Wick/StaticTheorem.lean b/HepLean/PerturbationTheory/Wick/StaticTheorem.lean index 1ba52d85..f94747ee 100644 --- a/HepLean/PerturbationTheory/Wick/StaticTheorem.lean +++ b/HepLean/PerturbationTheory/Wick/StaticTheorem.lean @@ -33,13 +33,13 @@ lemma static_wick_nil {A : Type} [Semiring A] [Algebra ℂ A] simp [ofListLift_empty] lemma static_wick_cons [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × f i) le] - {A : Type} [Semiring A] [Algebra ℂ A] (r : List 𝓕) (a : 𝓕) + {A : Type} [Semiring A] [Algebra ℂ A] (φs : List 𝓕) (φ : 𝓕) (F : FreeAlgebra ℂ (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F] (S : Contractions.Splitting f le) - (ih : F (ofListLift f r 1) = - ∑ c : Contractions r, c.toCenterTerm f q le F S * F (koszulOrder (fun i => q i.fst) le + (ih : F (ofListLift f φs 1) = + ∑ c : Contractions φs, c.toCenterTerm f q le F S * F (koszulOrder (fun i => q i.fst) le (ofListLift f c.normalize 1))) : - F (ofListLift f (a :: r) 1) = ∑ c : Contractions (a :: r), + F (ofListLift f (φ :: φs) 1) = ∑ c : Contractions (φ :: φs), c.toCenterTerm f q le F S * F (koszulOrder (fun i => q i.fst) le (ofListLift f c.normalize 1)) := by rw [ofListLift_cons_eq_ofListLift, map_mul, ih, Finset.mul_sum, @@ -47,7 +47,7 @@ lemma static_wick_cons [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × erw [Finset.sum_sigma] congr funext c - have hb := S.h𝓑 a + have hb := S.h𝓑 φ rw [← mul_assoc] have hi := c.toCenterTerm_center f q le F S rw [Subalgebra.mem_center_iff] at hi @@ -80,16 +80,16 @@ lemma static_wick_cons [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × funext n rw [← mul_assoc] rfl - exact S.h𝓑p a - exact S.h𝓑n a + exact S.h𝓑p φ + exact S.h𝓑n φ theorem static_wick_theorem [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × f i) le] - {A : Type} [Semiring A] [Algebra ℂ A] (r : List 𝓕) + {A : Type} [Semiring A] [Algebra ℂ A] (φs : List 𝓕) (F : FreeAlgebra ℂ (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F] (S : Contractions.Splitting f le) : - F (ofListLift f r 1) = ∑ c : Contractions r, c.toCenterTerm f q le F S * + F (ofListLift f φs 1) = ∑ c : Contractions φs, c.toCenterTerm f q le F S * F (koszulOrder (fun i => q i.fst) le (ofListLift f c.normalize 1)) := by - induction r with + induction φs with | nil => exact static_wick_nil q le F S | cons a r ih => exact static_wick_cons q le r a F S ih From eadb35447732dbd0ac584b4a198a1ff1632b74ea Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Sat, 4 Jan 2025 11:56:26 +0000 Subject: [PATCH 2/8] feat: equivalence between involutions and contractions --- HepLean/Mathematics/List.lean | 10 + .../Contractions/Basic.lean | 475 +++++++++++++++--- 2 files changed, 417 insertions(+), 68 deletions(-) diff --git a/HepLean/Mathematics/List.lean b/HepLean/Mathematics/List.lean index 28a06500..251e5e57 100644 --- a/HepLean/Mathematics/List.lean +++ b/HepLean/Mathematics/List.lean @@ -741,4 +741,14 @@ lemma optionEraseZ_some_length {I : Type} (l : List I) (a : I) (i : (Fin l.lengt (optionEraseZ l a (some i)).length = l.length - 1 := by simp [optionEraseZ, List.length_eraseIdx] +lemma optionEraseZ_ext {I : Type} {l l' : List I} {a a' : I} {i : Option (Fin l.length)} + {i' : Option (Fin l'.length)} (hl : l = l') (ha : a = a') + (hi : Option.map (Fin.cast (by rw [hl])) i = i') : + optionEraseZ l a i = optionEraseZ l' a' i' := by + subst hl + subst ha + cases hi + congr + simp + end HepLean.List diff --git a/HepLean/PerturbationTheory/Contractions/Basic.lean b/HepLean/PerturbationTheory/Contractions/Basic.lean index f59a5cf0..9ba10dcb 100644 --- a/HepLean/PerturbationTheory/Contractions/Basic.lean +++ b/HepLean/PerturbationTheory/Contractions/Basic.lean @@ -124,6 +124,14 @@ def consEquiv {φ : 𝓕} {φs : List 𝓕} : | ContractionsAux.cons (φsᵤₙ := aux') i c => rfl right_inv ci := by rfl +lemma consEquiv_ext {φs : List 𝓕} {c1 c2 : Contractions φs} + {n1 : Option (Fin c1.normalize.length)} {n2 : Option (Fin c2.normalize.length)} + (hc : c1 = c2) (hn : Option.map (finCongr (by rw [hc])) n1 = n2) : + (⟨c1, n1⟩ : (c : Contractions φs) × Option (Fin c.normalize.length)) = ⟨c2, n2⟩ := by + subst hc + subst hn + simp + /-- The type of contractions is decidable. -/ instance decidable : (φs : List 𝓕) → DecidableEq (Contractions φs) | [] => fun a b => @@ -574,44 +582,24 @@ def involutionCons (n : ℕ): subst hs exact ne_of_beq_false rfl -def uncontractedFromInduction : {l : List 𝓕} → (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) → - List 𝓕 - | [], _ => [] - | φ :: φs, f => - let f' := involutionCons φs.length f - let c' := uncontractedFromInduction f'.1 - if f'.2.1.isSome then - c' - else - φ :: c' - -lemma uncontractedFromInduction_length : {l : List 𝓕} → (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) → - (uncontractedFromInduction f).length = ∑ i, if f.1 i = i then 1 else 0 - | [] => by - intro f - rfl - | φ :: φs => by - intro f - let f' := involutionCons φs.length f - let c' := uncontractedFromInduction f'.1 - by_cases h : f'.2.1.isSome - · dsimp [uncontractedFromInduction] - rw [if_pos h] - rw [uncontractedFromInduction_length f'.1] - rw [Fin.sum_univ_succ] - simp [f', involutionCons] at h - rw [if_neg h] - simp - sorry - · sorry - - - +lemma involutionCons_ext {n : ℕ} {f1 f2 : (f : {f : Fin n → Fin n // Function.Involutive f}) × {i : Option (Fin n) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)}} + (h1 : f1.1 = f2.1) (h2 : f1.2 = Equiv.subtypeEquivRight (by rw [h1]; simp) f2.2) : f1 = f2 := by + cases f1 + cases f2 + simp at h1 h2 + subst h1 + rename_i fst snd snd_1 + simp_all only [Sigma.mk.inj_iff, heq_eq_eq, true_and] + obtain ⟨val, property⟩ := fst + obtain ⟨val_1, property_1⟩ := snd + obtain ⟨val_2, property_2⟩ := snd_1 + simp_all only + rfl -def uncontractedEquiv {l : List 𝓕} (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) : +def uncontractedEquiv' {l : List 𝓕} (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) : {i : Option (Fin l.length) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} ≃ - Option (Fin (uncontractedFromInduction f).length) := by + Option (Fin (Finset.univ.filter fun i => f.1 i = i).card) := by let e1 : {i : Option (Fin l.length) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} ≃ Option {i : Fin l.length // f.1 i = i} := { toFun := fun i => match i with @@ -637,44 +625,395 @@ def uncontractedEquiv {l : List 𝓕} (f : {f : Fin l.length → Fin l.length // let e2 : {i // i ∈ s} ≃ Fin (Finset.card s) := by refine (Finset.orderIsoOfFin _ ?_).symm.toEquiv simp [s] - have hcard : (uncontractedFromInduction f).length = Finset.card s := by - simp [s] - rw [Finset.card_filter] - rw [uncontractedFromInduction_length] - sorry + refine e1.trans (Equiv.optionCongr (e2'.trans (e2))) +lemma uncontractedEquiv'_none_image_zero {φs : List 𝓕} {φ : 𝓕} : + {f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}} + → uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2 = none + → f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ = ⟨0, Nat.zero_lt_succ φs.length⟩ := by + intro f h + simp only [Nat.succ_eq_add_one, involutionCons, Equiv.coe_fn_mk, uncontractedEquiv', + Option.isSome_some, Option.get_some, Option.isSome_none, Equiv.trans_apply, + Equiv.optionCongr_apply, Equiv.coe_trans, RelIso.coe_fn_toEquiv, Option.map_eq_none'] at h + simp_all only [List.length_cons, Fin.zero_eta] + obtain ⟨val, property⟩ := f + simp_all only [List.length_cons] + split at h + next i i_1 h_1 heq => + split at heq + next h_2 => simp_all only [reduceCtorEq] + next h_2 => simp_all only [reduceCtorEq] + next i h_1 heq => + split at heq + next h_2 => simp_all only + next h_2 => simp_all only [Subtype.mk.injEq, reduceCtorEq] +lemma uncontractedEquiv'_cast {l : List 𝓕} {f1 f2 : {f : Fin l.length → Fin l.length // Function.Involutive f}} + (hf : f1 = f2): + uncontractedEquiv' f1 = (Equiv.subtypeEquivRight (by rw [hf]; simp)).trans + ((uncontractedEquiv' f2).trans (Equiv.optionCongr (finCongr (by rw [hf])))):= by + subst hf + simp + rfl -def involutionEquiv : (l : List 𝓕) → Contractions l ≃ - {f : Fin l.length → Fin l.length // Function.Involutive f} - | [] => { - toFun := fun c => ⟨fun i => i, fun i => rfl⟩, - invFun := fun f => ⟨[], ContractionsAux.nil⟩, - left_inv := by - intro a - cases a - rename_i aux c - cases c +lemma uncontractedEquiv'_none_succ {φs : List 𝓕} {φ : 𝓕} : + {f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}} + → uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2 = none + → ∀ (x : Fin φs.length), f.1 x.succ = x.succ ↔ (involutionCons φs.length f).1.1 x = x := by + intro f h x + simp [involutionCons] + have hn' := uncontractedEquiv'_none_image_zero h + have hx : ¬ f.1 x.succ = ⟨0, Nat.zero_lt_succ φs.length⟩:= by + rw [← hn'] + exact fun hn => Fin.succ_ne_zero x (Function.Involutive.injective f.2 hn) + apply Iff.intro + · intro h2 h3 + rw [Fin.ext_iff] + simp [h2] + · intro h2 + have h2' := h2 hx + conv_rhs => rw [← h2'] + simp + + +lemma uncontractedEquiv'_isSome_image_zero {φs : List 𝓕} {φ : 𝓕} : + {f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}} + → (uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2).isSome + → ¬ f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ = ⟨0, Nat.zero_lt_succ φs.length⟩ := by + intro f hf + simp [uncontractedEquiv', involutionCons] at hf + simp_all only [List.length_cons, Fin.zero_eta] + obtain ⟨val, property⟩ := f + simp_all only [List.length_cons] + apply Aesop.BuiltinRules.not_intro + intro a + simp_all only [↓reduceDIte, Option.isSome_none, Bool.false_eq_true] + + + +def uncontractedFromInvolution: {φs : List 𝓕} → + (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) → + {l : List 𝓕 // l.length = (Finset.univ.filter fun i => f.1 i = i).card} + | [], _ => ⟨[], by simp⟩ + | φ :: φs, f => + let luc := uncontractedFromInvolution (involutionCons φs.length f).fst + let n' := uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2 + if hn : n' = none then + have hn' := uncontractedEquiv'_none_image_zero (φs := φs) (φ := φ) (f := f) hn + ⟨optionEraseZ luc φ none, by + simp [optionEraseZ] + rw [← luc.2] + conv_rhs => rw [Finset.card_filter] + rw [Fin.sum_univ_succ] + conv_rhs => erw [if_pos hn'] + ring_nf + simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, Nat.cast_id, + add_right_inj] + rw [Finset.card_filter] + apply congrArg + funext i + refine ite_congr ?h.h.h₁ (congrFun rfl) (congrFun rfl) + rw [uncontractedEquiv'_none_succ hn]⟩ + else + let n := n'.get (Option.isSome_iff_ne_none.mpr hn) + let np : Fin luc.1.length := ⟨n.1, by + rw [luc.2] + exact n.prop⟩ + ⟨optionEraseZ luc φ (some np), by + let k' := (involutionCons φs.length f).2 + have hkIsSome : (k'.1).isSome := by + simp [n', uncontractedEquiv' ] at hn + split at hn + · simp_all only [reduceCtorEq, not_false_eq_true, Nat.succ_eq_add_one, Option.isSome_some, k'] + · simp_all only [not_true_eq_false] + let k := k'.1.get hkIsSome + rw [optionEraseZ_some_length] + have hksucc : k.succ = f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ := by + simp [k, k', involutionCons] + have hzero : ⟨0, Nat.zero_lt_succ φs.length⟩ = f.1 k.succ := by + rw [hksucc] + rw [f.2] + have hkcons : ((involutionCons φs.length) f).1.1 k = k := by + exact k'.2 hkIsSome + have hksuccNe : f.1 k.succ ≠ k.succ := by + conv_rhs => rw [hksucc] + exact fun hn => Fin.succ_ne_zero k (Function.Involutive.injective f.2 hn ) + have hluc : 1 ≤ luc.1.length := by + simp + use k + simp [involutionCons] + rw [hksucc, f.2] + simp + rw [propext (Nat.sub_eq_iff_eq_add' hluc)] + have h0 : ¬ f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ = ⟨0, Nat.zero_lt_succ φs.length⟩ := by + exact Option.isSome_dite'.mp hkIsSome + conv_rhs => + rw [Finset.card_filter] + erw [Fin.sum_univ_succ] + erw [if_neg h0] + simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, List.length_cons, + Nat.cast_id, zero_add] + conv_rhs => lhs; rw [Eq.symm (Fintype.sum_ite_eq' k fun j => 1)] + rw [← Finset.sum_add_distrib] + rw [Finset.card_filter] + apply congrArg + funext i + by_cases hik : i = k + · subst hik + simp [hkcons, hksuccNe] + · simp [hik] + refine ite_congr ?_ (congrFun rfl) (congrFun rfl) + simp [involutionCons] + have hfi : f.1 i.succ ≠ ⟨0, Nat.zero_lt_succ φs.length⟩ := by + rw [hzero] + by_contra hn + have hik' := (Function.Involutive.injective f.2 hn) + simp only [List.length_cons, Fin.succ_inj] at hik' + exact hik hik' + apply Iff.intro + · intro h + have h' := h hfi + conv_rhs => rw [← h'] + simp + · intro h hfi + simp [Fin.ext_iff] + rw [h] + simp⟩ + +lemma uncontractedFromInvolution_cons {φs : List 𝓕} {φ : 𝓕} + (f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) : + uncontractedFromInvolution f = + optionEraseZ (uncontractedFromInvolution (involutionCons φs.length f).fst) φ + (Option.map (finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm)) + (uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2)) := by + let luc := uncontractedFromInvolution (involutionCons φs.length f).fst + let n' := uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2 + change _ = optionEraseZ luc φ + (Option.map (finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm)) n') + dsimp [uncontractedFromInvolution] + by_cases hn : n' = none + · have hn' := hn + simp [n'] at hn' + simp [hn'] + rw [hn] + rfl + · have hn' := hn + simp [n'] at hn' + simp [hn'] + congr + simp [n'] + simp_all only [Nat.succ_eq_add_one, not_false_eq_true, n', luc] + obtain ⟨val, property⟩ := f + obtain ⟨val_1, property_1⟩ := luc + simp_all only [Nat.succ_eq_add_one, List.length_cons] + ext a : 1 + simp_all only [Option.mem_def, Option.some.injEq, Option.map_eq_some', finCongr_apply] + apply Iff.intro + · intro a_1 + subst a_1 + apply Exists.intro + · apply And.intro + on_goal 2 => {rfl + } + · simp_all only [Option.some_get] + · intro a_1 + obtain ⟨w, h⟩ := a_1 + obtain ⟨left, right⟩ := h + subst right + simp_all only [Option.get_some] rfl - right_inv := by - intro f - ext i - exact Fin.elim0 i} - | φ :: φs => by - refine Equiv.trans consEquiv ?_ - refine Equiv.trans ?_ (involutionCons φs.length).symm - refine Equiv.sigmaCongr (involutionEquiv φs) (fun c => ?_) - exact { - toFun := fun i? => ⟨Option.map c.embedUncontracted i?, by - intro h - - sorry⟩ - invFun := fun i => sorry - left_inv := by - sorry - right_inv := by sorry - } + +def auxCongr : {φs: List 𝓕} → {φsᵤₙ φsᵤₙ' : List 𝓕} → (h : φsᵤₙ = φsᵤₙ') → + ContractionsAux φs φsᵤₙ ≃ ContractionsAux φs φsᵤₙ' + | _, _, _, Eq.refl _ => Equiv.refl _ + +lemma auxCongr_ext {φs: List 𝓕} {c c2 : Contractions φs} (h : c.1 = c2.1) + (hx : c.2 = auxCongr h.symm c2.2) : c = c2 := by + cases c + cases c2 + simp at h + subst h + simp [auxCongr] at hx + subst hx + rfl + + +lemma uncontractedEquiv'_cast' {l : List 𝓕} {f1 f2 : {f : Fin l.length → Fin l.length // Function.Involutive f}} + {N : ℕ} (hf : f1 = f2) (n : Option (Fin N)) (hn1 : N = (Finset.filter (fun i => f1.1 i = i) Finset.univ).card) + (hn2 : N = (Finset.filter (fun i => f2.1 i = i) Finset.univ).card): + HEq ((uncontractedEquiv' f1).symm (Option.map (finCongr hn1) n)) ((uncontractedEquiv' f2).symm (Option.map (finCongr hn2) n)) := by + subst hf + rfl + +def toInvolution' : {φs : List 𝓕} → (c : Contractions φs) → + {f : {f : Fin φs.length → Fin φs.length // Function.Involutive f} // + uncontractedFromInvolution f = c.1} + | [], ⟨[], ContractionsAux.nil⟩ => ⟨⟨fun i => i, by + intro i + simp⟩, by rfl⟩ + | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) n c⟩ => by + let ⟨⟨f', hf1⟩, hf2⟩ := toInvolution' ⟨aux, c⟩ + let n' : Option (Fin (uncontractedFromInvolution ⟨f', hf1⟩).1.length) := + Option.map (finCongr (by rw [hf2])) n + let F := (involutionCons φs.length).symm ⟨⟨f', hf1⟩, + (uncontractedEquiv' ⟨f', hf1⟩).symm + (Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ + refine ⟨F, ?_⟩ + have hF0 : ((involutionCons φs.length) F) = ⟨⟨f', hf1⟩, + (uncontractedEquiv' ⟨f', hf1⟩).symm + (Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ := by + simp [F] + have hF1 : ((involutionCons φs.length) F).fst = ⟨f', hf1⟩ := by + rw [hF0] + have hF2L : ((uncontractedFromInvolution ⟨f', hf1⟩)).1.length = + (Finset.filter (fun i => ((involutionCons φs.length) F).1.1 i = i) Finset.univ).card := by + apply Eq.trans ((uncontractedFromInvolution ⟨f', hf1⟩)).2 + congr + rw [hF1] + have hF2 : ((involutionCons φs.length) F).snd = (uncontractedEquiv' ((involutionCons φs.length) F).fst).symm + (Option.map (finCongr hF2L) n') := by + rw [@Sigma.subtype_ext_iff] at hF0 + ext1 + rw [hF0.2] + simp + congr 1 + · rw [hF1] + · refine uncontractedEquiv'_cast' ?_ n' _ _ + rw [hF1] + rw [uncontractedFromInvolution_cons] + have hx := (toInvolution' ⟨aux, c⟩).2 + simp at hx + simp + refine optionEraseZ_ext ?_ ?_ ?_ + · dsimp [F] + rw [Equiv.apply_symm_apply] + simp + rw [← hx] + simp_all only + · rfl + · simp [hF2] + dsimp [n'] + simp [finCongr] + simp only [Nat.succ_eq_add_one, id_eq, eq_mpr_eq_cast, F, n'] + ext a : 1 + simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans, + Fin.cast_eq_self, exists_eq_right] + +lemma toInvolution'_length {φs φsᵤₙ : List 𝓕} {c : ContractionsAux φs φsᵤₙ} : + φsᵤₙ.length = (Finset.filter (fun i => (toInvolution' ⟨φsᵤₙ, c⟩).1.1 i = i) Finset.univ).card + := by + have h2 := (toInvolution' ⟨φsᵤₙ, c⟩).2 + simp at h2 + conv_lhs => rw [← h2] + exact Mathlib.Vector.length_val (uncontractedFromInvolution (toInvolution' ⟨φsᵤₙ, c⟩).1) + +lemma toInvolution'_cons {φs φsᵤₙ : List 𝓕} {φ : 𝓕} + (c : ContractionsAux φs φsᵤₙ) (n : Option (Fin (φsᵤₙ.length))) : + (toInvolution' ⟨optionEraseZ φsᵤₙ φ n, ContractionsAux.cons n c⟩).1 + = (involutionCons φs.length).symm ⟨(toInvolution' ⟨φsᵤₙ, c⟩).1, + (uncontractedEquiv' (toInvolution' ⟨φsᵤₙ, c⟩).1).symm + (Option.map (finCongr (toInvolution'_length)) n)⟩ := by + dsimp [toInvolution'] + congr 3 + rw [Option.map_map] + simp [finCongr] + rfl + +lemma toInvolution_consEquiv {φs : List 𝓕} {φ : 𝓕} + (c : Contractions φs) (n : Option (Fin (c.normalize.length))) : + (toInvolution' ((consEquiv (φ := φ)).symm ⟨c, n⟩)).1 = + (involutionCons φs.length).symm ⟨(toInvolution' c).1, + (uncontractedEquiv' (toInvolution' c).1).symm + (Option.map (finCongr (toInvolution'_length)) n)⟩ := by + erw [toInvolution'_cons] + rfl + +def fromInvolutionAux : {l : List 𝓕} → + (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) → + ContractionsAux l (uncontractedFromInvolution f) + | [] => fun _ => ContractionsAux.nil + | _ :: φs => fun f => + let f' := involutionCons φs.length f + let c' := fromInvolutionAux f'.1 + let n' := Option.map (finCongr ((uncontractedFromInvolution f'.fst).2.symm)) + (uncontractedEquiv' f'.1 f'.2) + auxCongr (uncontractedFromInvolution_cons f).symm (ContractionsAux.cons n' c') + +def fromInvolution {φs : List 𝓕} (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) : + Contractions φs := ⟨uncontractedFromInvolution f, fromInvolutionAux f⟩ + +lemma fromInvolution_cons {φs : List 𝓕} {φ : 𝓕} + (f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) : + let f' := involutionCons φs.length f + fromInvolution f = consEquiv.symm + ⟨fromInvolution f'.1, Option.map (finCongr ((uncontractedFromInvolution f'.fst).2.symm)) + (uncontractedEquiv' f'.1 f'.2)⟩ := by + refine auxCongr_ext ?_ ?_ + · dsimp [fromInvolution] + rw [uncontractedFromInvolution_cons] + rfl + · dsimp [fromInvolution, fromInvolutionAux] + rfl + +lemma fromInvolution_of_involutionCons + {φs : List 𝓕} {φ : 𝓕} + (f : {f : Fin (φs ).length → Fin (φs).length // Function.Involutive f}) + (n : { i : Option (Fin φs.length) // ∀ (h : i.isSome = true), f.1 (i.get h) = i.get h }): + fromInvolution (φs := φ :: φs) ((involutionCons φs.length).symm ⟨f, n⟩) = + consEquiv.symm + ⟨fromInvolution f, Option.map (finCongr ((uncontractedFromInvolution f).2.symm)) + (uncontractedEquiv' f n)⟩ := by + rw [fromInvolution_cons] + congr 1 + simp + rw [Equiv.apply_symm_apply] + + + + +lemma toInvolution_fromInvolution : {φs : List 𝓕} → (c : Contractions φs) → + fromInvolution (toInvolution' c) = c + | [], ⟨[], ContractionsAux.nil⟩ => rfl + | φ :: φs, ⟨_, .cons (φsᵤₙ := φsᵤₙ) n c⟩ => by + rw [toInvolution'_cons] + rw [fromInvolution_of_involutionCons] + rw [Equiv.symm_apply_eq] + dsimp [consEquiv] + refine consEquiv_ext ?_ ?_ + · exact toInvolution_fromInvolution ⟨φsᵤₙ, c⟩ + · simp [finCongr] + ext a : 1 + simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans, + Fin.cast_eq_self, exists_eq_right] + +lemma fromInvolution_toInvolution : {φs : List 𝓕} → (f : {f : Fin (φs ).length → Fin (φs).length // Function.Involutive f}) + → toInvolution' (fromInvolution f) = f + | [], _ => by + ext x + exact Fin.elim0 x + | φ :: φs, f => by + rw [fromInvolution_cons] + rw [toInvolution_consEquiv] + erw [Equiv.symm_apply_eq] + have hx := fromInvolution_toInvolution ((involutionCons φs.length) f).fst + apply involutionCons_ext ?_ ?_ + · simp only [Nat.succ_eq_add_one, List.length_cons] + exact hx + · simp only [Nat.succ_eq_add_one, Option.map_map, List.length_cons] + rw [Equiv.symm_apply_eq] + conv_rhs => + lhs + rw [uncontractedEquiv'_cast hx] + simp [Nat.succ_eq_add_one,- eq_mpr_eq_cast, Equiv.trans_apply, -Equiv.optionCongr_apply] + rfl + +def equivInvolutions {φs : List 𝓕} : + Contractions φs ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f} where + toFun := fun c => toInvolution' c + invFun := fromInvolution + left_inv := toInvolution_fromInvolution + right_inv := fromInvolution_toInvolution /-- A structure specifying when a type `I` and a map `f :I → Type` corresponds to the splitting of a fields `φ` into a creation `n` and annihlation part `p`. -/ From 840d16a581a7cfa08b79f3f33bec11d3545abe3b Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Sat, 4 Jan 2025 14:22:58 +0000 Subject: [PATCH 3/8] feat: number of full contractions --- .../Contractions/Basic.lean | 324 ++++++++++++++++++ 1 file changed, 324 insertions(+) diff --git a/HepLean/PerturbationTheory/Contractions/Basic.lean b/HepLean/PerturbationTheory/Contractions/Basic.lean index 9ba10dcb..2e301072 100644 --- a/HepLean/PerturbationTheory/Contractions/Basic.lean +++ b/HepLean/PerturbationTheory/Contractions/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith -/ import HepLean.PerturbationTheory.Wick.OperatorMap +import Mathlib.Data.Nat.Factorial.DoubleFactorial /-! # Contractions of a list of fields @@ -1015,6 +1016,329 @@ def equivInvolutions {φs : List 𝓕} : left_inv := toInvolution_fromInvolution right_inv := fromInvolution_toInvolution + +lemma isFull_iff_uncontractedFromInvolution_empty {φs : List 𝓕} (c : Contractions φs) : + IsFull c ↔ (uncontractedFromInvolution (equivInvolutions c)).1 = [] := by + let l := toInvolution' c + erw [l.2] + rfl + +lemma isFull_iff_filter_card_involution_zero {φs : List 𝓕} (c : Contractions φs) : + IsFull c ↔ (Finset.univ.filter fun i => (equivInvolutions c).1 i = i).card = 0 := by + rw [isFull_iff_uncontractedFromInvolution_empty, List.ext_get_iff] + simp + +lemma isFull_iff_involution_no_fixed_points {φs : List 𝓕} (c : Contractions φs) : + IsFull c ↔ ∀ (i : Fin φs.length), (equivInvolutions c).1 i ≠ i := by + rw [isFull_iff_filter_card_involution_zero] + simp + rw [Finset.filter_eq_empty_iff] + apply Iff.intro + · intro h + intro i + refine h (Finset.mem_univ i) + · intro i h + exact fun a => i h + +def involutionNoFixed1 {n : ℕ} : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f + ∧ ∀ i, f i ≠ i} ≃ Σ (k : Fin (2 * n + 1)), + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f + ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} where + toFun f := ⟨(f.1 0).pred (f.2.2 0), ⟨f.1, f.2.1, by simpa using f.2.2⟩⟩ + invFun f := ⟨f.2.1, ⟨f.2.2.1, f.2.2.2.1⟩⟩ + left_inv f := by + rfl + right_inv f := by + simp + ext + · simp + rw [f.2.2.2.2] + simp + · simp + +def involutionNoFixed2 {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f + ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} ≃ + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive (e.symm ∘ f ∘ e) ∧ + (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} where + toFun f := ⟨e ∘ f.1 ∘ e.symm, by + intro i + simp + rw [f.2.1], by + simpa using f.2.2.1, by simpa using f.2.2.2⟩ + invFun f := ⟨e.symm ∘ f.1 ∘ e, by + intro i + simp + have hf2 := f.2.1 i + simpa using hf2, by + simpa using f.2.2.1, by simpa using f.2.2.2⟩ + left_inv f := by + ext i + simp + right_inv f := by + ext i + simp +def involutionNoFixed3 {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive (e.symm ∘ f ∘ e) ∧ + (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} ≃ + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} := by + refine Equiv.subtypeEquivRight ?_ + intro f + have h1 : Function.Involutive (⇑e.symm ∘ f ∘ ⇑e) ↔ Function.Involutive f := by + apply Iff.intro + · intro h i + have hi := h (e.symm i) + simpa using hi + · intro h i + have hi := h (e i) + simp [hi] + rw [h1] + simp + intro h1 h2 + apply Iff.intro + · intro h i + have hi := h (e.symm i) + simpa using hi + · intro h i + have hi := h (e i) + by_contra hn + nth_rewrite 2 [← hn] at hi + simp at hi + + +def involutionNoFixed4 {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} + ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f (e 0) = e k.succ} := by + refine Equiv.subtypeEquivRight ?_ + simp + intro f hi h1 + exact Equiv.symm_apply_eq e + +def involutionNoFixed5 {n : ℕ} (k : Fin (2 * n + 1)) : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f 0 = k.succ} + ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f 0 = 1} := by + refine Equiv.trans (involutionNoFixed2 k (Equiv.swap k.succ 1)) ?_ + refine Equiv.trans (involutionNoFixed3 k (Equiv.swap k.succ 1)) ?_ + refine Equiv.trans (involutionNoFixed4 k (Equiv.swap k.succ 1)) ?_ + refine Equiv.subtypeEquivRight ?_ + simp + intro f hi h1 + rw [Equiv.swap_apply_of_ne_of_ne] + · exact Ne.symm (Fin.succ_ne_zero k) + · exact Fin.zero_ne_one + +def involutionNoFixed6 {n : ℕ} : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f 0 = 1} ≃ + {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ + (∀ i, f i ≠ i)} where + toFun f := by + have hf1 : f.1 1 = 0 := by + have hf := f.2.2.2 + simp [← hf] + rw [f.2.1] + let f' := f.1 ∘ Fin.succ ∘ Fin.succ + have hf' (i : Fin (2 * n)) : f' i ≠ 0 := by + simp [f'] + simp [← hf1] + by_contra hn + have hn' := Function.Involutive.injective f.2.1 hn + simp [Fin.ext_iff] at hn' + let f'' := fun i => (f' i).pred (hf' i) + have hf'' (i : Fin (2 * n)) : f'' i ≠ 0 := by + simp [f''] + rw [@Fin.pred_eq_iff_eq_succ] + simp [f'] + simp [← f.2.2.2 ] + by_contra hn + have hn' := Function.Involutive.injective f.2.1 hn + simp [Fin.ext_iff] at hn' + let f''' := fun i => (f'' i).pred (hf'' i) + refine ⟨f''', ?_, ?_⟩ + · intro i + simp [f''', f'', f'] + simp [f.2.1 i.succ.succ] + · intro i + simp [f''', f'', f'] + rw [@Fin.pred_eq_iff_eq_succ] + rw [@Fin.pred_eq_iff_eq_succ] + exact f.2.2.1 i.succ.succ + invFun f := by + let f' := fun (i : Fin (2 * n.succ))=> + match i with + | ⟨0, h⟩ => 1 + | ⟨1, h⟩ => 0 + | ⟨(Nat.succ (Nat.succ n)), h⟩ => (f.1 ⟨n, by omega⟩).succ.succ + refine ⟨f', ?_, ?_, ?_⟩ + · intro i + match i with + | ⟨0, h⟩ => rfl + | ⟨1, h⟩ => rfl + | ⟨(Nat.succ (Nat.succ m)), h⟩ => + simp [f'] + split + · rename_i h + simp at h + exact False.elim (Fin.succ_ne_zero (f.1 ⟨m, _⟩).succ h) + · rename_i h + simp [Fin.ext_iff] at h + · rename_i h + rename_i x r + simp_all [Fin.ext_iff] + have hfn {a b : ℕ} {ha : a < 2 * n} {hb : b < 2 * n} + (hab : ↑(f.1 ⟨a, ha⟩) = b): ↑(f.1 ⟨b, hb⟩) = a := by + have ht : f.1 ⟨a, ha⟩ = ⟨b, hb⟩ := by + simp [hab, Fin.ext_iff] + rw [← ht, f.2.1] + exact hfn h + · intro i + match i with + | ⟨0, h⟩ => + simp [f'] + split + · rename_i h + simp + · rename_i h + simp [Fin.ext_iff] at h + · rename_i h + simp [Fin.ext_iff] at h + | ⟨1, h⟩ => + simp [f'] + split + · rename_i h + simp at h + · rename_i h + simp + · rename_i h + simp [Fin.ext_iff] at h + | ⟨(Nat.succ (Nat.succ m)), h⟩ => + simp [f', Fin.ext_iff] + have hf:= f.2.2 ⟨m, by exact Nat.add_lt_add_iff_right.mp h⟩ + simp [Fin.ext_iff] at hf + omega + · simp [f'] + split + · rename_i h + simp + · rename_i h + simp at h + · rename_i h + simp [Fin.ext_iff] at h + left_inv f := by + have hf1 : f.1 1 = 0 := by + have hf := f.2.2.2 + simp [← hf] + rw [f.2.1] + simp + ext i + simp + split + · simp + rw [f.2.2.2] + simp + · simp + rw [hf1] + simp + · rfl + right_inv f := by + simp + ext i + simp + split + · rename_i h + simp [Fin.ext_iff] at h + · rename_i h + simp [Fin.ext_iff] at h + · rename_i h + simp + congr + apply congrArg + simp_all [Fin.ext_iff] + + + +def involutionNoFixed7 {n : ℕ} (k : Fin (2 * n + 1)) : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f 0 = k.succ} + ≃ {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + refine Equiv.trans (involutionNoFixed5 k) involutionNoFixed6 + +def involutionNoFixed8 {n : ℕ} : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + ≃ Σ (_ : Fin (2 * n + 1)), {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + refine Equiv.trans involutionNoFixed1 ?_ + refine Equiv.sigmaCongrRight involutionNoFixed7 + +def involutionNoFixed9 {n : ℕ} : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + ≃ Fin (2 * n + 1) × {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + refine Equiv.trans involutionNoFixed8 ?_ + exact Equiv.sigmaEquivProd (Fin (2 * n + 1)) + { f // Function.Involutive f ∧ ∀ (i : Fin (2 * n)), f i ≠ i} + +instance {n : ℕ} : Fintype { f // Function.Involutive f ∧ ∀ (i : Fin ( n)), f i ≠ i } := by + haveI : DecidablePred fun x => Function.Involutive x := by + intro f + apply Fintype.decidableForallFintype (α := Fin (n)) + haveI : DecidablePred fun x => Function.Involutive x ∧ ∀ (i : Fin ( n)), x i ≠ i := by + intro x + apply instDecidableAnd + apply Subtype.fintype + +lemma involutionNoFixed_card_succ {n : ℕ} : + Fintype.card {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + = (2 * n + 1) * Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + rw [Fintype.card_congr (involutionNoFixed9)] + rw [Fintype.card_prod ] + congr + exact Fintype.card_fin (2 * n + 1) + + +open Nat + + +lemma involutionNoFixed_card : (n : ℕ) → + Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + = (2 * n - 1)‼ + | 0 => rfl + | Nat.succ n => by + rw [involutionNoFixed_card_succ] + rw [involutionNoFixed_card n] + exact Eq.symm (Nat.doubleFactorial_add_one (Nat.mul 2 n)) + +lemma involutionNoFixed_even : (n : ℕ) → (he : Even n) → + Fintype.card {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} = (n - 1)‼ := by + intro n he + obtain ⟨r, hr⟩ := he + have hr' : n = 2 * r := by omega + subst hr' + exact involutionNoFixed_card r + +def isFullInvolutionEquiv {φs : List 𝓕} : + {c : Contractions φs // IsFull c} ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f ∧ (∀ i, f i ≠ i)} where + toFun c := ⟨equivInvolutions c.1, by + apply And.intro (equivInvolutions c.1).2 + rw [← isFull_iff_involution_no_fixed_points] + exact c.2 + ⟩ + invFun f := ⟨equivInvolutions.symm ⟨f.1, f.2.1⟩, by + rw [isFull_iff_involution_no_fixed_points] + simpa using f.2.2⟩ + left_inv c := by simp + right_inv f := by simp + +lemma card_of_full_contractions {φs : List 𝓕} (he : Even φs.length ) : + Fintype.card {c : Contractions φs // IsFull c} = (φs.length - 1)‼ := by + rw [Fintype.card_congr (isFullInvolutionEquiv (φs := φs))] + exact involutionNoFixed_even φs.length he + + /-- A structure specifying when a type `I` and a map `f :I → Type` corresponds to the splitting of a fields `φ` into a creation `n` and annihlation part `p`. -/ structure Splitting (f : 𝓕 → Type) [∀ i, Fintype (f i)] From 1ab0c6f76904b49d21f9c1e09d26746f52caa4eb Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Sun, 5 Jan 2025 13:07:32 +0000 Subject: [PATCH 4/8] feat: Cardinality of involutions and refactor --- HepLean.lean | 2 + HepLean/Mathematics/Fin/Involutions.lean | 572 ++++++++ .../Contractions/Basic.lean | 1234 +---------------- .../PerturbationTheory/Contractions/Card.lean | 40 + .../Contractions/Involutions.lean | 367 ++++- .../Wick/StaticTheorem.lean | 10 +- 6 files changed, 1028 insertions(+), 1197 deletions(-) create mode 100644 HepLean/Mathematics/Fin/Involutions.lean create mode 100644 HepLean/PerturbationTheory/Contractions/Card.lean diff --git a/HepLean.lean b/HepLean.lean index 84522e2c..3cdfebc6 100644 --- a/HepLean.lean +++ b/HepLean.lean @@ -100,6 +100,7 @@ import HepLean.Lorentz.Weyl.Modules import HepLean.Lorentz.Weyl.Two import HepLean.Lorentz.Weyl.Unit import HepLean.Mathematics.Fin +import HepLean.Mathematics.Fin.Involutions import HepLean.Mathematics.LinearMaps import HepLean.Mathematics.List import HepLean.Mathematics.PiTensorProduct @@ -114,6 +115,7 @@ import HepLean.Meta.Notes.NoteFile import HepLean.Meta.Notes.ToHTML import HepLean.Meta.TransverseTactics import HepLean.PerturbationTheory.Contractions.Basic +import HepLean.PerturbationTheory.Contractions.Card import HepLean.PerturbationTheory.Contractions.Involutions import HepLean.PerturbationTheory.FeynmanDiagrams.Basic import HepLean.PerturbationTheory.FeynmanDiagrams.Instances.ComplexScalar diff --git a/HepLean/Mathematics/Fin/Involutions.lean b/HepLean/Mathematics/Fin/Involutions.lean new file mode 100644 index 00000000..606f4610 --- /dev/null +++ b/HepLean/Mathematics/Fin/Involutions.lean @@ -0,0 +1,572 @@ +/- +Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Tooby-Smith +-/ +import Mathlib.LinearAlgebra.PiTensorProduct +import Mathlib.Tactic.Polyrith +import Mathlib.Tactic.Linarith +import Mathlib.Data.Nat.Factorial.DoubleFactorial +/-! +# Fin involutions + +Some properties of involutions of `Fin n`. + +These involutions are used in e.g. proving results about Wick contractions. + +-/ +namespace HepLean.Fin + +open Nat + +def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involutive f } ≃ + (f : {f : Fin n → Fin n // Function.Involutive f}) × {i : Option (Fin n) // + ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} where + toFun f := ⟨⟨ + fun i => + if h : f.1 i.succ = 0 then i + else Fin.pred (f.1 i.succ) h , by + intro i + by_cases h : f.1 i.succ = 0 + · simp [h] + · simp [h] + simp [f.2 i.succ] + intro h + exact False.elim (Fin.succ_ne_zero i h)⟩, + ⟨if h : f.1 0 = 0 then none else Fin.pred (f.1 0) h , by + by_cases h0 : f.1 0 = 0 + · simp [h0] + · simp [h0] + refine fun h => False.elim (h (f.2 0))⟩⟩ + invFun f := ⟨ + if h : (f.2.1).isSome then + Fin.cons (f.2.1.get h).succ (Function.update (Fin.succ ∘ f.1.1) (f.2.1.get h) 0) + else + Fin.cons 0 (Fin.succ ∘ f.1.1) + , by + by_cases hs : (f.2.1).isSome + · simp only [Nat.succ_eq_add_one, hs, ↓reduceDIte, Fin.coe_eq_castSucc] + let a := f.2.1.get hs + change Function.Involutive (Fin.cons a.succ (Function.update (Fin.succ ∘ ↑f.fst) a 0)) + intro i + rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ + · subst hi + rw [Fin.cons_zero, Fin.cons_succ] + simp + · subst hj + rw [Fin.cons_succ] + by_cases hja : j = a + · subst hja + simp + · rw [Function.update_apply ] + rw [if_neg hja] + simp + have hf2 := f.2.2 hs + change f.1.1 a = a at hf2 + have hjf1 : f.1.1 j ≠ a := by + by_contra hn + have haj : j = f.1.1 a := by + rw [← hn] + rw [f.1.2] + rw [hf2] at haj + exact hja haj + rw [Function.update_apply, if_neg hjf1] + simp + rw [f.1.2] + · simp [hs] + intro i + rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ + · subst hi + simp + · subst hj + simp + rw [f.1.2]⟩ + left_inv f := by + match f with + | ⟨f, hf⟩ => + simp + ext i + by_cases h0 : f 0 = 0 + · simp [h0] + rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ + · subst hi + simp [h0] + · subst hj + simp [h0] + by_cases hj : f j.succ =0 + · rw [← h0] at hj + have hn := Function.Involutive.injective hf hj + exact False.elim (Fin.succ_ne_zero j hn) + · simp [hj] + rw [Fin.ext_iff] at hj + simp at hj + omega + · rw [if_neg h0] + by_cases hf' : i = f 0 + · subst hf' + simp + rw [hf] + simp + · rw [Function.update_apply, if_neg hf'] + rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ + · subst hi + simp + · subst hj + simp + by_cases hj : f j.succ =0 + · rw [← hj] at hf' + rw [hf] at hf' + simp at hf' + · simp [hj] + rw [Fin.ext_iff] at hj + simp at hj + omega + right_inv f := by + match f with + | ⟨⟨f, hf⟩, ⟨f0, hf0⟩⟩ => + ext i + · simp + by_cases hs : f0.isSome + · simp [hs] + by_cases hi : i = f0.get hs + · simp [hi, Function.update_apply] + exact Eq.symm (Fin.val_eq_of_eq (hf0 hs)) + · simp [hi] + split + · rename_i h + exact False.elim (Fin.succ_ne_zero (f i) h) + · rfl + · simp [hs] + split + · rename_i h + exact False.elim (Fin.succ_ne_zero (f i) h) + · rfl + · simp only [Nat.succ_eq_add_one, Option.mem_def, + Option.dite_none_left_eq_some, Option.some.injEq] + by_cases hs : f0.isSome + · simp only [hs, ↓reduceDIte] + simp [Fin.cons_zero] + have hx : ¬ (f0.get hs).succ = 0 := (Fin.succ_ne_zero (f0.get hs)) + simp [hx] + refine Iff.intro (fun hi => ?_) (fun hi => ?_) + · rw [← hi] + exact + Option.eq_some_of_isSome + (Eq.mpr_prop (Eq.refl (f0.isSome = true)) + (of_eq_true (Eq.trans (congrArg (fun x => x = true) hs) (eq_self true)))) + · subst hi + exact rfl + · simp [hs] + simp at hs + subst hs + exact ne_of_beq_false rfl + +lemma involutionCons_ext {n : ℕ} {f1 f2 : (f : {f : Fin n → Fin n // Function.Involutive f}) × + {i : Option (Fin n) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)}} + (h1 : f1.1 = f2.1) (h2 : f1.2 = Equiv.subtypeEquivRight (by rw [h1]; simp) f2.2) : f1 = f2 := by + cases f1 + cases f2 + simp at h1 h2 + subst h1 + rename_i fst snd snd_1 + simp_all only [Sigma.mk.inj_iff, heq_eq_eq, true_and] + obtain ⟨val, property⟩ := fst + obtain ⟨val_1, property_1⟩ := snd + obtain ⟨val_2, property_2⟩ := snd_1 + simp_all only + rfl + + +def involutionAddEquiv {n : ℕ} (f : {f : Fin n → Fin n // Function.Involutive f}) : + {i : Option (Fin n) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} ≃ + Option (Fin (Finset.univ.filter fun i => f.1 i = i).card) := by + let e1 : {i : Option (Fin n) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} + ≃ Option {i : Fin n // f.1 i = i} := + { toFun := fun i => match i with + | ⟨some i, h⟩ => some ⟨i, by simpa using h⟩ + | ⟨none, h⟩ => none + invFun := fun i => match i with + | some ⟨i, h⟩ => ⟨some i, by simpa using h⟩ + | none => ⟨none, by simp⟩ + left_inv := by + intro a + cases a + aesop + right_inv := by + intro a + cases a + rfl + simp_all only [Subtype.coe_eta] } + let s : Finset (Fin n) := Finset.univ.filter fun i => f.1 i = i + let e2' : { i : Fin n // f.1 i = i} ≃ {i // i ∈ s} := by + refine Equiv.subtypeEquivProp ?h + funext i + simp [s] + let e2 : {i // i ∈ s} ≃ Fin (Finset.card s) := by + refine (Finset.orderIsoOfFin _ ?_).symm.toEquiv + simp [s] + refine e1.trans (Equiv.optionCongr (e2'.trans (e2))) + + +lemma involutionAddEquiv_none_image_zero {n : ℕ} : + {f : {f : Fin n.succ → Fin n.succ // Function.Involutive f}} + → involutionAddEquiv (involutionCons n f).1 (involutionCons n f).2 = none + → f.1 ⟨0, Nat.zero_lt_succ n⟩ = ⟨0, Nat.zero_lt_succ n⟩ := by + intro f h + simp only [Nat.succ_eq_add_one, involutionCons, Equiv.coe_fn_mk, involutionAddEquiv, + Option.isSome_some, Option.get_some, Option.isSome_none, Equiv.trans_apply, + Equiv.optionCongr_apply, Equiv.coe_trans, RelIso.coe_fn_toEquiv, Option.map_eq_none'] at h + simp_all only [List.length_cons, Fin.zero_eta] + obtain ⟨val, property⟩ := f + simp_all only [List.length_cons] + split at h + next i i_1 h_1 heq => + split at heq + next h_2 => simp_all only [reduceCtorEq] + next h_2 => simp_all only [reduceCtorEq] + next i h_1 heq => + split at heq + next h_2 => simp_all only + next h_2 => simp_all only [Subtype.mk.injEq, reduceCtorEq] + +lemma involutionAddEquiv_cast {n : ℕ} {f1 f2 : {f : Fin n → Fin n // Function.Involutive f}} + (hf : f1 = f2): + involutionAddEquiv f1 = (Equiv.subtypeEquivRight (by rw [hf]; simp)).trans + ((involutionAddEquiv f2).trans (Equiv.optionCongr (finCongr (by rw [hf])))):= by + subst hf + simp + rfl + + +lemma involutionAddEquiv_cast' {m : ℕ} {f1 f2 : {f : Fin m → Fin m // Function.Involutive f}} + {N : ℕ} (hf : f1 = f2) (n : Option (Fin N)) (hn1 : N = (Finset.filter (fun i => f1.1 i = i) Finset.univ).card) + (hn2 : N = (Finset.filter (fun i => f2.1 i = i) Finset.univ).card): + HEq ((involutionAddEquiv f1).symm (Option.map (finCongr hn1) n)) ((involutionAddEquiv f2).symm (Option.map (finCongr hn2) n)) := by + subst hf + rfl + +lemma involutionAddEquiv_none_succ {n : ℕ} + {f : {f : Fin n.succ → Fin n.succ // Function.Involutive f}} + (h : involutionAddEquiv (involutionCons n f).1 (involutionCons n f).2 = none) + (x : Fin n) : f.1 x.succ = x.succ ↔ (involutionCons n f).1.1 x = x := by + simp [involutionCons] + have hn' := involutionAddEquiv_none_image_zero h + have hx : ¬ f.1 x.succ = ⟨0, Nat.zero_lt_succ n⟩:= by + rw [← hn'] + exact fun hn => Fin.succ_ne_zero x (Function.Involutive.injective f.2 hn) + apply Iff.intro + · intro h2 h3 + rw [Fin.ext_iff] + simp [h2] + · intro h2 + have h2' := h2 hx + conv_rhs => rw [← h2'] + simp + +lemma involutionAddEquiv_isSome_image_zero {n : ℕ} : + {f : {f : Fin n.succ → Fin n.succ // Function.Involutive f}} + → (involutionAddEquiv (involutionCons n f).1 (involutionCons n f).2).isSome + → ¬ f.1 ⟨0, Nat.zero_lt_succ n⟩ = ⟨0, Nat.zero_lt_succ n⟩ := by + intro f hf + simp [involutionAddEquiv, involutionCons] at hf + simp_all only [List.length_cons, Fin.zero_eta] + obtain ⟨val, property⟩ := f + simp_all only [List.length_cons] + apply Aesop.BuiltinRules.not_intro + intro a + simp_all only [↓reduceDIte, Option.isSome_none, Bool.false_eq_true] + +def involutionNoFixedEquivSum {n : ℕ} : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f + ∧ ∀ i, f i ≠ i} ≃ Σ (k : Fin (2 * n + 1)), + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f + ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} where + toFun f := ⟨(f.1 0).pred (f.2.2 0), ⟨f.1, f.2.1, by simpa using f.2.2⟩⟩ + invFun f := ⟨f.2.1, ⟨f.2.2.1, f.2.2.2.1⟩⟩ + left_inv f := by + rfl + right_inv f := by + simp + ext + · simp + rw [f.2.2.2.2] + simp + · simp + +/-! + +## Equivalences of involutions with no fixed points. + +The main aim of thes equivalences is to define `involutionNoFixedZeroEquivProd`. + +-/ + +def involutionNoFixedZeroSetEquivEquiv {n : ℕ} + (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f + ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} ≃ + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive (e.symm ∘ f ∘ e) ∧ + (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} where + toFun f := ⟨e ∘ f.1 ∘ e.symm, by + intro i + simp + rw [f.2.1], by + simpa using f.2.2.1, by simpa using f.2.2.2⟩ + invFun f := ⟨e.symm ∘ f.1 ∘ e, by + intro i + simp + have hf2 := f.2.1 i + simpa using hf2, by + simpa using f.2.2.1, by simpa using f.2.2.2⟩ + left_inv f := by + ext i + simp + right_inv f := by + ext i + simp + +def involutionNoFixedZeroSetEquivSetEquiv {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive (e.symm ∘ f ∘ e) ∧ + (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} ≃ + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} := by + refine Equiv.subtypeEquivRight ?_ + intro f + have h1 : Function.Involutive (⇑e.symm ∘ f ∘ ⇑e) ↔ Function.Involutive f := by + apply Iff.intro + · intro h i + have hi := h (e.symm i) + simpa using hi + · intro h i + have hi := h (e i) + simp [hi] + rw [h1] + simp + intro h1 h2 + apply Iff.intro + · intro h i + have hi := h (e.symm i) + simpa using hi + · intro h i + have hi := h (e i) + by_contra hn + nth_rewrite 2 [← hn] at hi + simp at hi + + +def involutionNoFixedZeroSetEquivEquiv' {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} + ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f (e 0) = e k.succ} := by + refine Equiv.subtypeEquivRight ?_ + simp + intro f hi h1 + exact Equiv.symm_apply_eq e + +def involutionNoFixedZeroSetEquivSetOne {n : ℕ} (k : Fin (2 * n + 1)) : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f 0 = k.succ} + ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f 0 = 1} := by + refine Equiv.trans (involutionNoFixedZeroSetEquivEquiv k (Equiv.swap k.succ 1)) ?_ + refine Equiv.trans (involutionNoFixedZeroSetEquivSetEquiv k (Equiv.swap k.succ 1)) ?_ + refine Equiv.trans (involutionNoFixedZeroSetEquivEquiv' k (Equiv.swap k.succ 1)) ?_ + refine Equiv.subtypeEquivRight ?_ + simp + intro f hi h1 + rw [Equiv.swap_apply_of_ne_of_ne] + · exact Ne.symm (Fin.succ_ne_zero k) + · exact Fin.zero_ne_one + +def involutionNoFixedSetOne {n : ℕ} : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f 0 = 1} ≃ + {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ + (∀ i, f i ≠ i)} where + toFun f := by + have hf1 : f.1 1 = 0 := by + have hf := f.2.2.2 + simp [← hf] + rw [f.2.1] + let f' := f.1 ∘ Fin.succ ∘ Fin.succ + have hf' (i : Fin (2 * n)) : f' i ≠ 0 := by + simp [f'] + simp [← hf1] + by_contra hn + have hn' := Function.Involutive.injective f.2.1 hn + simp [Fin.ext_iff] at hn' + let f'' := fun i => (f' i).pred (hf' i) + have hf'' (i : Fin (2 * n)) : f'' i ≠ 0 := by + simp [f''] + rw [@Fin.pred_eq_iff_eq_succ] + simp [f'] + simp [← f.2.2.2 ] + by_contra hn + have hn' := Function.Involutive.injective f.2.1 hn + simp [Fin.ext_iff] at hn' + let f''' := fun i => (f'' i).pred (hf'' i) + refine ⟨f''', ?_, ?_⟩ + · intro i + simp [f''', f'', f'] + simp [f.2.1 i.succ.succ] + · intro i + simp [f''', f'', f'] + rw [@Fin.pred_eq_iff_eq_succ] + rw [@Fin.pred_eq_iff_eq_succ] + exact f.2.2.1 i.succ.succ + invFun f := by + let f' := fun (i : Fin (2 * n.succ))=> + match i with + | ⟨0, h⟩ => 1 + | ⟨1, h⟩ => 0 + | ⟨(Nat.succ (Nat.succ n)), h⟩ => (f.1 ⟨n, by omega⟩).succ.succ + refine ⟨f', ?_, ?_, ?_⟩ + · intro i + match i with + | ⟨0, h⟩ => rfl + | ⟨1, h⟩ => rfl + | ⟨(Nat.succ (Nat.succ m)), h⟩ => + simp [f'] + split + · rename_i h + simp at h + exact False.elim (Fin.succ_ne_zero (f.1 ⟨m, _⟩).succ h) + · rename_i h + simp [Fin.ext_iff] at h + · rename_i h + rename_i x r + simp_all [Fin.ext_iff] + have hfn {a b : ℕ} {ha : a < 2 * n} {hb : b < 2 * n} + (hab : ↑(f.1 ⟨a, ha⟩) = b): ↑(f.1 ⟨b, hb⟩) = a := by + have ht : f.1 ⟨a, ha⟩ = ⟨b, hb⟩ := by + simp [hab, Fin.ext_iff] + rw [← ht, f.2.1] + exact hfn h + · intro i + match i with + | ⟨0, h⟩ => + simp [f'] + split + · rename_i h + simp + · rename_i h + simp [Fin.ext_iff] at h + · rename_i h + simp [Fin.ext_iff] at h + | ⟨1, h⟩ => + simp [f'] + split + · rename_i h + simp at h + · rename_i h + simp + · rename_i h + simp [Fin.ext_iff] at h + | ⟨(Nat.succ (Nat.succ m)), h⟩ => + simp [f', Fin.ext_iff] + have hf:= f.2.2 ⟨m, by exact Nat.add_lt_add_iff_right.mp h⟩ + simp [Fin.ext_iff] at hf + omega + · simp [f'] + split + · rename_i h + simp + · rename_i h + simp at h + · rename_i h + simp [Fin.ext_iff] at h + left_inv f := by + have hf1 : f.1 1 = 0 := by + have hf := f.2.2.2 + simp [← hf] + rw [f.2.1] + simp + ext i + simp + split + · simp + rw [f.2.2.2] + simp + · simp + rw [hf1] + simp + · rfl + right_inv f := by + simp + ext i + simp + split + · rename_i h + simp [Fin.ext_iff] at h + · rename_i h + simp [Fin.ext_iff] at h + · rename_i h + simp + congr + apply congrArg + simp_all [Fin.ext_iff] + + +def involutionNoFixedZeroSetEquiv {n : ℕ} (k : Fin (2 * n + 1)) : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f 0 = k.succ} + ≃ {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + refine Equiv.trans (involutionNoFixedZeroSetEquivSetOne k) involutionNoFixedSetOne + +def involutionNoFixedEquivSumSame {n : ℕ} : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + ≃ Σ (_ : Fin (2 * n + 1)), {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + refine Equiv.trans involutionNoFixedEquivSum ?_ + refine Equiv.sigmaCongrRight involutionNoFixedZeroSetEquiv + +def involutionNoFixedZeroEquivProd {n : ℕ} : + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + ≃ Fin (2 * n + 1) × {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + refine Equiv.trans involutionNoFixedEquivSumSame ?_ + exact Equiv.sigmaEquivProd (Fin (2 * n + 1)) + { f // Function.Involutive f ∧ ∀ (i : Fin (2 * n)), f i ≠ i} + +/-! + +## Cardinality + +-/ + + +instance {n : ℕ} : Fintype { f // Function.Involutive f ∧ ∀ (i : Fin n), f i ≠ i } := by + haveI : DecidablePred fun x => Function.Involutive x := by + intro f + apply Fintype.decidableForallFintype (α := Fin n) + haveI : DecidablePred fun x => Function.Involutive x ∧ ∀ (i : Fin n), x i ≠ i := by + intro x + apply instDecidableAnd + apply Subtype.fintype + +lemma involutionNoFixed_card_succ {n : ℕ} : + Fintype.card {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + = (2 * n + 1) * Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + rw [Fintype.card_congr (involutionNoFixedZeroEquivProd)] + rw [Fintype.card_prod ] + congr + exact Fintype.card_fin (2 * n + 1) + + +lemma involutionNoFixed_card_mul_two : (n : ℕ) → + Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + = (2 * n - 1)‼ + | 0 => rfl + | Nat.succ n => by + rw [involutionNoFixed_card_succ] + rw [involutionNoFixed_card_mul_two n] + exact Eq.symm (Nat.doubleFactorial_add_one (Nat.mul 2 n)) + +lemma involutionNoFixed_card_even : (n : ℕ) → (he : Even n) → + Fintype.card {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} = (n - 1)‼ := by + intro n he + obtain ⟨r, hr⟩ := he + have hr' : n = 2 * r := by omega + subst hr' + exact involutionNoFixed_card_mul_two r + +end HepLean.Fin diff --git a/HepLean/PerturbationTheory/Contractions/Basic.lean b/HepLean/PerturbationTheory/Contractions/Basic.lean index 2e301072..c57e1d1a 100644 --- a/HepLean/PerturbationTheory/Contractions/Basic.lean +++ b/HepLean/PerturbationTheory/Contractions/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith -/ import HepLean.PerturbationTheory.Wick.OperatorMap -import Mathlib.Data.Nat.Factorial.DoubleFactorial +import HepLean.Mathematics.Fin.Involutions /-! # Contractions of a list of fields @@ -14,6 +14,7 @@ import Mathlib.Data.Nat.Factorial.DoubleFactorial namespace Wick open HepLean.List +open HepLean.Fin open FieldStatistic variable {𝓕 : Type} @@ -32,8 +33,42 @@ namespace Contractions variable {l : List 𝓕} (c : Contractions l) +def auxCongr : {φs: List 𝓕} → {φsᵤₙ φsᵤₙ' : List 𝓕} → (h : φsᵤₙ = φsᵤₙ') → + ContractionsAux φs φsᵤₙ ≃ ContractionsAux φs φsᵤₙ' + | _, _, _, Eq.refl _ => Equiv.refl _ + +lemma auxCongr_ext {φs: List 𝓕} {c c2 : Contractions φs} (h : c.1 = c2.1) + (hx : c.2 = auxCongr h.symm c2.2) : c = c2 := by + cases c + cases c2 + simp at h + subst h + simp [auxCongr] at hx + subst hx + rfl + /-- The list of uncontracted fields. -/ -def normalize : List 𝓕 := c.1 +def uncontracted : List 𝓕 := c.1 + +lemma uncontracted_length_even_iff : {l : List 𝓕} → (c : Contractions l) → + Even l.length ↔ Even c.uncontracted.length + | [], ⟨[], ContractionsAux.nil⟩ => by + simp [uncontracted] + | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => by + simp only [List.length_cons, uncontracted, optionEraseZ] + rw [Nat.even_add_one, Nat.even_add_one] + rw [uncontracted_length_even_iff ⟨aux, c⟩] + rfl + | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩=> by + simp only [List.length_cons, uncontracted, optionEraseZ_some_length] + rw [Nat.even_sub, Nat.even_add_one] + · simp only [Nat.not_even_iff_odd, Nat.not_even_one, iff_false] + rw [← Nat.not_even_iff_odd, ← Nat.not_even_iff_odd] + rw [uncontracted_length_even_iff ⟨aux, c⟩] + rfl + · refine Nat.one_le_iff_ne_zero.mpr (fun hn => ?_) + rw [hn] at n + exact Fin.elim0 n lemma contractions_nil (a : Contractions ([] : List 𝓕)) : a = ⟨[], ContractionsAux.nil⟩ := by cases a @@ -41,7 +76,7 @@ lemma contractions_nil (a : Contractions ([] : List 𝓕)) : a = ⟨[], Contract cases c rfl -def embedUncontracted {l : List 𝓕} (c : Contractions l) : Fin c.normalize.length → Fin l.length := +def embedUncontracted {l : List 𝓕} (c : Contractions l) : Fin c.uncontracted.length → Fin l.length := match l, c with | [], ⟨[], ContractionsAux.nil⟩ => Fin.elim0 | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => @@ -49,9 +84,9 @@ def embedUncontracted {l : List 𝓕} (c : Contractions l) : Fin c.normalize.len | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => by let lc := embedUncontracted ⟨aux, c⟩ refine Fin.succ ∘ lc ∘ Fin.cast ?_ ∘ Fin.succAbove ⟨n, by - rw [normalize, optionEraseZ_some_length] + rw [uncontracted, optionEraseZ_some_length] omega⟩ - simp only [normalize, optionEraseZ_some_length] + simp only [uncontracted, optionEraseZ_some_length] have hq : aux.length ≠ 0 := by by_contra hn rw [hn] at n @@ -80,8 +115,6 @@ lemma embedUncontracted_injective {l : List 𝓕} (c : Contractions l) : refine Function.Injective.comp (Fin.cast_injective (embedUncontracted.proof_5 φ φs aux i c)) Fin.succAbove_right_injective - - /-- Establishes uniqueness of contractions for a single field, showing that any contraction of a single field must be equivalent to the trivial contraction with no pairing. -/ lemma contractions_single {i : 𝓕} (a : Contractions [i]) : a = @@ -107,17 +140,17 @@ def nilEquiv : Contractions ([] : List 𝓕) ≃ Unit where right_inv _ := rfl /-- The equivalence between contractions of `a :: l` and contractions of - `Contractions l` paired with an optional element of `Fin (c.normalize).length` specifying + `Contractions l` paired with an optional element of `Fin (c.uncontracted).length` specifying what `a` contracts with if any. -/ def consEquiv {φ : 𝓕} {φs : List 𝓕} : - Contractions (φ :: φs) ≃ (c : Contractions φs) × Option (Fin c.normalize.length) where + Contractions (φ :: φs) ≃ (c : Contractions φs) × Option (Fin c.uncontracted.length) where toFun c := match c with | ⟨aux, c⟩ => match c with | ContractionsAux.cons (φsᵤₙ := aux') i c => ⟨⟨aux', c⟩, i⟩ invFun ci := - ⟨(optionEraseZ (ci.fst.normalize) φ ci.2), ContractionsAux.cons (φ := φ) ci.2 ci.1.2⟩ + ⟨(optionEraseZ (ci.fst.uncontracted) φ ci.2), ContractionsAux.cons (φ := φ) ci.2 ci.1.2⟩ left_inv c := by match c with | ⟨aux, c⟩ => @@ -126,9 +159,9 @@ def consEquiv {φ : 𝓕} {φs : List 𝓕} : right_inv ci := by rfl lemma consEquiv_ext {φs : List 𝓕} {c1 c2 : Contractions φs} - {n1 : Option (Fin c1.normalize.length)} {n2 : Option (Fin c2.normalize.length)} + {n1 : Option (Fin c1.uncontracted.length)} {n2 : Option (Fin c2.uncontracted.length)} (hc : c1 = c2) (hn : Option.map (finCongr (by rw [hc])) n1 = n2) : - (⟨c1, n1⟩ : (c : Contractions φs) × Option (Fin c.normalize.length)) = ⟨c2, n2⟩ := by + (⟨c1, n1⟩ : (c : Contractions φs) × Option (Fin c.uncontracted.length)) = ⟨c2, n2⟩ := by subst hc subst hn simp @@ -142,7 +175,7 @@ instance decidable : (φs : List 𝓕) → DecidableEq (Contractions φs) | ContractionsAux.nil, ContractionsAux.nil => isTrue rfl | _ :: φs => haveI : DecidableEq (Contractions φs) := decidable φs - haveI : DecidableEq ((c : Contractions φs) × Option (Fin (c.normalize).length)) := + haveI : DecidableEq ((c : Contractions φs) × Option (Fin (c.uncontracted).length)) := Sigma.instDecidableEqSigma Equiv.decidableEq consEquiv @@ -156,1189 +189,20 @@ instance fintype : (φs : List 𝓕) → Fintype (Contractions φs) exact contractions_nil a} | φ :: φs => haveI : Fintype (Contractions φs) := fintype φs - haveI : Fintype ((c : Contractions φs) × Option (Fin (c.normalize).length)) := + haveI : Fintype ((c : Contractions φs) × Option (Fin (c.uncontracted).length)) := Sigma.instFintype Fintype.ofEquiv _ consEquiv.symm /-- A contraction is a full contraction if there normalizing list of fields is empty. -/ -def IsFull : Prop := c.normalize = [] +def IsFull : Prop := c.uncontracted = [] /-- Provides a decidable instance for determining if a contraction is full (i.e., all fields are paired). -/ instance isFull_decidable : Decidable c.IsFull := by - have hn : c.IsFull ↔ c.normalize.length = 0 := by + have hn : c.IsFull ↔ c.uncontracted.length = 0 := by simp [IsFull] apply decidable_of_decidable_of_iff hn.symm -def toOptionList : {l : List 𝓕} → (c : Contractions l) → List (Option (Fin l.length)) - | [], ⟨[], ContractionsAux.nil⟩ => [] - | _ :: _, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => none :: - List.map (Option.map Fin.succ) (toOptionList ⟨aux, c⟩) - | _ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => - some (Fin.succ (embedUncontracted ⟨aux, c⟩ n)) :: - List.set ((List.map (Option.map Fin.succ) (toOptionList ⟨aux, c⟩))) - (embedUncontracted ⟨aux, c⟩ n) (some ⟨0, Nat.zero_lt_succ φs.length⟩) - -lemma toOptionList_length {l : List 𝓕} (c : Contractions l) : c.toOptionList.length = l.length := by - match l, c with - | [], ⟨[], ContractionsAux.nil⟩ => - dsimp [toOptionList] - | _ :: _, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => - dsimp [toOptionList] - rw [List.length_map, toOptionList_length ⟨aux, c⟩] - | _ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => - dsimp [toOptionList] - rw [List.length_set, List.length_map, toOptionList_length ⟨aux, c⟩] - -def toFinOptionMap {l : List 𝓕} (c : Contractions l) : Fin l.length → Option (Fin l.length) := - c.toOptionList.get ∘ Fin.cast (toOptionList_length c).symm - -lemma toFinOptionMap_neq_self {l : List 𝓕} (c : Contractions l) (i : Fin l.length) : - c.toFinOptionMap i ≠ some i := by - match l, c with - | [], ⟨[], ContractionsAux.nil⟩ => - exact Fin.elim0 i - | _ :: _, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => - dsimp [toFinOptionMap, toOptionList] - match i with - | ⟨0, _⟩ => - exact Option.noConfusion - | ⟨n + 1, h⟩ => - simp only [List.getElem_cons_succ, List.getElem_map, List.length_cons] - have hn := toFinOptionMap_neq_self ⟨aux, c⟩ ⟨n, Nat.succ_lt_succ_iff.mp h⟩ - simp only [Option.map_eq_some', not_exists, not_and] - intro x hx - simp only [toFinOptionMap, Function.comp_apply, Fin.cast_mk, List.get_eq_getElem, hx, ne_eq, - Option.some.injEq] at hn - rw [Fin.ext_iff] at hn ⊢ - simp_all - | _ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => - dsimp [toFinOptionMap, toOptionList] - match i with - | ⟨0, _⟩ => - simp only [List.getElem_cons_zero, List.length_cons, Fin.zero_eta, Option.some.injEq, ne_eq] - exact Fin.succ_ne_zero (embedUncontracted ⟨aux, c⟩ n) - | ⟨n + 1, h⟩ => - simp only [List.getElem_cons_succ, List.length_cons, ne_eq] - rw [List.getElem_set] - split - · exact ne_of_beq_false rfl - · simp only [List.getElem_map, Option.map_eq_some', not_exists, not_and] - intro x hx - have hn := toFinOptionMap_neq_self ⟨aux, c⟩ ⟨n, Nat.succ_lt_succ_iff.mp h⟩ - simp only [toFinOptionMap, Function.comp_apply, Fin.cast_mk, List.get_eq_getElem, hx, ne_eq, - Option.some.injEq] at hn - rw [Fin.ext_iff] at hn ⊢ - simp_all - -@[simp] -lemma toFinOptionMap_embedUncontracted {l : List 𝓕} (c : Contractions l) - (i : Fin c.normalize.length) : c.toFinOptionMap (embedUncontracted c i) = none := by - match l, c with - | [], ⟨[], ContractionsAux.nil⟩ => - exact Fin.elim0 i - | _ :: _, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => - dsimp [toFinOptionMap, toOptionList, embedUncontracted] - match i with - | ⟨0, _⟩ => - rfl - | ⟨n + 1, h⟩ => - rw [show ⟨n + 1, h⟩ = Fin.succ ⟨n, Nat.succ_lt_succ_iff.mp h⟩ by rfl] - rw [Fin.cons_succ] - simp only [Function.comp_apply, Fin.val_succ, List.getElem_cons_succ, List.getElem_map, - Option.map_eq_none'] - exact toFinOptionMap_embedUncontracted ⟨aux, c⟩ ⟨n, Nat.succ_lt_succ_iff.mp h⟩ - | _ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => - dsimp [toFinOptionMap, toOptionList, embedUncontracted] - rw [List.getElem_set] - split - · rename_i hs - have hx := embedUncontracted_injective ⟨aux, c⟩ (Fin.val_injective hs) - refine False.elim ?_ - have hn := Fin.succAbove_ne ⟨n, embedUncontracted.proof_6 _ φs aux n c⟩ i - simp [Fin.ext_iff] at hx - simp [Fin.ext_iff] at hn - exact hn (id (Eq.symm hx)) - · simp - exact toFinOptionMap_embedUncontracted ⟨aux, c⟩ _ - -lemma toFinOptionMap_involutive {l : List 𝓕} (c : Contractions l) (i j : Fin l.length) : - c.toFinOptionMap i = some j → c.toFinOptionMap j = some i := by - match l, c with - | [], ⟨[], ContractionsAux.nil⟩ => - exact Fin.elim0 i - | _ :: _, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => - dsimp [toFinOptionMap, toOptionList] - match i, j with - | ⟨0, _⟩, j => - exact Option.noConfusion - | ⟨n + 1, h⟩, ⟨0, _⟩ => - simp - intro x hx - exact Fin.succ_ne_zero x - | ⟨n + 1, hn⟩, ⟨m + 1, hm⟩ => - intro h1 - have hnm := toFinOptionMap_involutive ⟨aux, c⟩ ⟨n, Nat.succ_lt_succ_iff.mp hn⟩ - ⟨m, Nat.succ_lt_succ_iff.mp hm⟩ - simp - simp [toFinOptionMap] at hnm - have hnm' := hnm (by - simp at h1 - obtain ⟨a, ha, ha2⟩ := h1 - rw [ha] - simp only [Option.some.injEq] - rw [Fin.ext_iff] at ha2 ⊢ - simp_all) - use ⟨n, Nat.succ_lt_succ_iff.mp hn⟩ - simpa using hnm' - | _ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => - dsimp [toFinOptionMap, toOptionList] - match i, j with - | ⟨0, _⟩, j => - intro hj - simp only [List.getElem_cons_zero, Option.some.injEq] at hj - subst hj - simp - | ⟨n' + 1, h⟩, ⟨0, _⟩ => - intro hj - simp at hj - simp - rw [List.getElem_set] at hj - simp_all only [List.length_cons, lt_add_iff_pos_left, add_pos_iff, zero_lt_one, or_true, List.getElem_map, - ite_eq_left_iff, Option.map_eq_some'] - simp [Fin.ext_iff] - by_contra hn - have hn' := hj hn - obtain ⟨a, ha, ha2⟩ := hn' - exact Fin.succ_ne_zero a ha2 - | ⟨n' + 1, hn⟩, ⟨m + 1, hm⟩ => - simp - rw [List.getElem_set, List.getElem_set] - simp - split - · intro h - simp [Fin.ext_iff] at h - rename_i hn' - intro h1 - split - · rename_i hnx - have hnm := toFinOptionMap_involutive ⟨aux, c⟩ ⟨n', Nat.succ_lt_succ_iff.mp hn⟩ - ⟨m, Nat.succ_lt_succ_iff.mp hm⟩ - subst hnx - simp at hnm - refine False.elim (hnm ?_) - simp at h1 - obtain ⟨a, ha, ha2⟩ := h1 - have ha' : (embedUncontracted ⟨aux, c⟩ n) = a := by - rw [Fin.ext_iff] at ha2 ⊢ - simp_all - rw [ha'] - rw [← ha] - rfl - · rename_i hnx - have hnm := toFinOptionMap_involutive ⟨aux, c⟩ ⟨n', Nat.succ_lt_succ_iff.mp hn⟩ - ⟨m, Nat.succ_lt_succ_iff.mp hm⟩ (by - dsimp [toFinOptionMap] - simp at h1 - obtain ⟨a, ha, ha2⟩ := h1 - have ha' : a = ⟨m, by exact Nat.succ_lt_succ_iff.mp hm⟩ := by - rw [Fin.ext_iff] at ha2 ⊢ - simp_all - rw [← ha', ← ha]) - change Option.map Fin.succ (toFinOptionMap ⟨aux, c⟩ ⟨m, Nat.succ_lt_succ_iff.mp hm⟩) = _ - rw [hnm] - rfl - -def toInvolution {l : List 𝓕} (c : Contractions l) : Fin l.length → Fin l.length := - fun i => - if h : Option.isSome (c.toFinOptionMap i) then Option.get (c.toFinOptionMap i) h else i - -lemma toInvolution_of_isSome {l : List 𝓕} {c : Contractions l} {i : Fin l.length} - (hs : Option.isSome (c.toFinOptionMap i)) : - c.toInvolution i = Option.get (c.toFinOptionMap i) hs := by - dsimp [toInvolution] - rw [dif_pos hs] - -lemma toInvolution_of_eq_none {l : List 𝓕} {c : Contractions l} {i : Fin l.length} - (hs : (c.toFinOptionMap i) = none) : - c.toInvolution i = i := by - dsimp [toInvolution] - rw [dif_neg] - simp_all - -lemma toInvolution_image_isSome {l : List 𝓕} {c : Contractions l} {i : Fin l.length} - (hs : Option.isSome (c.toFinOptionMap i)) : - Option.isSome (c.toFinOptionMap (c.toInvolution i)) := by - dsimp [toInvolution] - rw [dif_pos hs] - have hx := toFinOptionMap_involutive c i ((c.toFinOptionMap i).get hs) - simp at hx - rw [hx] - rfl - -lemma toInvolution_involutive {l : List 𝓕} (c : Contractions l) : - Function.Involutive c.toInvolution := by - intro i - by_cases h : Option.isSome (c.toFinOptionMap i) - · have hx := toFinOptionMap_involutive c i ((c.toFinOptionMap i).get h) - rw [toInvolution_of_isSome (toInvolution_image_isSome h)] - simp only [Option.some_get, forall_const] at hx - simp only [toInvolution_of_isSome h, hx, Option.get_some] - · simp at h - rw [toInvolution_of_eq_none h] - rw [toInvolution_of_eq_none h] - -def involutionEquiv1 {l : List 𝓕} : - {f : Fin l.length → Fin l.length // Function.Involutive f } ≃ - {f : Fin l.length → Option (Fin l.length) // (∀ i, f i ≠ some i) ∧ - ∀ i j, f i = some j → f j = some i} where - toFun f := ⟨fun i => if h : f.1 i = i then none else f.1 i, - fun i => by - simp, - fun i j => by - simp - intro hi hij - subst hij - rw [f.2] - simp - exact fun a => hi (id (Eq.symm a))⟩ - invFun f := ⟨fun i => if h : (f.1 i).isSome then Option.get (f.1 i) h else i, - fun i => by - by_cases h : Option.isSome (f.1 i) - · simp [h] - have hf := f.2.2 i (Option.get (f.1 i) h) (by simp) - simp [hf] - · simp - rw [dif_neg] - · rw [dif_neg h] - · rw [dif_neg h] - exact h⟩ - left_inv f := by - simp - ext i - simp - by_cases hf : f.1 i = i - · simp [hf] - · simp [hf] - right_inv f := by - simp - ext1 - simp - funext i - obtain ⟨val, property⟩ := f - obtain ⟨left, right⟩ := property - simp_all only [ne_eq] - split - next h => - ext a : 1 - simp_all only [Option.mem_def, reduceCtorEq, false_iff] - apply Aesop.BuiltinRules.not_intro - intro a_1 - simp_all only [Option.isSome_some, Option.get_some, forall_const] - next h => - simp_all only [not_forall] - obtain ⟨w, h⟩ := h - simp_all only [↓reduceDIte, Option.some_get] - -def involutionCons (n : ℕ): - {f : Fin n.succ → Fin n.succ // Function.Involutive f } ≃ - (f : {f : Fin n → Fin n // Function.Involutive f}) × {i : Option (Fin n) // - ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} where - toFun f := ⟨⟨ - fun i => - if h : f.1 i.succ = 0 then i - else Fin.pred (f.1 i.succ) h , by - intro i - by_cases h : f.1 i.succ = 0 - · simp [h] - · simp [h] - simp [f.2 i.succ] - intro h - exact False.elim (Fin.succ_ne_zero i h)⟩, - ⟨if h : f.1 0 = 0 then none else Fin.pred (f.1 0) h , by - by_cases h0 : f.1 0 = 0 - · simp [h0] - · simp [h0] - refine fun h => False.elim (h (f.2 0))⟩⟩ - invFun f := ⟨ - if h : (f.2.1).isSome then - Fin.cons (f.2.1.get h).succ (Function.update (Fin.succ ∘ f.1.1) (f.2.1.get h) 0) - else - Fin.cons 0 (Fin.succ ∘ f.1.1) - , by - by_cases hs : (f.2.1).isSome - · simp only [Nat.succ_eq_add_one, hs, ↓reduceDIte, Fin.coe_eq_castSucc] - let a := f.2.1.get hs - change Function.Involutive (Fin.cons a.succ (Function.update (Fin.succ ∘ ↑f.fst) a 0)) - intro i - rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ - · subst hi - rw [Fin.cons_zero, Fin.cons_succ] - simp - · subst hj - rw [Fin.cons_succ] - by_cases hja : j = a - · subst hja - simp - · rw [Function.update_apply ] - rw [if_neg hja] - simp - have hf2 := f.2.2 hs - change f.1.1 a = a at hf2 - have hjf1 : f.1.1 j ≠ a := by - by_contra hn - have haj : j = f.1.1 a := by - rw [← hn] - rw [f.1.2] - rw [hf2] at haj - exact hja haj - rw [Function.update_apply, if_neg hjf1] - simp - rw [f.1.2] - · simp [hs] - intro i - rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ - · subst hi - simp - · subst hj - simp - rw [f.1.2]⟩ - left_inv f := by - match f with - | ⟨f, hf⟩ => - simp - ext i - by_cases h0 : f 0 = 0 - · simp [h0] - rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ - · subst hi - simp [h0] - · subst hj - simp [h0] - by_cases hj : f j.succ =0 - · rw [← h0] at hj - have hn := Function.Involutive.injective hf hj - exact False.elim (Fin.succ_ne_zero j hn) - · simp [hj] - rw [Fin.ext_iff] at hj - simp at hj - omega - · rw [if_neg h0] - by_cases hf' : i = f 0 - · subst hf' - simp - rw [hf] - simp - · rw [Function.update_apply, if_neg hf'] - rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ - · subst hi - simp - · subst hj - simp - by_cases hj : f j.succ =0 - · rw [← hj] at hf' - rw [hf] at hf' - simp at hf' - · simp [hj] - rw [Fin.ext_iff] at hj - simp at hj - omega - right_inv f := by - match f with - | ⟨⟨f, hf⟩, ⟨f0, hf0⟩⟩ => - ext i - · simp - by_cases hs : f0.isSome - · simp [hs] - by_cases hi : i = f0.get hs - · simp [hi, Function.update_apply] - exact Eq.symm (Fin.val_eq_of_eq (hf0 hs)) - · simp [hi] - split - · rename_i h - exact False.elim (Fin.succ_ne_zero (f i) h) - · rfl - · simp [hs] - split - · rename_i h - exact False.elim (Fin.succ_ne_zero (f i) h) - · rfl - · simp only [Nat.succ_eq_add_one, Option.mem_def, - Option.dite_none_left_eq_some, Option.some.injEq] - by_cases hs : f0.isSome - · simp only [hs, ↓reduceDIte] - simp [Fin.cons_zero] - have hx : ¬ (f0.get hs).succ = 0 := (Fin.succ_ne_zero (f0.get hs)) - simp [hx] - refine Iff.intro (fun hi => ?_) (fun hi => ?_) - · rw [← hi] - exact - Option.eq_some_of_isSome - (Eq.mpr_prop (Eq.refl (f0.isSome = true)) - (of_eq_true (Eq.trans (congrArg (fun x => x = true) hs) (eq_self true)))) - · subst hi - exact rfl - · simp [hs] - simp at hs - subst hs - exact ne_of_beq_false rfl - -lemma involutionCons_ext {n : ℕ} {f1 f2 : (f : {f : Fin n → Fin n // Function.Involutive f}) × {i : Option (Fin n) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)}} - (h1 : f1.1 = f2.1) (h2 : f1.2 = Equiv.subtypeEquivRight (by rw [h1]; simp) f2.2) : f1 = f2 := by - cases f1 - cases f2 - simp at h1 h2 - subst h1 - rename_i fst snd snd_1 - simp_all only [Sigma.mk.inj_iff, heq_eq_eq, true_and] - obtain ⟨val, property⟩ := fst - obtain ⟨val_1, property_1⟩ := snd - obtain ⟨val_2, property_2⟩ := snd_1 - simp_all only - rfl - -def uncontractedEquiv' {l : List 𝓕} (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) : - {i : Option (Fin l.length) // - ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} ≃ - Option (Fin (Finset.univ.filter fun i => f.1 i = i).card) := by - let e1 : {i : Option (Fin l.length) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} - ≃ Option {i : Fin l.length // f.1 i = i} := - { toFun := fun i => match i with - | ⟨some i, h⟩ => some ⟨i, by simpa using h⟩ - | ⟨none, h⟩ => none - invFun := fun i => match i with - | some ⟨i, h⟩ => ⟨some i, by simpa using h⟩ - | none => ⟨none, by simp⟩ - left_inv := by - intro a - cases a - aesop - right_inv := by - intro a - cases a - rfl - simp_all only [Subtype.coe_eta] } - let s : Finset (Fin l.length) := Finset.univ.filter fun i => f.1 i = i - let e2' : { i : Fin l.length // f.1 i = i} ≃ {i // i ∈ s} := by - refine Equiv.subtypeEquivProp ?h - funext i - simp [s] - let e2 : {i // i ∈ s} ≃ Fin (Finset.card s) := by - refine (Finset.orderIsoOfFin _ ?_).symm.toEquiv - simp [s] - refine e1.trans (Equiv.optionCongr (e2'.trans (e2))) - - -lemma uncontractedEquiv'_none_image_zero {φs : List 𝓕} {φ : 𝓕} : - {f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}} - → uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2 = none - → f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ = ⟨0, Nat.zero_lt_succ φs.length⟩ := by - intro f h - simp only [Nat.succ_eq_add_one, involutionCons, Equiv.coe_fn_mk, uncontractedEquiv', - Option.isSome_some, Option.get_some, Option.isSome_none, Equiv.trans_apply, - Equiv.optionCongr_apply, Equiv.coe_trans, RelIso.coe_fn_toEquiv, Option.map_eq_none'] at h - simp_all only [List.length_cons, Fin.zero_eta] - obtain ⟨val, property⟩ := f - simp_all only [List.length_cons] - split at h - next i i_1 h_1 heq => - split at heq - next h_2 => simp_all only [reduceCtorEq] - next h_2 => simp_all only [reduceCtorEq] - next i h_1 heq => - split at heq - next h_2 => simp_all only - next h_2 => simp_all only [Subtype.mk.injEq, reduceCtorEq] - -lemma uncontractedEquiv'_cast {l : List 𝓕} {f1 f2 : {f : Fin l.length → Fin l.length // Function.Involutive f}} - (hf : f1 = f2): - uncontractedEquiv' f1 = (Equiv.subtypeEquivRight (by rw [hf]; simp)).trans - ((uncontractedEquiv' f2).trans (Equiv.optionCongr (finCongr (by rw [hf])))):= by - subst hf - simp - rfl - -lemma uncontractedEquiv'_none_succ {φs : List 𝓕} {φ : 𝓕} : - {f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}} - → uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2 = none - → ∀ (x : Fin φs.length), f.1 x.succ = x.succ ↔ (involutionCons φs.length f).1.1 x = x := by - intro f h x - simp [involutionCons] - have hn' := uncontractedEquiv'_none_image_zero h - have hx : ¬ f.1 x.succ = ⟨0, Nat.zero_lt_succ φs.length⟩:= by - rw [← hn'] - exact fun hn => Fin.succ_ne_zero x (Function.Involutive.injective f.2 hn) - apply Iff.intro - · intro h2 h3 - rw [Fin.ext_iff] - simp [h2] - · intro h2 - have h2' := h2 hx - conv_rhs => rw [← h2'] - simp - - -lemma uncontractedEquiv'_isSome_image_zero {φs : List 𝓕} {φ : 𝓕} : - {f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}} - → (uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2).isSome - → ¬ f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ = ⟨0, Nat.zero_lt_succ φs.length⟩ := by - intro f hf - simp [uncontractedEquiv', involutionCons] at hf - simp_all only [List.length_cons, Fin.zero_eta] - obtain ⟨val, property⟩ := f - simp_all only [List.length_cons] - apply Aesop.BuiltinRules.not_intro - intro a - simp_all only [↓reduceDIte, Option.isSome_none, Bool.false_eq_true] - - - -def uncontractedFromInvolution: {φs : List 𝓕} → - (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) → - {l : List 𝓕 // l.length = (Finset.univ.filter fun i => f.1 i = i).card} - | [], _ => ⟨[], by simp⟩ - | φ :: φs, f => - let luc := uncontractedFromInvolution (involutionCons φs.length f).fst - let n' := uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2 - if hn : n' = none then - have hn' := uncontractedEquiv'_none_image_zero (φs := φs) (φ := φ) (f := f) hn - ⟨optionEraseZ luc φ none, by - simp [optionEraseZ] - rw [← luc.2] - conv_rhs => rw [Finset.card_filter] - rw [Fin.sum_univ_succ] - conv_rhs => erw [if_pos hn'] - ring_nf - simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, Nat.cast_id, - add_right_inj] - rw [Finset.card_filter] - apply congrArg - funext i - refine ite_congr ?h.h.h₁ (congrFun rfl) (congrFun rfl) - rw [uncontractedEquiv'_none_succ hn]⟩ - else - let n := n'.get (Option.isSome_iff_ne_none.mpr hn) - let np : Fin luc.1.length := ⟨n.1, by - rw [luc.2] - exact n.prop⟩ - ⟨optionEraseZ luc φ (some np), by - let k' := (involutionCons φs.length f).2 - have hkIsSome : (k'.1).isSome := by - simp [n', uncontractedEquiv' ] at hn - split at hn - · simp_all only [reduceCtorEq, not_false_eq_true, Nat.succ_eq_add_one, Option.isSome_some, k'] - · simp_all only [not_true_eq_false] - let k := k'.1.get hkIsSome - rw [optionEraseZ_some_length] - have hksucc : k.succ = f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ := by - simp [k, k', involutionCons] - have hzero : ⟨0, Nat.zero_lt_succ φs.length⟩ = f.1 k.succ := by - rw [hksucc] - rw [f.2] - have hkcons : ((involutionCons φs.length) f).1.1 k = k := by - exact k'.2 hkIsSome - have hksuccNe : f.1 k.succ ≠ k.succ := by - conv_rhs => rw [hksucc] - exact fun hn => Fin.succ_ne_zero k (Function.Involutive.injective f.2 hn ) - have hluc : 1 ≤ luc.1.length := by - simp - use k - simp [involutionCons] - rw [hksucc, f.2] - simp - rw [propext (Nat.sub_eq_iff_eq_add' hluc)] - have h0 : ¬ f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ = ⟨0, Nat.zero_lt_succ φs.length⟩ := by - exact Option.isSome_dite'.mp hkIsSome - conv_rhs => - rw [Finset.card_filter] - erw [Fin.sum_univ_succ] - erw [if_neg h0] - simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, List.length_cons, - Nat.cast_id, zero_add] - conv_rhs => lhs; rw [Eq.symm (Fintype.sum_ite_eq' k fun j => 1)] - rw [← Finset.sum_add_distrib] - rw [Finset.card_filter] - apply congrArg - funext i - by_cases hik : i = k - · subst hik - simp [hkcons, hksuccNe] - · simp [hik] - refine ite_congr ?_ (congrFun rfl) (congrFun rfl) - simp [involutionCons] - have hfi : f.1 i.succ ≠ ⟨0, Nat.zero_lt_succ φs.length⟩ := by - rw [hzero] - by_contra hn - have hik' := (Function.Involutive.injective f.2 hn) - simp only [List.length_cons, Fin.succ_inj] at hik' - exact hik hik' - apply Iff.intro - · intro h - have h' := h hfi - conv_rhs => rw [← h'] - simp - · intro h hfi - simp [Fin.ext_iff] - rw [h] - simp⟩ - -lemma uncontractedFromInvolution_cons {φs : List 𝓕} {φ : 𝓕} - (f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) : - uncontractedFromInvolution f = - optionEraseZ (uncontractedFromInvolution (involutionCons φs.length f).fst) φ - (Option.map (finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm)) - (uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2)) := by - let luc := uncontractedFromInvolution (involutionCons φs.length f).fst - let n' := uncontractedEquiv' (involutionCons φs.length f).1 (involutionCons φs.length f).2 - change _ = optionEraseZ luc φ - (Option.map (finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm)) n') - dsimp [uncontractedFromInvolution] - by_cases hn : n' = none - · have hn' := hn - simp [n'] at hn' - simp [hn'] - rw [hn] - rfl - · have hn' := hn - simp [n'] at hn' - simp [hn'] - congr - simp [n'] - simp_all only [Nat.succ_eq_add_one, not_false_eq_true, n', luc] - obtain ⟨val, property⟩ := f - obtain ⟨val_1, property_1⟩ := luc - simp_all only [Nat.succ_eq_add_one, List.length_cons] - ext a : 1 - simp_all only [Option.mem_def, Option.some.injEq, Option.map_eq_some', finCongr_apply] - apply Iff.intro - · intro a_1 - subst a_1 - apply Exists.intro - · apply And.intro - on_goal 2 => {rfl - } - · simp_all only [Option.some_get] - · intro a_1 - obtain ⟨w, h⟩ := a_1 - obtain ⟨left, right⟩ := h - subst right - simp_all only [Option.get_some] - rfl - -def auxCongr : {φs: List 𝓕} → {φsᵤₙ φsᵤₙ' : List 𝓕} → (h : φsᵤₙ = φsᵤₙ') → - ContractionsAux φs φsᵤₙ ≃ ContractionsAux φs φsᵤₙ' - | _, _, _, Eq.refl _ => Equiv.refl _ - -lemma auxCongr_ext {φs: List 𝓕} {c c2 : Contractions φs} (h : c.1 = c2.1) - (hx : c.2 = auxCongr h.symm c2.2) : c = c2 := by - cases c - cases c2 - simp at h - subst h - simp [auxCongr] at hx - subst hx - rfl - - -lemma uncontractedEquiv'_cast' {l : List 𝓕} {f1 f2 : {f : Fin l.length → Fin l.length // Function.Involutive f}} - {N : ℕ} (hf : f1 = f2) (n : Option (Fin N)) (hn1 : N = (Finset.filter (fun i => f1.1 i = i) Finset.univ).card) - (hn2 : N = (Finset.filter (fun i => f2.1 i = i) Finset.univ).card): - HEq ((uncontractedEquiv' f1).symm (Option.map (finCongr hn1) n)) ((uncontractedEquiv' f2).symm (Option.map (finCongr hn2) n)) := by - subst hf - rfl - -def toInvolution' : {φs : List 𝓕} → (c : Contractions φs) → - {f : {f : Fin φs.length → Fin φs.length // Function.Involutive f} // - uncontractedFromInvolution f = c.1} - | [], ⟨[], ContractionsAux.nil⟩ => ⟨⟨fun i => i, by - intro i - simp⟩, by rfl⟩ - | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) n c⟩ => by - let ⟨⟨f', hf1⟩, hf2⟩ := toInvolution' ⟨aux, c⟩ - let n' : Option (Fin (uncontractedFromInvolution ⟨f', hf1⟩).1.length) := - Option.map (finCongr (by rw [hf2])) n - let F := (involutionCons φs.length).symm ⟨⟨f', hf1⟩, - (uncontractedEquiv' ⟨f', hf1⟩).symm - (Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ - refine ⟨F, ?_⟩ - have hF0 : ((involutionCons φs.length) F) = ⟨⟨f', hf1⟩, - (uncontractedEquiv' ⟨f', hf1⟩).symm - (Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ := by - simp [F] - have hF1 : ((involutionCons φs.length) F).fst = ⟨f', hf1⟩ := by - rw [hF0] - have hF2L : ((uncontractedFromInvolution ⟨f', hf1⟩)).1.length = - (Finset.filter (fun i => ((involutionCons φs.length) F).1.1 i = i) Finset.univ).card := by - apply Eq.trans ((uncontractedFromInvolution ⟨f', hf1⟩)).2 - congr - rw [hF1] - have hF2 : ((involutionCons φs.length) F).snd = (uncontractedEquiv' ((involutionCons φs.length) F).fst).symm - (Option.map (finCongr hF2L) n') := by - rw [@Sigma.subtype_ext_iff] at hF0 - ext1 - rw [hF0.2] - simp - congr 1 - · rw [hF1] - · refine uncontractedEquiv'_cast' ?_ n' _ _ - rw [hF1] - rw [uncontractedFromInvolution_cons] - have hx := (toInvolution' ⟨aux, c⟩).2 - simp at hx - simp - refine optionEraseZ_ext ?_ ?_ ?_ - · dsimp [F] - rw [Equiv.apply_symm_apply] - simp - rw [← hx] - simp_all only - · rfl - · simp [hF2] - dsimp [n'] - simp [finCongr] - simp only [Nat.succ_eq_add_one, id_eq, eq_mpr_eq_cast, F, n'] - ext a : 1 - simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans, - Fin.cast_eq_self, exists_eq_right] - -lemma toInvolution'_length {φs φsᵤₙ : List 𝓕} {c : ContractionsAux φs φsᵤₙ} : - φsᵤₙ.length = (Finset.filter (fun i => (toInvolution' ⟨φsᵤₙ, c⟩).1.1 i = i) Finset.univ).card - := by - have h2 := (toInvolution' ⟨φsᵤₙ, c⟩).2 - simp at h2 - conv_lhs => rw [← h2] - exact Mathlib.Vector.length_val (uncontractedFromInvolution (toInvolution' ⟨φsᵤₙ, c⟩).1) - -lemma toInvolution'_cons {φs φsᵤₙ : List 𝓕} {φ : 𝓕} - (c : ContractionsAux φs φsᵤₙ) (n : Option (Fin (φsᵤₙ.length))) : - (toInvolution' ⟨optionEraseZ φsᵤₙ φ n, ContractionsAux.cons n c⟩).1 - = (involutionCons φs.length).symm ⟨(toInvolution' ⟨φsᵤₙ, c⟩).1, - (uncontractedEquiv' (toInvolution' ⟨φsᵤₙ, c⟩).1).symm - (Option.map (finCongr (toInvolution'_length)) n)⟩ := by - dsimp [toInvolution'] - congr 3 - rw [Option.map_map] - simp [finCongr] - rfl - -lemma toInvolution_consEquiv {φs : List 𝓕} {φ : 𝓕} - (c : Contractions φs) (n : Option (Fin (c.normalize.length))) : - (toInvolution' ((consEquiv (φ := φ)).symm ⟨c, n⟩)).1 = - (involutionCons φs.length).symm ⟨(toInvolution' c).1, - (uncontractedEquiv' (toInvolution' c).1).symm - (Option.map (finCongr (toInvolution'_length)) n)⟩ := by - erw [toInvolution'_cons] - rfl - -def fromInvolutionAux : {l : List 𝓕} → - (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) → - ContractionsAux l (uncontractedFromInvolution f) - | [] => fun _ => ContractionsAux.nil - | _ :: φs => fun f => - let f' := involutionCons φs.length f - let c' := fromInvolutionAux f'.1 - let n' := Option.map (finCongr ((uncontractedFromInvolution f'.fst).2.symm)) - (uncontractedEquiv' f'.1 f'.2) - auxCongr (uncontractedFromInvolution_cons f).symm (ContractionsAux.cons n' c') - -def fromInvolution {φs : List 𝓕} (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) : - Contractions φs := ⟨uncontractedFromInvolution f, fromInvolutionAux f⟩ - -lemma fromInvolution_cons {φs : List 𝓕} {φ : 𝓕} - (f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) : - let f' := involutionCons φs.length f - fromInvolution f = consEquiv.symm - ⟨fromInvolution f'.1, Option.map (finCongr ((uncontractedFromInvolution f'.fst).2.symm)) - (uncontractedEquiv' f'.1 f'.2)⟩ := by - refine auxCongr_ext ?_ ?_ - · dsimp [fromInvolution] - rw [uncontractedFromInvolution_cons] - rfl - · dsimp [fromInvolution, fromInvolutionAux] - rfl - -lemma fromInvolution_of_involutionCons - {φs : List 𝓕} {φ : 𝓕} - (f : {f : Fin (φs ).length → Fin (φs).length // Function.Involutive f}) - (n : { i : Option (Fin φs.length) // ∀ (h : i.isSome = true), f.1 (i.get h) = i.get h }): - fromInvolution (φs := φ :: φs) ((involutionCons φs.length).symm ⟨f, n⟩) = - consEquiv.symm - ⟨fromInvolution f, Option.map (finCongr ((uncontractedFromInvolution f).2.symm)) - (uncontractedEquiv' f n)⟩ := by - rw [fromInvolution_cons] - congr 1 - simp - rw [Equiv.apply_symm_apply] - - - - -lemma toInvolution_fromInvolution : {φs : List 𝓕} → (c : Contractions φs) → - fromInvolution (toInvolution' c) = c - | [], ⟨[], ContractionsAux.nil⟩ => rfl - | φ :: φs, ⟨_, .cons (φsᵤₙ := φsᵤₙ) n c⟩ => by - rw [toInvolution'_cons] - rw [fromInvolution_of_involutionCons] - rw [Equiv.symm_apply_eq] - dsimp [consEquiv] - refine consEquiv_ext ?_ ?_ - · exact toInvolution_fromInvolution ⟨φsᵤₙ, c⟩ - · simp [finCongr] - ext a : 1 - simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans, - Fin.cast_eq_self, exists_eq_right] - -lemma fromInvolution_toInvolution : {φs : List 𝓕} → (f : {f : Fin (φs ).length → Fin (φs).length // Function.Involutive f}) - → toInvolution' (fromInvolution f) = f - | [], _ => by - ext x - exact Fin.elim0 x - | φ :: φs, f => by - rw [fromInvolution_cons] - rw [toInvolution_consEquiv] - erw [Equiv.symm_apply_eq] - have hx := fromInvolution_toInvolution ((involutionCons φs.length) f).fst - apply involutionCons_ext ?_ ?_ - · simp only [Nat.succ_eq_add_one, List.length_cons] - exact hx - · simp only [Nat.succ_eq_add_one, Option.map_map, List.length_cons] - rw [Equiv.symm_apply_eq] - conv_rhs => - lhs - rw [uncontractedEquiv'_cast hx] - simp [Nat.succ_eq_add_one,- eq_mpr_eq_cast, Equiv.trans_apply, -Equiv.optionCongr_apply] - rfl - -def equivInvolutions {φs : List 𝓕} : - Contractions φs ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f} where - toFun := fun c => toInvolution' c - invFun := fromInvolution - left_inv := toInvolution_fromInvolution - right_inv := fromInvolution_toInvolution - - -lemma isFull_iff_uncontractedFromInvolution_empty {φs : List 𝓕} (c : Contractions φs) : - IsFull c ↔ (uncontractedFromInvolution (equivInvolutions c)).1 = [] := by - let l := toInvolution' c - erw [l.2] - rfl - -lemma isFull_iff_filter_card_involution_zero {φs : List 𝓕} (c : Contractions φs) : - IsFull c ↔ (Finset.univ.filter fun i => (equivInvolutions c).1 i = i).card = 0 := by - rw [isFull_iff_uncontractedFromInvolution_empty, List.ext_get_iff] - simp - -lemma isFull_iff_involution_no_fixed_points {φs : List 𝓕} (c : Contractions φs) : - IsFull c ↔ ∀ (i : Fin φs.length), (equivInvolutions c).1 i ≠ i := by - rw [isFull_iff_filter_card_involution_zero] - simp - rw [Finset.filter_eq_empty_iff] - apply Iff.intro - · intro h - intro i - refine h (Finset.mem_univ i) - · intro i h - exact fun a => i h - -def involutionNoFixed1 {n : ℕ} : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f - ∧ ∀ i, f i ≠ i} ≃ Σ (k : Fin (2 * n + 1)), - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f - ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} where - toFun f := ⟨(f.1 0).pred (f.2.2 0), ⟨f.1, f.2.1, by simpa using f.2.2⟩⟩ - invFun f := ⟨f.2.1, ⟨f.2.2.1, f.2.2.2.1⟩⟩ - left_inv f := by - rfl - right_inv f := by - simp - ext - · simp - rw [f.2.2.2.2] - simp - · simp - -def involutionNoFixed2 {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f - ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} ≃ - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive (e.symm ∘ f ∘ e) ∧ - (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} where - toFun f := ⟨e ∘ f.1 ∘ e.symm, by - intro i - simp - rw [f.2.1], by - simpa using f.2.2.1, by simpa using f.2.2.2⟩ - invFun f := ⟨e.symm ∘ f.1 ∘ e, by - intro i - simp - have hf2 := f.2.1 i - simpa using hf2, by - simpa using f.2.2.1, by simpa using f.2.2.2⟩ - left_inv f := by - ext i - simp - right_inv f := by - ext i - simp -def involutionNoFixed3 {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive (e.symm ∘ f ∘ e) ∧ - (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} ≃ - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} := by - refine Equiv.subtypeEquivRight ?_ - intro f - have h1 : Function.Involutive (⇑e.symm ∘ f ∘ ⇑e) ↔ Function.Involutive f := by - apply Iff.intro - · intro h i - have hi := h (e.symm i) - simpa using hi - · intro h i - have hi := h (e i) - simp [hi] - rw [h1] - simp - intro h1 h2 - apply Iff.intro - · intro h i - have hi := h (e.symm i) - simpa using hi - · intro h i - have hi := h (e i) - by_contra hn - nth_rewrite 2 [← hn] at hi - simp at hi - - -def involutionNoFixed4 {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} - ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f (e 0) = e k.succ} := by - refine Equiv.subtypeEquivRight ?_ - simp - intro f hi h1 - exact Equiv.symm_apply_eq e - -def involutionNoFixed5 {n : ℕ} (k : Fin (2 * n + 1)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f 0 = k.succ} - ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f 0 = 1} := by - refine Equiv.trans (involutionNoFixed2 k (Equiv.swap k.succ 1)) ?_ - refine Equiv.trans (involutionNoFixed3 k (Equiv.swap k.succ 1)) ?_ - refine Equiv.trans (involutionNoFixed4 k (Equiv.swap k.succ 1)) ?_ - refine Equiv.subtypeEquivRight ?_ - simp - intro f hi h1 - rw [Equiv.swap_apply_of_ne_of_ne] - · exact Ne.symm (Fin.succ_ne_zero k) - · exact Fin.zero_ne_one - -def involutionNoFixed6 {n : ℕ} : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f 0 = 1} ≃ - {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ - (∀ i, f i ≠ i)} where - toFun f := by - have hf1 : f.1 1 = 0 := by - have hf := f.2.2.2 - simp [← hf] - rw [f.2.1] - let f' := f.1 ∘ Fin.succ ∘ Fin.succ - have hf' (i : Fin (2 * n)) : f' i ≠ 0 := by - simp [f'] - simp [← hf1] - by_contra hn - have hn' := Function.Involutive.injective f.2.1 hn - simp [Fin.ext_iff] at hn' - let f'' := fun i => (f' i).pred (hf' i) - have hf'' (i : Fin (2 * n)) : f'' i ≠ 0 := by - simp [f''] - rw [@Fin.pred_eq_iff_eq_succ] - simp [f'] - simp [← f.2.2.2 ] - by_contra hn - have hn' := Function.Involutive.injective f.2.1 hn - simp [Fin.ext_iff] at hn' - let f''' := fun i => (f'' i).pred (hf'' i) - refine ⟨f''', ?_, ?_⟩ - · intro i - simp [f''', f'', f'] - simp [f.2.1 i.succ.succ] - · intro i - simp [f''', f'', f'] - rw [@Fin.pred_eq_iff_eq_succ] - rw [@Fin.pred_eq_iff_eq_succ] - exact f.2.2.1 i.succ.succ - invFun f := by - let f' := fun (i : Fin (2 * n.succ))=> - match i with - | ⟨0, h⟩ => 1 - | ⟨1, h⟩ => 0 - | ⟨(Nat.succ (Nat.succ n)), h⟩ => (f.1 ⟨n, by omega⟩).succ.succ - refine ⟨f', ?_, ?_, ?_⟩ - · intro i - match i with - | ⟨0, h⟩ => rfl - | ⟨1, h⟩ => rfl - | ⟨(Nat.succ (Nat.succ m)), h⟩ => - simp [f'] - split - · rename_i h - simp at h - exact False.elim (Fin.succ_ne_zero (f.1 ⟨m, _⟩).succ h) - · rename_i h - simp [Fin.ext_iff] at h - · rename_i h - rename_i x r - simp_all [Fin.ext_iff] - have hfn {a b : ℕ} {ha : a < 2 * n} {hb : b < 2 * n} - (hab : ↑(f.1 ⟨a, ha⟩) = b): ↑(f.1 ⟨b, hb⟩) = a := by - have ht : f.1 ⟨a, ha⟩ = ⟨b, hb⟩ := by - simp [hab, Fin.ext_iff] - rw [← ht, f.2.1] - exact hfn h - · intro i - match i with - | ⟨0, h⟩ => - simp [f'] - split - · rename_i h - simp - · rename_i h - simp [Fin.ext_iff] at h - · rename_i h - simp [Fin.ext_iff] at h - | ⟨1, h⟩ => - simp [f'] - split - · rename_i h - simp at h - · rename_i h - simp - · rename_i h - simp [Fin.ext_iff] at h - | ⟨(Nat.succ (Nat.succ m)), h⟩ => - simp [f', Fin.ext_iff] - have hf:= f.2.2 ⟨m, by exact Nat.add_lt_add_iff_right.mp h⟩ - simp [Fin.ext_iff] at hf - omega - · simp [f'] - split - · rename_i h - simp - · rename_i h - simp at h - · rename_i h - simp [Fin.ext_iff] at h - left_inv f := by - have hf1 : f.1 1 = 0 := by - have hf := f.2.2.2 - simp [← hf] - rw [f.2.1] - simp - ext i - simp - split - · simp - rw [f.2.2.2] - simp - · simp - rw [hf1] - simp - · rfl - right_inv f := by - simp - ext i - simp - split - · rename_i h - simp [Fin.ext_iff] at h - · rename_i h - simp [Fin.ext_iff] at h - · rename_i h - simp - congr - apply congrArg - simp_all [Fin.ext_iff] - - - -def involutionNoFixed7 {n : ℕ} (k : Fin (2 * n + 1)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f 0 = k.succ} - ≃ {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by - refine Equiv.trans (involutionNoFixed5 k) involutionNoFixed6 - -def involutionNoFixed8 {n : ℕ} : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} - ≃ Σ (_ : Fin (2 * n + 1)), {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by - refine Equiv.trans involutionNoFixed1 ?_ - refine Equiv.sigmaCongrRight involutionNoFixed7 - -def involutionNoFixed9 {n : ℕ} : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} - ≃ Fin (2 * n + 1) × {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by - refine Equiv.trans involutionNoFixed8 ?_ - exact Equiv.sigmaEquivProd (Fin (2 * n + 1)) - { f // Function.Involutive f ∧ ∀ (i : Fin (2 * n)), f i ≠ i} - -instance {n : ℕ} : Fintype { f // Function.Involutive f ∧ ∀ (i : Fin ( n)), f i ≠ i } := by - haveI : DecidablePred fun x => Function.Involutive x := by - intro f - apply Fintype.decidableForallFintype (α := Fin (n)) - haveI : DecidablePred fun x => Function.Involutive x ∧ ∀ (i : Fin ( n)), x i ≠ i := by - intro x - apply instDecidableAnd - apply Subtype.fintype - -lemma involutionNoFixed_card_succ {n : ℕ} : - Fintype.card {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} - = (2 * n + 1) * Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by - rw [Fintype.card_congr (involutionNoFixed9)] - rw [Fintype.card_prod ] - congr - exact Fintype.card_fin (2 * n + 1) - - -open Nat - - -lemma involutionNoFixed_card : (n : ℕ) → - Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} - = (2 * n - 1)‼ - | 0 => rfl - | Nat.succ n => by - rw [involutionNoFixed_card_succ] - rw [involutionNoFixed_card n] - exact Eq.symm (Nat.doubleFactorial_add_one (Nat.mul 2 n)) - -lemma involutionNoFixed_even : (n : ℕ) → (he : Even n) → - Fintype.card {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} = (n - 1)‼ := by - intro n he - obtain ⟨r, hr⟩ := he - have hr' : n = 2 * r := by omega - subst hr' - exact involutionNoFixed_card r - -def isFullInvolutionEquiv {φs : List 𝓕} : - {c : Contractions φs // IsFull c} ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f ∧ (∀ i, f i ≠ i)} where - toFun c := ⟨equivInvolutions c.1, by - apply And.intro (equivInvolutions c.1).2 - rw [← isFull_iff_involution_no_fixed_points] - exact c.2 - ⟩ - invFun f := ⟨equivInvolutions.symm ⟨f.1, f.2.1⟩, by - rw [isFull_iff_involution_no_fixed_points] - simpa using f.2.2⟩ - left_inv c := by simp - right_inv f := by simp - -lemma card_of_full_contractions {φs : List 𝓕} (he : Even φs.length ) : - Fintype.card {c : Contractions φs // IsFull c} = (φs.length - 1)‼ := by - rw [Fintype.card_congr (isFullInvolutionEquiv (φs := φs))] - exact involutionNoFixed_even φs.length he - - /-- A structure specifying when a type `I` and a map `f :I → Type` corresponds to the splitting of a fields `φ` into a creation `n` and annihlation part `p`. -/ structure Splitting (f : 𝓕 → Type) [∀ i, Fintype (f i)] diff --git a/HepLean/PerturbationTheory/Contractions/Card.lean b/HepLean/PerturbationTheory/Contractions/Card.lean new file mode 100644 index 00000000..bee0d2fb --- /dev/null +++ b/HepLean/PerturbationTheory/Contractions/Card.lean @@ -0,0 +1,40 @@ +/- +Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Tooby-Smith +-/ +import HepLean.PerturbationTheory.Contractions.Involutions +/-! + +# Cardinality of full contractions + +-/ + +namespace Wick +namespace Contractions + +open HepLean.Fin +open Nat + +/-- There are `(φs.length - 1)‼` full contractions of a list `φs` with an even number of fields. -/ +lemma card_of_full_contractions_even {φs : List 𝓕} (he : Even φs.length ) : + Fintype.card {c : Contractions φs // IsFull c} = (φs.length - 1)‼ := by + rw [Fintype.card_congr (isFullInvolutionEquiv (φs := φs))] + exact involutionNoFixed_card_even φs.length he + +/-- There are no full contractions of a list with an odd number of fields. -/ +lemma card_of_full_contractions_odd {φs : List 𝓕} (ho : Odd φs.length ) : + Fintype.card {c : Contractions φs // IsFull c} = 0 := by + rw [Fintype.card_eq_zero_iff, isEmpty_subtype] + intro c + simp only [IsFull] + by_contra hn + have hc := uncontracted_length_even_iff c + rw [hn] at hc + simp at hc + rw [← Nat.not_odd_iff_even] at hc + exact hc ho + +end Contractions + +end Wick diff --git a/HepLean/PerturbationTheory/Contractions/Involutions.lean b/HepLean/PerturbationTheory/Contractions/Involutions.lean index 5b0b9118..c5d67587 100644 --- a/HepLean/PerturbationTheory/Contractions/Involutions.lean +++ b/HepLean/PerturbationTheory/Contractions/Involutions.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith -/ import HepLean.PerturbationTheory.Contractions.Basic -import HepLean.Meta.Informal.Basic /-! # Involutions @@ -23,6 +22,7 @@ is given by the OEIS sequence A000085. namespace Wick open HepLean.List +open HepLean.Fin open FieldStatistic variable {𝓕 : Type} @@ -30,13 +30,366 @@ namespace Contractions variable {l : List 𝓕} -informal_definition equivInvolution where - math :≈ "There is an isomorphism between the type of contractions of a list `l` and - the type of involutions from `Fin l.length` to `Fin l.length." +/-! + +## From Involution. + +-/ + +def uncontractedFromInvolution : {φs : List 𝓕} → + (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) → + {l : List 𝓕 // l.length = (Finset.univ.filter fun i => f.1 i = i).card} + | [], _ => ⟨[], by simp⟩ + | φ :: φs, f => + let luc := uncontractedFromInvolution (involutionCons φs.length f).fst + let n' := involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2 + if hn : n' = none then + have hn' := involutionAddEquiv_none_image_zero (n := φs.length) (f := f) hn + ⟨optionEraseZ luc φ none, by + simp [optionEraseZ] + rw [← luc.2] + conv_rhs => rw [Finset.card_filter] + rw [Fin.sum_univ_succ] + conv_rhs => erw [if_pos hn'] + ring_nf + simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, Nat.cast_id, + add_right_inj] + rw [Finset.card_filter] + apply congrArg + funext i + refine ite_congr ?h.h.h₁ (congrFun rfl) (congrFun rfl) + rw [involutionAddEquiv_none_succ hn]⟩ + else + let n := n'.get (Option.isSome_iff_ne_none.mpr hn) + let np : Fin luc.1.length := ⟨n.1, by + rw [luc.2] + exact n.prop⟩ + ⟨optionEraseZ luc φ (some np), by + let k' := (involutionCons φs.length f).2 + have hkIsSome : (k'.1).isSome := by + simp [n', involutionAddEquiv ] at hn + split at hn + · simp_all only [reduceCtorEq, not_false_eq_true, Nat.succ_eq_add_one, Option.isSome_some, k'] + · simp_all only [not_true_eq_false] + let k := k'.1.get hkIsSome + rw [optionEraseZ_some_length] + have hksucc : k.succ = f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ := by + simp [k, k', involutionCons] + have hzero : ⟨0, Nat.zero_lt_succ φs.length⟩ = f.1 k.succ := by + rw [hksucc] + rw [f.2] + have hkcons : ((involutionCons φs.length) f).1.1 k = k := by + exact k'.2 hkIsSome + have hksuccNe : f.1 k.succ ≠ k.succ := by + conv_rhs => rw [hksucc] + exact fun hn => Fin.succ_ne_zero k (Function.Involutive.injective f.2 hn ) + have hluc : 1 ≤ luc.1.length := by + simp + use k + simp [involutionCons] + rw [hksucc, f.2] + simp + rw [propext (Nat.sub_eq_iff_eq_add' hluc)] + have h0 : ¬ f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ = ⟨0, Nat.zero_lt_succ φs.length⟩ := by + exact Option.isSome_dite'.mp hkIsSome + conv_rhs => + rw [Finset.card_filter] + erw [Fin.sum_univ_succ] + erw [if_neg h0] + simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, List.length_cons, + Nat.cast_id, zero_add] + conv_rhs => lhs; rw [Eq.symm (Fintype.sum_ite_eq' k fun j => 1)] + rw [← Finset.sum_add_distrib] + rw [Finset.card_filter] + apply congrArg + funext i + by_cases hik : i = k + · subst hik + simp [hkcons, hksuccNe] + · simp [hik] + refine ite_congr ?_ (congrFun rfl) (congrFun rfl) + simp [involutionCons] + have hfi : f.1 i.succ ≠ ⟨0, Nat.zero_lt_succ φs.length⟩ := by + rw [hzero] + by_contra hn + have hik' := (Function.Involutive.injective f.2 hn) + simp only [List.length_cons, Fin.succ_inj] at hik' + exact hik hik' + apply Iff.intro + · intro h + have h' := h hfi + conv_rhs => rw [← h'] + simp + · intro h hfi + simp [Fin.ext_iff] + rw [h] + simp⟩ + +lemma uncontractedFromInvolution_cons {φs : List 𝓕} {φ : 𝓕} + (f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) : + uncontractedFromInvolution f = + optionEraseZ (uncontractedFromInvolution (involutionCons φs.length f).fst) φ + (Option.map (finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm)) + (involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2)) := by + let luc := uncontractedFromInvolution (involutionCons φs.length f).fst + let n' := involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2 + change _ = optionEraseZ luc φ + (Option.map (finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm)) n') + dsimp [uncontractedFromInvolution] + by_cases hn : n' = none + · have hn' := hn + simp [n'] at hn' + simp [hn'] + rw [hn] + rfl + · have hn' := hn + simp [n'] at hn' + simp [hn'] + congr + simp [n'] + simp_all only [Nat.succ_eq_add_one, not_false_eq_true, n', luc] + obtain ⟨val, property⟩ := f + obtain ⟨val_1, property_1⟩ := luc + simp_all only [Nat.succ_eq_add_one, List.length_cons] + ext a : 1 + simp_all only [Option.mem_def, Option.some.injEq, Option.map_eq_some', finCongr_apply] + apply Iff.intro + · intro a_1 + subst a_1 + apply Exists.intro + · apply And.intro + on_goal 2 => {rfl + } + · simp_all only [Option.some_get] + · intro a_1 + obtain ⟨w, h⟩ := a_1 + obtain ⟨left, right⟩ := h + subst right + simp_all only [Option.get_some] + rfl + +def fromInvolutionAux : {l : List 𝓕} → + (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) → + ContractionsAux l (uncontractedFromInvolution f) + | [] => fun _ => ContractionsAux.nil + | _ :: φs => fun f => + let f' := involutionCons φs.length f + let c' := fromInvolutionAux f'.1 + let n' := Option.map (finCongr ((uncontractedFromInvolution f'.fst).2.symm)) + (involutionAddEquiv f'.1 f'.2) + auxCongr (uncontractedFromInvolution_cons f).symm (ContractionsAux.cons n' c') + +def fromInvolution {φs : List 𝓕} (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) : + Contractions φs := ⟨uncontractedFromInvolution f, fromInvolutionAux f⟩ + +lemma fromInvolution_cons {φs : List 𝓕} {φ : 𝓕} + (f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) : + let f' := involutionCons φs.length f + fromInvolution f = consEquiv.symm + ⟨fromInvolution f'.1, Option.map (finCongr ((uncontractedFromInvolution f'.fst).2.symm)) + (involutionAddEquiv f'.1 f'.2)⟩ := by + refine auxCongr_ext ?_ ?_ + · dsimp [fromInvolution] + rw [uncontractedFromInvolution_cons] + rfl + · dsimp [fromInvolution, fromInvolutionAux] + rfl + +lemma fromInvolution_of_involutionCons + {φs : List 𝓕} {φ : 𝓕} + (f : {f : Fin (φs ).length → Fin (φs).length // Function.Involutive f}) + (n : { i : Option (Fin φs.length) // ∀ (h : i.isSome = true), f.1 (i.get h) = i.get h }): + fromInvolution (φs := φ :: φs) ((involutionCons φs.length).symm ⟨f, n⟩) = + consEquiv.symm + ⟨fromInvolution f, Option.map (finCongr ((uncontractedFromInvolution f).2.symm)) + (involutionAddEquiv f n)⟩ := by + rw [fromInvolution_cons] + congr 1 + simp + rw [Equiv.apply_symm_apply] + + +/-! + +## To Involution. + +-/ + +def toInvolution : {φs : List 𝓕} → (c : Contractions φs) → + {f : {f : Fin φs.length → Fin φs.length // Function.Involutive f} // + uncontractedFromInvolution f = c.1} + | [], ⟨[], ContractionsAux.nil⟩ => ⟨⟨fun i => i, by + intro i + simp⟩, by rfl⟩ + | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) n c⟩ => by + let ⟨⟨f', hf1⟩, hf2⟩ := toInvolution ⟨aux, c⟩ + let n' : Option (Fin (uncontractedFromInvolution ⟨f', hf1⟩).1.length) := + Option.map (finCongr (by rw [hf2])) n + let F := (involutionCons φs.length).symm ⟨⟨f', hf1⟩, + (involutionAddEquiv ⟨f', hf1⟩).symm + (Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ + refine ⟨F, ?_⟩ + have hF0 : ((involutionCons φs.length) F) = ⟨⟨f', hf1⟩, + (involutionAddEquiv ⟨f', hf1⟩).symm + (Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ := by + simp [F] + have hF1 : ((involutionCons φs.length) F).fst = ⟨f', hf1⟩ := by + rw [hF0] + have hF2L : ((uncontractedFromInvolution ⟨f', hf1⟩)).1.length = + (Finset.filter (fun i => ((involutionCons φs.length) F).1.1 i = i) Finset.univ).card := by + apply Eq.trans ((uncontractedFromInvolution ⟨f', hf1⟩)).2 + congr + rw [hF1] + have hF2 : ((involutionCons φs.length) F).snd = (involutionAddEquiv ((involutionCons φs.length) F).fst).symm + (Option.map (finCongr hF2L) n') := by + rw [@Sigma.subtype_ext_iff] at hF0 + ext1 + rw [hF0.2] + simp + congr 1 + · rw [hF1] + · refine involutionAddEquiv_cast' ?_ n' _ _ + rw [hF1] + rw [uncontractedFromInvolution_cons] + have hx := (toInvolution ⟨aux, c⟩).2 + simp at hx + simp + refine optionEraseZ_ext ?_ ?_ ?_ + · dsimp [F] + rw [Equiv.apply_symm_apply] + simp + rw [← hx] + simp_all only + · rfl + · simp [hF2] + dsimp [n'] + simp [finCongr] + simp only [Nat.succ_eq_add_one, id_eq, eq_mpr_eq_cast, F, n'] + ext a : 1 + simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans, + Fin.cast_eq_self, exists_eq_right] + +lemma toInvolution_length {φs φsᵤₙ : List 𝓕} {c : ContractionsAux φs φsᵤₙ} : + φsᵤₙ.length = (Finset.filter (fun i => (toInvolution ⟨φsᵤₙ, c⟩).1.1 i = i) Finset.univ).card + := by + have h2 := (toInvolution ⟨φsᵤₙ, c⟩).2 + simp at h2 + conv_lhs => rw [← h2] + exact Mathlib.Vector.length_val (uncontractedFromInvolution (toInvolution ⟨φsᵤₙ, c⟩).1) + +lemma toInvolution_cons {φs φsᵤₙ : List 𝓕} {φ : 𝓕} + (c : ContractionsAux φs φsᵤₙ) (n : Option (Fin (φsᵤₙ.length))) : + (toInvolution ⟨optionEraseZ φsᵤₙ φ n, ContractionsAux.cons n c⟩).1 + = (involutionCons φs.length).symm ⟨(toInvolution ⟨φsᵤₙ, c⟩).1, + (involutionAddEquiv (toInvolution ⟨φsᵤₙ, c⟩).1).symm + (Option.map (finCongr (toInvolution_length)) n)⟩ := by + dsimp [toInvolution] + congr 3 + rw [Option.map_map] + simp [finCongr] + rfl + +lemma toInvolution_consEquiv {φs : List 𝓕} {φ : 𝓕} + (c : Contractions φs) (n : Option (Fin (c.uncontracted.length))) : + (toInvolution ((consEquiv (φ := φ)).symm ⟨c, n⟩)).1 = + (involutionCons φs.length).symm ⟨(toInvolution c).1, + (involutionAddEquiv (toInvolution c).1).symm + (Option.map (finCongr (toInvolution_length)) n)⟩ := by + erw [toInvolution_cons] + rfl + +/-! + +## Involution equiv. + +-/ + +lemma toInvolution_fromInvolution : {φs : List 𝓕} → (c : Contractions φs) → + fromInvolution (toInvolution c) = c + | [], ⟨[], ContractionsAux.nil⟩ => rfl + | φ :: φs, ⟨_, .cons (φsᵤₙ := φsᵤₙ) n c⟩ => by + rw [toInvolution_cons] + rw [fromInvolution_of_involutionCons] + rw [Equiv.symm_apply_eq] + dsimp [consEquiv] + refine consEquiv_ext ?_ ?_ + · exact toInvolution_fromInvolution ⟨φsᵤₙ, c⟩ + · simp [finCongr] + ext a : 1 + simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans, + Fin.cast_eq_self, exists_eq_right] + +lemma fromInvolution_toInvolution : {φs : List 𝓕} → (f : {f : Fin (φs ).length → Fin (φs).length // Function.Involutive f}) + → toInvolution (fromInvolution f) = f + | [], _ => by + ext x + exact Fin.elim0 x + | φ :: φs, f => by + rw [fromInvolution_cons] + rw [toInvolution_consEquiv] + erw [Equiv.symm_apply_eq] + have hx := fromInvolution_toInvolution ((involutionCons φs.length) f).fst + apply involutionCons_ext ?_ ?_ + · simp only [Nat.succ_eq_add_one, List.length_cons] + exact hx + · simp only [Nat.succ_eq_add_one, Option.map_map, List.length_cons] + rw [Equiv.symm_apply_eq] + conv_rhs => + lhs + rw [involutionAddEquiv_cast hx] + simp [Nat.succ_eq_add_one,- eq_mpr_eq_cast, Equiv.trans_apply, -Equiv.optionCongr_apply] + rfl + +def equivInvolutions {φs : List 𝓕} : + Contractions φs ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f} where + toFun := fun c => toInvolution c + invFun := fromInvolution + left_inv := toInvolution_fromInvolution + right_inv := fromInvolution_toInvolution + + +/-! + +## Full contractions and involutions. +-/ +lemma isFull_iff_uncontractedFromInvolution_empty {φs : List 𝓕} (c : Contractions φs) : + IsFull c ↔ (uncontractedFromInvolution (equivInvolutions c)).1 = [] := by + let l := toInvolution c + erw [l.2] + rfl + +lemma isFull_iff_filter_card_involution_zero {φs : List 𝓕} (c : Contractions φs) : + IsFull c ↔ (Finset.univ.filter fun i => (equivInvolutions c).1 i = i).card = 0 := by + rw [isFull_iff_uncontractedFromInvolution_empty, List.ext_get_iff] + simp + +lemma isFull_iff_involution_no_fixed_points {φs : List 𝓕} (c : Contractions φs) : + IsFull c ↔ ∀ (i : Fin φs.length), (equivInvolutions c).1 i ≠ i := by + rw [isFull_iff_filter_card_involution_zero] + simp + rw [Finset.filter_eq_empty_iff] + apply Iff.intro + · intro h + intro i + refine h (Finset.mem_univ i) + · intro i h + exact fun a => i h + + +open Nat in +def isFullInvolutionEquiv {φs : List 𝓕} : + {c : Contractions φs // IsFull c} ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f ∧ (∀ i, f i ≠ i)} where + toFun c := ⟨equivInvolutions c.1, by + apply And.intro (equivInvolutions c.1).2 + rw [← isFull_iff_involution_no_fixed_points] + exact c.2 + ⟩ + invFun f := ⟨equivInvolutions.symm ⟨f.1, f.2.1⟩, by + rw [isFull_iff_involution_no_fixed_points] + simpa using f.2.2⟩ + left_inv c := by simp + right_inv f := by simp -informal_definition equivFullInvolution where - math :≈ "There is an isomorphism from the type of full contractions of a list `l` - and the type of fixed-point free involutions from `Fin l.length` to `Fin l.length." end Contractions diff --git a/HepLean/PerturbationTheory/Wick/StaticTheorem.lean b/HepLean/PerturbationTheory/Wick/StaticTheorem.lean index f94747ee..16bad3ad 100644 --- a/HepLean/PerturbationTheory/Wick/StaticTheorem.lean +++ b/HepLean/PerturbationTheory/Wick/StaticTheorem.lean @@ -25,11 +25,11 @@ lemma static_wick_nil {A : Type} [Semiring A] [Algebra ℂ A] (S : Contractions.Splitting f le) : F (ofListLift f [] 1) = ∑ c : Contractions [], c.toCenterTerm f q le F S * - F (koszulOrder (fun i => q i.fst) le (ofListLift f c.normalize 1)) := by + F (koszulOrder (fun i => q i.fst) le (ofListLift f c.uncontracted 1)) := by rw [← Contractions.nilEquiv.symm.sum_comp] simp only [Finset.univ_unique, PUnit.default_eq_unit, Contractions.nilEquiv, Equiv.coe_fn_symm_mk, Finset.sum_const, Finset.card_singleton, one_smul] - dsimp [Contractions.normalize, Contractions.toCenterTerm] + dsimp [Contractions.uncontracted, Contractions.toCenterTerm] simp [ofListLift_empty] lemma static_wick_cons [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × f i) le] @@ -38,10 +38,10 @@ lemma static_wick_cons [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × (S : Contractions.Splitting f le) (ih : F (ofListLift f φs 1) = ∑ c : Contractions φs, c.toCenterTerm f q le F S * F (koszulOrder (fun i => q i.fst) le - (ofListLift f c.normalize 1))) : + (ofListLift f c.uncontracted 1))) : F (ofListLift f (φ :: φs) 1) = ∑ c : Contractions (φ :: φs), c.toCenterTerm f q le F S * - F (koszulOrder (fun i => q i.fst) le (ofListLift f c.normalize 1)) := by + F (koszulOrder (fun i => q i.fst) le (ofListLift f c.uncontracted 1)) := by rw [ofListLift_cons_eq_ofListLift, map_mul, ih, Finset.mul_sum, ← Contractions.consEquiv.symm.sum_comp] erw [Finset.sum_sigma] @@ -88,7 +88,7 @@ theorem static_wick_theorem [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕 (F : FreeAlgebra ℂ (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F] (S : Contractions.Splitting f le) : F (ofListLift f φs 1) = ∑ c : Contractions φs, c.toCenterTerm f q le F S * - F (koszulOrder (fun i => q i.fst) le (ofListLift f c.normalize 1)) := by + F (koszulOrder (fun i => q i.fst) le (ofListLift f c.uncontracted 1)) := by induction φs with | nil => exact static_wick_nil q le F S | cons a r ih => exact static_wick_cons q le r a F S ih From 9184f6087c21552342251c8efd13382ff080cc76 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Sun, 5 Jan 2025 16:00:30 +0000 Subject: [PATCH 5/8] feat: Update free simp --- HepLean/Meta/AllFilePaths.lean | 8 +++--- scripts/MetaPrograms/free_simps.lean | 43 ++++++++++++++++++++++++---- 2 files changed, 42 insertions(+), 9 deletions(-) diff --git a/HepLean/Meta/AllFilePaths.lean b/HepLean/Meta/AllFilePaths.lean index afd869ec..1f19d822 100644 --- a/HepLean/Meta/AllFilePaths.lean +++ b/HepLean/Meta/AllFilePaths.lean @@ -16,15 +16,15 @@ open Lean Elab System /-- The recursive function underlying `allFilePaths`. -/ partial def allFilePaths.go (prev : Array FilePath) - (root : Name) (path : FilePath) : IO (Array FilePath) := do + (root : String) (path : FilePath) : IO (Array FilePath) := do let mut r := prev for entry in ← path.readDir do if ← entry.path.isDir then - r ← go r (root.mkStr entry.fileName) entry.path + r ← go r (root ++ "/" ++ entry.fileName) entry.path else - r := r.push (root.toString.replace "." "/" ++ "/" ++ entry.fileName) + r := r.push (root ++ "/" ++ entry.fileName) pure r /-- Gets an array of all file paths in `HepLean`. -/ partial def allFilePaths : IO (Array FilePath) := do - allFilePaths.go (#[] : Array FilePath) `HepLean ("./HepLean" : FilePath) + allFilePaths.go (#[] : Array FilePath) "./HepLean" ("./HepLean" : FilePath) diff --git a/scripts/MetaPrograms/free_simps.lean b/scripts/MetaPrograms/free_simps.lean index 9e32c9d4..791b68a4 100644 --- a/scripts/MetaPrograms/free_simps.lean +++ b/scripts/MetaPrograms/free_simps.lean @@ -71,9 +71,42 @@ unsafe def processAllFiles : IO Unit := do println! x | .error _ => println! "Error" -unsafe def main (args : List String) : IO Unit := do +unsafe def processFileArray (files : Array FilePath) : IO Unit := do + if files.toList.length = 0 then + return () + if files.toList.length = 1 then + let path? := files.get? 0 + match path? with + | some path => + transverseTactics path visitTacticInfo + return () + | none => + return () + let tasks := files.map fun f => + ((IO.asTask $ IO.Process.run + {cmd := "lake", args := #["exe", "free_simps","-file", f.toString]}), f) + tasks.toList.enum.forM fun (n, (t, path)) => do + let tn ← IO.wait (← t) + match tn with + | .ok x => + if x ≠ "" then + println! "{n} of {tasks.toList.length}: {path}" + println! x + | .error _ => println! "Error" + IO.println "Finished!" + +unsafe def argToArrayFilePath (args : List String) : IO (Array FilePath) := do match args with - | [path] => transverseTactics path visitTacticInfo - | _ => - processAllFiles - IO.println "Finished" + | ["-file", path] => return #[path] + | ["-dir", dir] => do + let files ← allFilePaths.go (#[] : Array FilePath) dir (dir : FilePath) + return files + | [] => do + let files ← allFilePaths + return files + | _ => do + panic! "Invalid arguments" + +unsafe def main (args : List String) : IO Unit := do + let files ← argToArrayFilePath args + processFileArray files From 7026d8ce6639c59157c5d419aa4d8cc4e5fa07c8 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Sun, 5 Jan 2025 16:46:15 +0000 Subject: [PATCH 6/8] refactor: free simps --- HepLean/Mathematics/Fin.lean | 12 +- HepLean/Mathematics/Fin/Involutions.lean | 132 +++++++++--------- HepLean/Mathematics/List.lean | 2 +- .../Contractions/Basic.lean | 18 +-- .../PerturbationTheory/Contractions/Card.lean | 2 +- .../Contractions/Involutions.lean | 128 ++++++++--------- scripts/MetaPrograms/free_simps.lean | 2 + 7 files changed, 150 insertions(+), 146 deletions(-) diff --git a/HepLean/Mathematics/Fin.lean b/HepLean/Mathematics/Fin.lean index 166dd8f1..d1427084 100644 --- a/HepLean/Mathematics/Fin.lean +++ b/HepLean/Mathematics/Fin.lean @@ -101,7 +101,7 @@ lemma succAbove_succAbove_predAboveI (i : Fin n.succ.succ) (j : Fin n.succ) (x : · rw [Fin.succAbove_of_castSucc_lt _ _] exact hx1 · rw [Fin.lt_def] at h1 hx1 ⊢ - simp_all + simp_all only [Nat.succ_eq_add_one, Fin.coe_castSucc] omega · exact Nat.lt_trans hx1 h1 · simp only [not_lt] at hx1 @@ -115,7 +115,7 @@ lemma succAbove_succAbove_predAboveI (i : Fin n.succ.succ) (j : Fin n.succ) (x : · rfl · exact hx1 · rw [Fin.lt_def] at hx2 ⊢ - simp_all + simp_all only [Nat.succ_eq_add_one, Fin.coe_castSucc, Fin.val_succ] omega · simp only [Nat.succ_eq_add_one, not_lt] at hx2 rw [Fin.succAbove_of_le_castSucc _ _ hx2] @@ -135,12 +135,12 @@ lemma succAbove_succAbove_predAboveI (i : Fin n.succ.succ) (j : Fin n.succ) (x : · nth_rewrite 2 [Fin.succAbove_of_le_castSucc _ _] · rw [Fin.succAbove_of_le_castSucc _ _] rw [Fin.le_def] at hx1 ⊢ - simp_all + simp_all only [Nat.succ_eq_add_one, Fin.coe_castSucc, Fin.val_succ, add_le_add_iff_right] · rw [Fin.le_def] at h1 hx1 ⊢ - simp_all + simp_all only [Nat.succ_eq_add_one, Fin.coe_castSucc] omega · rw [Fin.le_def] at hx1 h1 ⊢ - simp_all + simp_all only [Nat.succ_eq_add_one, Fin.coe_castSucc, Fin.val_succ] omega · simp only [Nat.succ_eq_add_one, not_le] at hx1 rw [Fin.lt_def] at hx1 @@ -151,7 +151,7 @@ lemma succAbove_succAbove_predAboveI (i : Fin n.succ.succ) (j : Fin n.succ) (x : nth_rewrite 2 [Fin.succAbove_of_castSucc_lt _ _] · rw [Fin.succAbove_of_castSucc_lt _ _] rw [Fin.lt_def] at hx2 ⊢ - simp_all + simp_all only [Nat.succ_eq_add_one, Fin.coe_castSucc, Fin.val_succ] omega · rw [Fin.lt_def] at hx2 ⊢ simp_all diff --git a/HepLean/Mathematics/Fin/Involutions.lean b/HepLean/Mathematics/Fin/Involutions.lean index 606f4610..8baed460 100644 --- a/HepLean/Mathematics/Fin/Involutions.lean +++ b/HepLean/Mathematics/Fin/Involutions.lean @@ -29,14 +29,15 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu intro i by_cases h : f.1 i.succ = 0 · simp [h] - · simp [h] - simp [f.2 i.succ] + · simp only [succ_eq_add_one, h, ↓reduceDIte, Fin.succ_pred] + simp only [f.2 i.succ, Fin.pred_succ, dite_eq_ite, ite_eq_right_iff] intro h exact False.elim (Fin.succ_ne_zero i h)⟩, ⟨if h : f.1 0 = 0 then none else Fin.pred (f.1 0) h , by by_cases h0 : f.1 0 = 0 · simp [h0] - · simp [h0] + · simp only [succ_eq_add_one, h0, ↓reduceDIte, Option.isSome_some, Option.get_some, + Fin.succ_pred, dite_eq_left_iff, Fin.pred_inj, forall_const] refine fun h => False.elim (h (f.2 0))⟩⟩ invFun f := ⟨ if h : (f.2.1).isSome then @@ -60,7 +61,7 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu simp · rw [Function.update_apply ] rw [if_neg hja] - simp + simp only [Function.comp_apply, Fin.cons_succ] have hf2 := f.2.2 hs change f.1.1 a = a at hf2 have hjf1 : f.1.1 j ≠ a := by @@ -71,40 +72,41 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu rw [hf2] at haj exact hja haj rw [Function.update_apply, if_neg hjf1] - simp + simp only [Function.comp_apply, Fin.succ_inj] rw [f.1.2] - · simp [hs] + · simp only [succ_eq_add_one, hs, Bool.false_eq_true, ↓reduceDIte] intro i rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ · subst hi simp · subst hj - simp + simp only [Fin.cons_succ, Function.comp_apply, Fin.succ_inj] rw [f.1.2]⟩ left_inv f := by match f with | ⟨f, hf⟩ => - simp + simp only [succ_eq_add_one, Option.isSome_dite', Option.get_dite', Fin.succ_pred, + Fin.cons_update, dite_eq_ite, ite_not, Subtype.mk.injEq] ext i by_cases h0 : f 0 = 0 - · simp [h0] + · simp only [h0, ↓reduceIte] rcases Fin.eq_zero_or_eq_succ i with hi | ⟨j, hj⟩ · subst hi simp [h0] · subst hj - simp [h0] + simp only [Fin.cons_succ, Function.comp_apply, Fin.val_succ] by_cases hj : f j.succ =0 · rw [← h0] at hj have hn := Function.Involutive.injective hf hj exact False.elim (Fin.succ_ne_zero j hn) - · simp [hj] + · simp only [hj, ↓reduceDIte, Fin.coe_pred] rw [Fin.ext_iff] at hj - simp at hj + simp only [succ_eq_add_one, Fin.val_zero] at hj omega · rw [if_neg h0] by_cases hf' : i = f 0 · subst hf' - simp + simp only [Function.update_same, Fin.val_zero] rw [hf] simp · rw [Function.update_apply, if_neg hf'] @@ -112,31 +114,33 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu · subst hi simp · subst hj - simp + simp only [Fin.cons_succ, Function.comp_apply, Fin.val_succ] by_cases hj : f j.succ =0 · rw [← hj] at hf' rw [hf] at hf' - simp at hf' - · simp [hj] + simp only [not_true_eq_false] at hf' + · simp only [hj, ↓reduceDIte, Fin.coe_pred] rw [Fin.ext_iff] at hj - simp at hj + simp only [succ_eq_add_one, Fin.val_zero] at hj omega right_inv f := by match f with | ⟨⟨f, hf⟩, ⟨f0, hf0⟩⟩ => ext i - · simp + · simp only [succ_eq_add_one, Fin.cons_update] by_cases hs : f0.isSome - · simp [hs] + · simp only [hs, ↓reduceDIte] by_cases hi : i = f0.get hs - · simp [hi, Function.update_apply] + · simp only [Function.update_apply, hi, ↓reduceIte, ↓reduceDIte] exact Eq.symm (Fin.val_eq_of_eq (hf0 hs)) - · simp [hi] + · simp only [ne_eq, Fin.succ_inj, hi, not_false_eq_true, Function.update_noteq, + Fin.cons_succ, Function.comp_apply, Fin.pred_succ, dite_eq_ite] split · rename_i h exact False.elim (Fin.succ_ne_zero (f i) h) · rfl - · simp [hs] + · simp only [hs, Bool.false_eq_true, ↓reduceDIte, Fin.cons_succ, Function.comp_apply, + Fin.pred_succ, dite_eq_ite] split · rename_i h exact False.elim (Fin.succ_ne_zero (f i) h) @@ -145,9 +149,9 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu Option.dite_none_left_eq_some, Option.some.injEq] by_cases hs : f0.isSome · simp only [hs, ↓reduceDIte] - simp [Fin.cons_zero] + simp only [Fin.cons_zero, Fin.pred_succ, exists_prop] have hx : ¬ (f0.get hs).succ = 0 := (Fin.succ_ne_zero (f0.get hs)) - simp [hx] + simp only [hx, not_false_eq_true, true_and] refine Iff.intro (fun hi => ?_) (fun hi => ?_) · rw [← hi] exact @@ -156,8 +160,9 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu (of_eq_true (Eq.trans (congrArg (fun x => x = true) hs) (eq_self true)))) · subst hi exact rfl - · simp [hs] - simp at hs + · simp only [hs, Bool.false_eq_true, ↓reduceDIte, Fin.cons_zero, not_true_eq_false, + IsEmpty.exists_iff, false_iff] + simp only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none] at hs subst hs exact ne_of_beq_false rfl @@ -166,7 +171,7 @@ lemma involutionCons_ext {n : ℕ} {f1 f2 : (f : {f : Fin n → Fin n // Functi (h1 : f1.1 = f2.1) (h2 : f1.2 = Equiv.subtypeEquivRight (by rw [h1]; simp) f2.2) : f1 = f2 := by cases f1 cases f2 - simp at h1 h2 + simp only at h1 h2 subst h1 rename_i fst snd snd_1 simp_all only [Sigma.mk.inj_iff, heq_eq_eq, true_and] @@ -234,7 +239,7 @@ lemma involutionAddEquiv_cast {n : ℕ} {f1 f2 : {f : Fin n → Fin n // Functio involutionAddEquiv f1 = (Equiv.subtypeEquivRight (by rw [hf]; simp)).trans ((involutionAddEquiv f2).trans (Equiv.optionCongr (finCongr (by rw [hf])))):= by subst hf - simp + simp only [finCongr_refl, Equiv.optionCongr_refl, Equiv.trans_refl] rfl @@ -249,7 +254,7 @@ lemma involutionAddEquiv_none_succ {n : ℕ} {f : {f : Fin n.succ → Fin n.succ // Function.Involutive f}} (h : involutionAddEquiv (involutionCons n f).1 (involutionCons n f).2 = none) (x : Fin n) : f.1 x.succ = x.succ ↔ (involutionCons n f).1.1 x = x := by - simp [involutionCons] + simp only [succ_eq_add_one, involutionCons, Fin.cons_update, Equiv.coe_fn_mk, dite_eq_left_iff] have hn' := involutionAddEquiv_none_image_zero h have hx : ¬ f.1 x.succ = ⟨0, Nat.zero_lt_succ n⟩:= by rw [← hn'] @@ -268,7 +273,9 @@ lemma involutionAddEquiv_isSome_image_zero {n : ℕ} : → (involutionAddEquiv (involutionCons n f).1 (involutionCons n f).2).isSome → ¬ f.1 ⟨0, Nat.zero_lt_succ n⟩ = ⟨0, Nat.zero_lt_succ n⟩ := by intro f hf - simp [involutionAddEquiv, involutionCons] at hf + simp only [succ_eq_add_one, involutionCons, Equiv.coe_fn_mk, involutionAddEquiv, + Option.isSome_some, Option.get_some, Option.isSome_none, Equiv.trans_apply, + Equiv.optionCongr_apply, Equiv.coe_trans, RelIso.coe_fn_toEquiv, Option.isSome_map'] at hf simp_all only [List.length_cons, Fin.zero_eta] obtain ⟨val, property⟩ := f simp_all only [List.length_cons] @@ -286,9 +293,9 @@ def involutionNoFixedEquivSum {n : ℕ} : left_inv f := by rfl right_inv f := by - simp + simp only [succ_eq_add_one, ne_eq, mul_eq] ext - · simp + · simp only [Fin.coe_pred] rw [f.2.2.2.2] simp · simp @@ -309,12 +316,12 @@ def involutionNoFixedZeroSetEquivEquiv {n : ℕ} (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} where toFun f := ⟨e ∘ f.1 ∘ e.symm, by intro i - simp + simp only [succ_eq_add_one, ne_eq, Function.comp_apply, Equiv.symm_apply_apply] rw [f.2.1], by simpa using f.2.2.1, by simpa using f.2.2.2⟩ invFun f := ⟨e.symm ∘ f.1 ∘ e, by intro i - simp + simp only [succ_eq_add_one, Function.comp_apply, ne_eq, Equiv.apply_symm_apply] have hf2 := f.2.1 i simpa using hf2, by simpa using f.2.2.1, by simpa using f.2.2.2⟩ @@ -341,7 +348,7 @@ def involutionNoFixedZeroSetEquivSetEquiv {n : ℕ} (k : Fin (2 * n + 1)) (e : F have hi := h (e i) simp [hi] rw [h1] - simp + simp only [succ_eq_add_one, Function.comp_apply, ne_eq, and_congr_right_iff, and_congr_left_iff] intro h1 h2 apply Iff.intro · intro h i @@ -360,7 +367,7 @@ def involutionNoFixedZeroSetEquivEquiv' {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i) ∧ f (e 0) = e k.succ} := by refine Equiv.subtypeEquivRight ?_ - simp + simp only [succ_eq_add_one, ne_eq, Function.comp_apply, and_congr_right_iff] intro f hi h1 exact Equiv.symm_apply_eq e @@ -373,7 +380,7 @@ def involutionNoFixedZeroSetEquivSetOne {n : ℕ} (k : Fin (2 * n + 1)) : refine Equiv.trans (involutionNoFixedZeroSetEquivSetEquiv k (Equiv.swap k.succ 1)) ?_ refine Equiv.trans (involutionNoFixedZeroSetEquivEquiv' k (Equiv.swap k.succ 1)) ?_ refine Equiv.subtypeEquivRight ?_ - simp + simp only [succ_eq_add_one, ne_eq, Equiv.swap_apply_left, and_congr_right_iff] intro f hi h1 rw [Equiv.swap_apply_of_ne_of_ne] · exact Ne.symm (Fin.succ_ne_zero k) @@ -387,33 +394,32 @@ def involutionNoFixedSetOne {n : ℕ} : toFun f := by have hf1 : f.1 1 = 0 := by have hf := f.2.2.2 - simp [← hf] + simp only [succ_eq_add_one, ne_eq, ← hf] rw [f.2.1] let f' := f.1 ∘ Fin.succ ∘ Fin.succ have hf' (i : Fin (2 * n)) : f' i ≠ 0 := by - simp [f'] - simp [← hf1] + simp only [succ_eq_add_one, mul_eq, ne_eq, Function.comp_apply, f'] + simp only [← hf1, succ_eq_add_one, ne_eq] by_contra hn have hn' := Function.Involutive.injective f.2.1 hn simp [Fin.ext_iff] at hn' let f'' := fun i => (f' i).pred (hf' i) have hf'' (i : Fin (2 * n)) : f'' i ≠ 0 := by - simp [f''] + simp only [mul_eq, ne_eq, f''] rw [@Fin.pred_eq_iff_eq_succ] - simp [f'] - simp [← f.2.2.2 ] + simp only [mul_eq, succ_eq_add_one, ne_eq, Function.comp_apply, Fin.succ_zero_eq_one, f'] + simp only [← f.2.2.2, succ_eq_add_one, ne_eq] by_contra hn have hn' := Function.Involutive.injective f.2.1 hn simp [Fin.ext_iff] at hn' let f''' := fun i => (f'' i).pred (hf'' i) refine ⟨f''', ?_, ?_⟩ · intro i - simp [f''', f'', f'] + simp only [mul_eq, succ_eq_add_one, ne_eq, Function.comp_apply, Fin.succ_pred, f''', f'', f'] simp [f.2.1 i.succ.succ] · intro i - simp [f''', f'', f'] - rw [@Fin.pred_eq_iff_eq_succ] - rw [@Fin.pred_eq_iff_eq_succ] + simp only [mul_eq, succ_eq_add_one, ne_eq, Function.comp_apply, f''', f'', f'] + rw [Fin.pred_eq_iff_eq_succ, Fin.pred_eq_iff_eq_succ] exact f.2.2.1 i.succ.succ invFun f := by let f' := fun (i : Fin (2 * n.succ))=> @@ -427,16 +433,16 @@ def involutionNoFixedSetOne {n : ℕ} : | ⟨0, h⟩ => rfl | ⟨1, h⟩ => rfl | ⟨(Nat.succ (Nat.succ m)), h⟩ => - simp [f'] + simp only [succ_eq_add_one, ne_eq, f'] split · rename_i h - simp at h + simp only [succ_eq_add_one, Fin.zero_eta] at h exact False.elim (Fin.succ_ne_zero (f.1 ⟨m, _⟩).succ h) · rename_i h simp [Fin.ext_iff] at h · rename_i h rename_i x r - simp_all [Fin.ext_iff] + simp_all only [succ_eq_add_one, Fin.ext_iff, Fin.val_succ, add_left_inj] have hfn {a b : ℕ} {ha : a < 2 * n} {hb : b < 2 * n} (hab : ↑(f.1 ⟨a, ha⟩) = b): ↑(f.1 ⟨b, hb⟩) = a := by have ht : f.1 ⟨a, ha⟩ = ⟨b, hb⟩ := by @@ -446,7 +452,7 @@ def involutionNoFixedSetOne {n : ℕ} : · intro i match i with | ⟨0, h⟩ => - simp [f'] + simp only [succ_eq_add_one, ne_eq, Fin.zero_eta, f'] split · rename_i h simp @@ -455,7 +461,7 @@ def involutionNoFixedSetOne {n : ℕ} : · rename_i h simp [Fin.ext_iff] at h | ⟨1, h⟩ => - simp [f'] + simp only [succ_eq_add_one, ne_eq, Fin.mk_one, f'] split · rename_i h simp at h @@ -464,11 +470,11 @@ def involutionNoFixedSetOne {n : ℕ} : · rename_i h simp [Fin.ext_iff] at h | ⟨(Nat.succ (Nat.succ m)), h⟩ => - simp [f', Fin.ext_iff] + simp only [succ_eq_add_one, ne_eq, Fin.ext_iff, Fin.val_succ, add_left_inj, f'] have hf:= f.2.2 ⟨m, by exact Nat.add_lt_add_iff_right.mp h⟩ - simp [Fin.ext_iff] at hf + simp only [ne_eq, Fin.ext_iff] at hf omega - · simp [f'] + · simp only [succ_eq_add_one, ne_eq, f'] split · rename_i h simp @@ -479,30 +485,30 @@ def involutionNoFixedSetOne {n : ℕ} : left_inv f := by have hf1 : f.1 1 = 0 := by have hf := f.2.2.2 - simp [← hf] + simp only [succ_eq_add_one, ne_eq, ← hf] rw [f.2.1] - simp + simp only [succ_eq_add_one, ne_eq, mul_eq, Function.comp_apply, Fin.succ_mk, Fin.succ_pred] ext i - simp + simp only split - · simp + · simp only [Fin.val_one, succ_eq_add_one, Fin.zero_eta] rw [f.2.2.2] simp - · simp + · simp only [Fin.val_zero, succ_eq_add_one, Fin.mk_one] rw [hf1] simp · rfl right_inv f := by - simp + simp only [ne_eq, mul_eq, succ_eq_add_one, Function.comp_apply] ext i - simp + simp only [Fin.coe_pred] split · rename_i h simp [Fin.ext_iff] at h · rename_i h simp [Fin.ext_iff] at h · rename_i h - simp + simp only [Fin.val_succ, add_tsub_cancel_right] congr apply congrArg simp_all [Fin.ext_iff] diff --git a/HepLean/Mathematics/List.lean b/HepLean/Mathematics/List.lean index 251e5e57..ef5ed81b 100644 --- a/HepLean/Mathematics/List.lean +++ b/HepLean/Mathematics/List.lean @@ -45,7 +45,7 @@ lemma takeWile_eraseIdx {I : Type} (P : I → Prop) [DecidablePred P] : | a :: b :: l, Nat.succ n, h => by simp only [Nat.succ_eq_add_one, List.eraseIdx_cons_succ] by_cases hPa : P a - · dsimp [List.takeWhile] + · dsimp only [List.takeWhile] simp only [hPa, decide_True, List.eraseIdx_cons_succ, List.cons.injEq, true_and] rw [takeWile_eraseIdx] rfl diff --git a/HepLean/PerturbationTheory/Contractions/Basic.lean b/HepLean/PerturbationTheory/Contractions/Basic.lean index c57e1d1a..f4af2b31 100644 --- a/HepLean/PerturbationTheory/Contractions/Basic.lean +++ b/HepLean/PerturbationTheory/Contractions/Basic.lean @@ -41,9 +41,9 @@ lemma auxCongr_ext {φs: List 𝓕} {c c2 : Contractions φs} (h : c.1 = c2.1) (hx : c.2 = auxCongr h.symm c2.2) : c = c2 := by cases c cases c2 - simp at h + simp only at h subst h - simp [auxCongr] at hx + simp only [auxCongr, Equiv.refl_apply] at hx subst hx rfl @@ -97,11 +97,11 @@ lemma embedUncontracted_injective {l : List 𝓕} (c : Contractions l) : Function.Injective c.embedUncontracted := by match l, c with | [], ⟨[], ContractionsAux.nil⟩ => - dsimp [embedUncontracted] + dsimp only [List.length_nil, embedUncontracted] intro i exact Fin.elim0 i | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => - dsimp [embedUncontracted] + dsimp only [List.length_cons, embedUncontracted, Fin.zero_eta] refine Fin.cons_injective_iff.mpr ?_ apply And.intro · simp only [Set.mem_range, Function.comp_apply, not_exists] @@ -109,7 +109,7 @@ lemma embedUncontracted_injective {l : List 𝓕} (c : Contractions l) : · exact Function.Injective.comp (Fin.succ_injective φs.length) (embedUncontracted_injective ⟨aux, c⟩) | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some i) c⟩ => - dsimp [embedUncontracted] + dsimp only [List.length_cons, embedUncontracted] refine Function.Injective.comp (Fin.succ_injective φs.length) ?hf refine Function.Injective.comp (embedUncontracted_injective ⟨aux, c⟩) ?hf.hf refine Function.Injective.comp (Fin.cast_injective (embedUncontracted.proof_5 φ φs aux i c)) @@ -251,7 +251,7 @@ lemma toCenterTerm_none (f : 𝓕 → Type) [∀ i, Fintype (f i)] toCenterTerm f q le F c S := by rw [consEquiv] simp only [Equiv.coe_fn_symm_mk] - dsimp [toCenterTerm] + dsimp only [toCenterTerm] rfl /-- Proves that the part of the term gained from Wick contractions is in @@ -264,13 +264,13 @@ lemma toCenterTerm_center (f : 𝓕 → Type) [∀ i, Fintype (f i)] {φs : List 𝓕} → (c : Contractions φs) → (S : Splitting f le) → (c.toCenterTerm f q le F S) ∈ Subalgebra.center ℂ A | [], ⟨[], .nil⟩, _ => by - dsimp [toCenterTerm] + dsimp only [toCenterTerm] exact Subalgebra.one_mem (Subalgebra.center ℂ A) | _ :: _, ⟨_, .cons (φsᵤₙ := aux') none c⟩, S => by - dsimp [toCenterTerm] + dsimp only [toCenterTerm] exact toCenterTerm_center f q le F ⟨aux', c⟩ S | a :: _, ⟨_, .cons (φsᵤₙ := aux') (some n) c⟩, S => by - dsimp [toCenterTerm] + dsimp only [toCenterTerm, List.get_eq_getElem] refine Subalgebra.mul_mem (Subalgebra.center ℂ A) ?hx ?hy exact toCenterTerm_center f q le F ⟨aux', c⟩ S apply Subalgebra.smul_mem diff --git a/HepLean/PerturbationTheory/Contractions/Card.lean b/HepLean/PerturbationTheory/Contractions/Card.lean index bee0d2fb..0edb7f97 100644 --- a/HepLean/PerturbationTheory/Contractions/Card.lean +++ b/HepLean/PerturbationTheory/Contractions/Card.lean @@ -31,7 +31,7 @@ lemma card_of_full_contractions_odd {φs : List 𝓕} (ho : Odd φs.length ) : by_contra hn have hc := uncontracted_length_even_iff c rw [hn] at hc - simp at hc + simp only [List.length_nil, even_zero, iff_true] at hc rw [← Nat.not_odd_iff_even] at hc exact hc ho diff --git a/HepLean/PerturbationTheory/Contractions/Involutions.lean b/HepLean/PerturbationTheory/Contractions/Involutions.lean index c5d67587..d624ad12 100644 --- a/HepLean/PerturbationTheory/Contractions/Involutions.lean +++ b/HepLean/PerturbationTheory/Contractions/Involutions.lean @@ -36,6 +36,8 @@ variable {l : List 𝓕} -/ +/-- Given an involution the uncontracted fields associated with it (corresponding + to the fixed points of that involution). -/ def uncontractedFromInvolution : {φs : List 𝓕} → (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) → {l : List 𝓕 // l.length = (Finset.univ.filter fun i => f.1 i = i).card} @@ -43,10 +45,11 @@ def uncontractedFromInvolution : {φs : List 𝓕} → | φ :: φs, f => let luc := uncontractedFromInvolution (involutionCons φs.length f).fst let n' := involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2 + let np : Option (Fin luc.1.length) := Option.map (finCongr luc.2.symm) n' if hn : n' = none then have hn' := involutionAddEquiv_none_image_zero (n := φs.length) (f := f) hn ⟨optionEraseZ luc φ none, by - simp [optionEraseZ] + simp only [optionEraseZ, Nat.succ_eq_add_one, List.length_cons, Mathlib.Vector.length_val] rw [← luc.2] conv_rhs => rw [Finset.card_filter] rw [Fin.sum_univ_succ] @@ -61,13 +64,13 @@ def uncontractedFromInvolution : {φs : List 𝓕} → rw [involutionAddEquiv_none_succ hn]⟩ else let n := n'.get (Option.isSome_iff_ne_none.mpr hn) - let np : Fin luc.1.length := ⟨n.1, by - rw [luc.2] - exact n.prop⟩ + let np : Fin luc.1.length := Fin.cast luc.2.symm n ⟨optionEraseZ luc φ (some np), by let k' := (involutionCons φs.length f).2 have hkIsSome : (k'.1).isSome := by - simp [n', involutionAddEquiv ] at hn + simp only [Nat.succ_eq_add_one, involutionAddEquiv, Option.isSome_some, Option.get_some, + Option.isSome_none, Equiv.trans_apply, Equiv.coe_fn_mk, Equiv.optionCongr_apply, + Equiv.coe_trans, RelIso.coe_fn_toEquiv, Option.map_eq_none', n'] at hn split at hn · simp_all only [reduceCtorEq, not_false_eq_true, Nat.succ_eq_add_one, Option.isSome_some, k'] · simp_all only [not_true_eq_false] @@ -76,39 +79,34 @@ def uncontractedFromInvolution : {φs : List 𝓕} → have hksucc : k.succ = f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ := by simp [k, k', involutionCons] have hzero : ⟨0, Nat.zero_lt_succ φs.length⟩ = f.1 k.succ := by - rw [hksucc] - rw [f.2] - have hkcons : ((involutionCons φs.length) f).1.1 k = k := by - exact k'.2 hkIsSome + rw [hksucc, f.2] have hksuccNe : f.1 k.succ ≠ k.succ := by conv_rhs => rw [hksucc] exact fun hn => Fin.succ_ne_zero k (Function.Involutive.injective f.2 hn ) have hluc : 1 ≤ luc.1.length := by - simp + simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, Finset.one_le_card] use k - simp [involutionCons] + simp only [involutionCons, Nat.succ_eq_add_one, Fin.cons_update, Equiv.coe_fn_mk, + dite_eq_left_iff, Finset.mem_filter, Finset.mem_univ, true_and] rw [hksucc, f.2] simp rw [propext (Nat.sub_eq_iff_eq_add' hluc)] - have h0 : ¬ f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ = ⟨0, Nat.zero_lt_succ φs.length⟩ := by - exact Option.isSome_dite'.mp hkIsSome conv_rhs => rw [Finset.card_filter] - erw [Fin.sum_univ_succ] - erw [if_neg h0] + erw [Fin.sum_univ_succ, if_neg (Option.isSome_dite'.mp hkIsSome)] simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, List.length_cons, Nat.cast_id, zero_add] conv_rhs => lhs; rw [Eq.symm (Fintype.sum_ite_eq' k fun j => 1)] - rw [← Finset.sum_add_distrib] - rw [Finset.card_filter] + rw [← Finset.sum_add_distrib, Finset.card_filter] apply congrArg funext i by_cases hik : i = k · subst hik - simp [hkcons, hksuccNe] - · simp [hik] + simp only [k'.2 hkIsSome, Nat.succ_eq_add_one, ↓reduceIte, hksuccNe, add_zero] + · simp only [hik, ↓reduceIte, zero_add] refine ite_congr ?_ (congrFun rfl) (congrFun rfl) - simp [involutionCons] + simp only [involutionCons, Nat.succ_eq_add_one, Fin.cons_update, Equiv.coe_fn_mk, + dite_eq_left_iff, eq_iff_iff] have hfi : f.1 i.succ ≠ ⟨0, Nat.zero_lt_succ φs.length⟩ := by rw [hzero] by_contra hn @@ -121,7 +119,7 @@ def uncontractedFromInvolution : {φs : List 𝓕} → conv_rhs => rw [← h'] simp · intro h hfi - simp [Fin.ext_iff] + simp only [Fin.ext_iff, Fin.coe_pred] rw [h] simp⟩ @@ -135,18 +133,17 @@ lemma uncontractedFromInvolution_cons {φs : List 𝓕} {φ : 𝓕} let n' := involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2 change _ = optionEraseZ luc φ (Option.map (finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm)) n') - dsimp [uncontractedFromInvolution] + dsimp only [List.length_cons, uncontractedFromInvolution, Nat.succ_eq_add_one, Fin.zero_eta] by_cases hn : n' = none · have hn' := hn - simp [n'] at hn' - simp [hn'] - rw [hn] + simp only [Nat.succ_eq_add_one, n'] at hn' + simp only [hn', ↓reduceDIte, hn] rfl · have hn' := hn - simp [n'] at hn' - simp [hn'] + simp only [Nat.succ_eq_add_one, n'] at hn' + simp only [hn', ↓reduceDIte] congr - simp [n'] + simp only [Nat.succ_eq_add_one, n'] simp_all only [Nat.succ_eq_add_one, not_false_eq_true, n', luc] obtain ⟨val, property⟩ := f obtain ⟨val_1, property_1⟩ := luc @@ -166,10 +163,10 @@ lemma uncontractedFromInvolution_cons {φs : List 𝓕} {φ : 𝓕} obtain ⟨left, right⟩ := h subst right simp_all only [Option.get_some] - rfl +/-- The `ContractionsAux` associated to an involution. -/ def fromInvolutionAux : {l : List 𝓕} → - (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) → + (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) → ContractionsAux l (uncontractedFromInvolution f) | [] => fun _ => ContractionsAux.nil | _ :: φs => fun f => @@ -179,6 +176,7 @@ def fromInvolutionAux : {l : List 𝓕} → (involutionAddEquiv f'.1 f'.2) auxCongr (uncontractedFromInvolution_cons f).symm (ContractionsAux.cons n' c') +/-- The contraction associated with an involution. -/ def fromInvolution {φs : List 𝓕} (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) : Contractions φs := ⟨uncontractedFromInvolution f, fromInvolutionAux f⟩ @@ -189,32 +187,31 @@ lemma fromInvolution_cons {φs : List 𝓕} {φ : 𝓕} ⟨fromInvolution f'.1, Option.map (finCongr ((uncontractedFromInvolution f'.fst).2.symm)) (involutionAddEquiv f'.1 f'.2)⟩ := by refine auxCongr_ext ?_ ?_ - · dsimp [fromInvolution] + · dsimp only [fromInvolution, List.length_cons, Nat.succ_eq_add_one] rw [uncontractedFromInvolution_cons] rfl - · dsimp [fromInvolution, fromInvolutionAux] + · dsimp only [fromInvolution, List.length_cons, fromInvolutionAux, Nat.succ_eq_add_one, id_eq, + eq_mpr_eq_cast] rfl -lemma fromInvolution_of_involutionCons - {φs : List 𝓕} {φ : 𝓕} +lemma fromInvolution_of_involutionCons {φs : List 𝓕} {φ : 𝓕} (f : {f : Fin (φs ).length → Fin (φs).length // Function.Involutive f}) (n : { i : Option (Fin φs.length) // ∀ (h : i.isSome = true), f.1 (i.get h) = i.get h }): fromInvolution (φs := φ :: φs) ((involutionCons φs.length).symm ⟨f, n⟩) = - consEquiv.symm - ⟨fromInvolution f, Option.map (finCongr ((uncontractedFromInvolution f).2.symm)) + consEquiv.symm ⟨fromInvolution f, Option.map (finCongr ((uncontractedFromInvolution f).2.symm)) (involutionAddEquiv f n)⟩ := by rw [fromInvolution_cons] congr 1 - simp + simp only [Nat.succ_eq_add_one, Sigma.mk.inj_iff, Equiv.apply_symm_apply, true_and] rw [Equiv.apply_symm_apply] - /-! ## To Involution. -/ +/-- The involution associated with a contraction. -/ def toInvolution : {φs : List 𝓕} → (c : Contractions φs) → {f : {f : Fin φs.length → Fin φs.length // Function.Involutive f} // uncontractedFromInvolution f = c.1} @@ -245,35 +242,35 @@ def toInvolution : {φs : List 𝓕} → (c : Contractions φs) → rw [@Sigma.subtype_ext_iff] at hF0 ext1 rw [hF0.2] - simp + simp only [Nat.succ_eq_add_one] congr 1 · rw [hF1] · refine involutionAddEquiv_cast' ?_ n' _ _ rw [hF1] rw [uncontractedFromInvolution_cons] have hx := (toInvolution ⟨aux, c⟩).2 - simp at hx - simp + simp only at hx + simp only [Nat.succ_eq_add_one] refine optionEraseZ_ext ?_ ?_ ?_ - · dsimp [F] + · dsimp only [F] rw [Equiv.apply_symm_apply] - simp + simp only rw [← hx] simp_all only · rfl - · simp [hF2] - dsimp [n'] - simp [finCongr] + · simp only [hF2, Nat.succ_eq_add_one, Equiv.apply_symm_apply, Option.map_map] + dsimp only [id_eq, eq_mpr_eq_cast, Nat.succ_eq_add_one, n'] + simp only [finCongr, Equiv.coe_fn_mk, Option.map_map] simp only [Nat.succ_eq_add_one, id_eq, eq_mpr_eq_cast, F, n'] ext a : 1 simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans, Fin.cast_eq_self, exists_eq_right] -lemma toInvolution_length {φs φsᵤₙ : List 𝓕} {c : ContractionsAux φs φsᵤₙ} : - φsᵤₙ.length = (Finset.filter (fun i => (toInvolution ⟨φsᵤₙ, c⟩).1.1 i = i) Finset.univ).card - := by +lemma toInvolution_length_uncontracted {φs φsᵤₙ : List 𝓕} {c : ContractionsAux φs φsᵤₙ} : + φsᵤₙ.length = + (Finset.filter (fun i => (toInvolution ⟨φsᵤₙ, c⟩).1.1 i = i) Finset.univ).card := by have h2 := (toInvolution ⟨φsᵤₙ, c⟩).2 - simp at h2 + simp only at h2 conv_lhs => rw [← h2] exact Mathlib.Vector.length_val (uncontractedFromInvolution (toInvolution ⟨φsᵤₙ, c⟩).1) @@ -282,11 +279,11 @@ lemma toInvolution_cons {φs φsᵤₙ : List 𝓕} {φ : 𝓕} (toInvolution ⟨optionEraseZ φsᵤₙ φ n, ContractionsAux.cons n c⟩).1 = (involutionCons φs.length).symm ⟨(toInvolution ⟨φsᵤₙ, c⟩).1, (involutionAddEquiv (toInvolution ⟨φsᵤₙ, c⟩).1).symm - (Option.map (finCongr (toInvolution_length)) n)⟩ := by - dsimp [toInvolution] + (Option.map (finCongr (toInvolution_length_uncontracted)) n)⟩ := by + dsimp only [List.length_cons, toInvolution, Nat.succ_eq_add_one, subset_refl, Set.coe_inclusion] congr 3 rw [Option.map_map] - simp [finCongr] + simp only [finCongr, Equiv.coe_fn_mk] rfl lemma toInvolution_consEquiv {φs : List 𝓕} {φ : 𝓕} @@ -294,7 +291,7 @@ lemma toInvolution_consEquiv {φs : List 𝓕} {φ : 𝓕} (toInvolution ((consEquiv (φ := φ)).symm ⟨c, n⟩)).1 = (involutionCons φs.length).symm ⟨(toInvolution c).1, (involutionAddEquiv (toInvolution c).1).symm - (Option.map (finCongr (toInvolution_length)) n)⟩ := by + (Option.map (finCongr (toInvolution_length_uncontracted)) n)⟩ := by erw [toInvolution_cons] rfl @@ -308,13 +305,11 @@ lemma toInvolution_fromInvolution : {φs : List 𝓕} → (c : Contractions φs) fromInvolution (toInvolution c) = c | [], ⟨[], ContractionsAux.nil⟩ => rfl | φ :: φs, ⟨_, .cons (φsᵤₙ := φsᵤₙ) n c⟩ => by - rw [toInvolution_cons] - rw [fromInvolution_of_involutionCons] - rw [Equiv.symm_apply_eq] - dsimp [consEquiv] + rw [toInvolution_cons, fromInvolution_of_involutionCons, Equiv.symm_apply_eq] + dsimp only [consEquiv, Equiv.coe_fn_mk] refine consEquiv_ext ?_ ?_ · exact toInvolution_fromInvolution ⟨φsᵤₙ, c⟩ - · simp [finCongr] + · simp only [finCongr, Equiv.coe_fn_mk, Equiv.apply_symm_apply, Option.map_map] ext a : 1 simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans, Fin.cast_eq_self, exists_eq_right] @@ -337,9 +332,11 @@ lemma fromInvolution_toInvolution : {φs : List 𝓕} → (f : {f : Fin (φs ). conv_rhs => lhs rw [involutionAddEquiv_cast hx] - simp [Nat.succ_eq_add_one,- eq_mpr_eq_cast, Equiv.trans_apply, -Equiv.optionCongr_apply] + simp only [Nat.succ_eq_add_one, Equiv.trans_apply] rfl +/-- The equivalence between contractions and involutions. + Note: This shows that the type of contractions only depends on the length of the list `φs`. -/ def equivInvolutions {φs : List 𝓕} : Contractions φs ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f} where toFun := fun c => toInvolution c @@ -347,11 +344,11 @@ def equivInvolutions {φs : List 𝓕} : left_inv := toInvolution_fromInvolution right_inv := fromInvolution_toInvolution - /-! ## Full contractions and involutions. -/ + lemma isFull_iff_uncontractedFromInvolution_empty {φs : List 𝓕} (c : Contractions φs) : IsFull c ↔ (uncontractedFromInvolution (equivInvolutions c)).1 = [] := by let l := toInvolution c @@ -366,7 +363,7 @@ lemma isFull_iff_filter_card_involution_zero {φs : List 𝓕} (c : Contraction lemma isFull_iff_involution_no_fixed_points {φs : List 𝓕} (c : Contractions φs) : IsFull c ↔ ∀ (i : Fin φs.length), (equivInvolutions c).1 i ≠ i := by rw [isFull_iff_filter_card_involution_zero] - simp + simp only [Finset.card_eq_zero, ne_eq] rw [Finset.filter_eq_empty_iff] apply Iff.intro · intro h @@ -376,9 +373,9 @@ lemma isFull_iff_involution_no_fixed_points {φs : List 𝓕} (c : Contractions exact fun a => i h -open Nat in -def isFullInvolutionEquiv {φs : List 𝓕} : - {c : Contractions φs // IsFull c} ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f ∧ (∀ i, f i ≠ i)} where +/-- The equivalence between full contractions and fixed-point free involutions. -/ +def isFullInvolutionEquiv {φs : List 𝓕} : {c : Contractions φs // IsFull c} ≃ + {f : Fin φs.length → Fin φs.length // Function.Involutive f ∧ (∀ i, f i ≠ i)} where toFun c := ⟨equivInvolutions c.1, by apply And.intro (equivInvolutions c.1).2 rw [← isFull_iff_involution_no_fixed_points] @@ -392,5 +389,4 @@ def isFullInvolutionEquiv {φs : List 𝓕} : end Contractions - end Wick diff --git a/scripts/MetaPrograms/free_simps.lean b/scripts/MetaPrograms/free_simps.lean index 791b68a4..b6ade599 100644 --- a/scripts/MetaPrograms/free_simps.lean +++ b/scripts/MetaPrograms/free_simps.lean @@ -37,6 +37,8 @@ or is merely a tactic combinator (e.g. `by`, `;`, multiline tactics, parenthesiz def isSimp (t : TacticInfo) : Bool := match t.name? with | some ``Lean.Parser.Tactic.simp => true + | some ``Lean.Parser.Tactic.dsimp => true + | some ``Lean.Parser.Tactic.simpAll => true | _ => false end Lean.Elab.TacticInfo From 408a676bbda7c45332328c59a911859918c25f10 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Sun, 5 Jan 2025 17:00:36 +0000 Subject: [PATCH 7/8] refactor: style lint --- .../ComplexTensor/PauliMatrices/Basis.lean | 2 +- HepLean/Mathematics/Fin/Involutions.lean | 116 +++++++++--------- HepLean/Mathematics/List.lean | 3 +- .../Contractions/Basic.lean | 25 ++-- .../PerturbationTheory/Contractions/Card.lean | 4 +- .../Contractions/Involutions.lean | 63 +++++----- .../Wick/CreateAnnihilateSection.lean | 3 +- .../Wick/Signs/InsertSign.lean | 2 +- .../Wick/Signs/KoszulSign.lean | 4 +- .../Wick/StaticTheorem.lean | 2 +- 10 files changed, 111 insertions(+), 113 deletions(-) diff --git a/HepLean/Lorentz/ComplexTensor/PauliMatrices/Basis.lean b/HepLean/Lorentz/ComplexTensor/PauliMatrices/Basis.lean index 3186a12f..84a73ede 100644 --- a/HepLean/Lorentz/ComplexTensor/PauliMatrices/Basis.lean +++ b/HepLean/Lorentz/ComplexTensor/PauliMatrices/Basis.lean @@ -485,7 +485,7 @@ lemma pauliMatrix_contr_down_3 : rw [contrBasisVectorMul_pos _] conv => lhs; rhs; rhs; lhs; - rw [contrBasisVectorMul_pos _] + rw [contrBasisVectorMul_pos _] conv => lhs simp only [_root_.zero_smul, one_smul, _root_.smul_zero, _root_.add_zero, _root_.zero_add] diff --git a/HepLean/Mathematics/Fin/Involutions.lean b/HepLean/Mathematics/Fin/Involutions.lean index 8baed460..b5d76894 100644 --- a/HepLean/Mathematics/Fin/Involutions.lean +++ b/HepLean/Mathematics/Fin/Involutions.lean @@ -25,7 +25,7 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu toFun f := ⟨⟨ fun i => if h : f.1 i.succ = 0 then i - else Fin.pred (f.1 i.succ) h , by + else Fin.pred (f.1 i.succ) h, by intro i by_cases h : f.1 i.succ = 0 · simp [h] @@ -33,7 +33,7 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu simp only [f.2 i.succ, Fin.pred_succ, dite_eq_ite, ite_eq_right_iff] intro h exact False.elim (Fin.succ_ne_zero i h)⟩, - ⟨if h : f.1 0 = 0 then none else Fin.pred (f.1 0) h , by + ⟨if h : f.1 0 = 0 then none else Fin.pred (f.1 0) h, by by_cases h0 : f.1 0 = 0 · simp [h0] · simp only [succ_eq_add_one, h0, ↓reduceDIte, Option.isSome_some, Option.get_some, @@ -43,8 +43,7 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu if h : (f.2.1).isSome then Fin.cons (f.2.1.get h).succ (Function.update (Fin.succ ∘ f.1.1) (f.2.1.get h) 0) else - Fin.cons 0 (Fin.succ ∘ f.1.1) - , by + Fin.cons 0 (Fin.succ ∘ f.1.1), by by_cases hs : (f.2.1).isSome · simp only [Nat.succ_eq_add_one, hs, ↓reduceDIte, Fin.coe_eq_castSucc] let a := f.2.1.get hs @@ -59,7 +58,7 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu by_cases hja : j = a · subst hja simp - · rw [Function.update_apply ] + · rw [Function.update_apply] rw [if_neg hja] simp only [Function.comp_apply, Fin.cons_succ] have hf2 := f.2.2 hs @@ -145,12 +144,12 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu · rename_i h exact False.elim (Fin.succ_ne_zero (f i) h) · rfl - · simp only [Nat.succ_eq_add_one, Option.mem_def, + · simp only [Nat.succ_eq_add_one, Option.mem_def, Option.dite_none_left_eq_some, Option.some.injEq] by_cases hs : f0.isSome · simp only [hs, ↓reduceDIte] simp only [Fin.cons_zero, Fin.pred_succ, exists_prop] - have hx : ¬ (f0.get hs).succ = 0 := (Fin.succ_ne_zero (f0.get hs)) + have hx : ¬ (f0.get hs).succ = 0 := (Fin.succ_ne_zero (f0.get hs)) simp only [hx, not_false_eq_true, true_and] refine Iff.intro (fun hi => ?_) (fun hi => ?_) · rw [← hi] @@ -166,7 +165,7 @@ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involu subst hs exact ne_of_beq_false rfl -lemma involutionCons_ext {n : ℕ} {f1 f2 : (f : {f : Fin n → Fin n // Function.Involutive f}) × +lemma involutionCons_ext {n : ℕ} {f1 f2 : (f : {f : Fin n → Fin n // Function.Involutive f}) × {i : Option (Fin n) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)}} (h1 : f1.1 = f2.1) (h2 : f1.2 = Equiv.subtypeEquivRight (by rw [h1]; simp) f2.2) : f1 = f2 := by cases f1 @@ -181,10 +180,9 @@ lemma involutionCons_ext {n : ℕ} {f1 f2 : (f : {f : Fin n → Fin n // Functi simp_all only rfl - def involutionAddEquiv {n : ℕ} (f : {f : Fin n → Fin n // Function.Involutive f}) : {i : Option (Fin n) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} ≃ - Option (Fin (Finset.univ.filter fun i => f.1 i = i).card) := by + Option (Fin (Finset.univ.filter fun i => f.1 i = i).card) := by let e1 : {i : Option (Fin n) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} ≃ Option {i : Fin n // f.1 i = i} := { toFun := fun i => match i with @@ -208,14 +206,13 @@ def involutionAddEquiv {n : ℕ} (f : {f : Fin n → Fin n // Function.Involutiv funext i simp [s] let e2 : {i // i ∈ s} ≃ Fin (Finset.card s) := by - refine (Finset.orderIsoOfFin _ ?_).symm.toEquiv - simp [s] + refine (Finset.orderIsoOfFin _ ?_).symm.toEquiv + simp [s] refine e1.trans (Equiv.optionCongr (e2'.trans (e2))) - lemma involutionAddEquiv_none_image_zero {n : ℕ} : {f : {f : Fin n.succ → Fin n.succ // Function.Involutive f}} - → involutionAddEquiv (involutionCons n f).1 (involutionCons n f).2 = none + → involutionAddEquiv (involutionCons n f).1 (involutionCons n f).2 = none → f.1 ⟨0, Nat.zero_lt_succ n⟩ = ⟨0, Nat.zero_lt_succ n⟩ := by intro f h simp only [Nat.succ_eq_add_one, involutionCons, Equiv.coe_fn_mk, involutionAddEquiv, @@ -235,28 +232,29 @@ lemma involutionAddEquiv_none_image_zero {n : ℕ} : next h_2 => simp_all only [Subtype.mk.injEq, reduceCtorEq] lemma involutionAddEquiv_cast {n : ℕ} {f1 f2 : {f : Fin n → Fin n // Function.Involutive f}} - (hf : f1 = f2): - involutionAddEquiv f1 = (Equiv.subtypeEquivRight (by rw [hf]; simp)).trans - ((involutionAddEquiv f2).trans (Equiv.optionCongr (finCongr (by rw [hf])))):= by + (hf : f1 = f2) : + involutionAddEquiv f1 = (Equiv.subtypeEquivRight (by rw [hf]; simp)).trans + ((involutionAddEquiv f2).trans (Equiv.optionCongr (finCongr (by rw [hf])))) := by subst hf simp only [finCongr_refl, Equiv.optionCongr_refl, Equiv.trans_refl] rfl - lemma involutionAddEquiv_cast' {m : ℕ} {f1 f2 : {f : Fin m → Fin m // Function.Involutive f}} - {N : ℕ} (hf : f1 = f2) (n : Option (Fin N)) (hn1 : N = (Finset.filter (fun i => f1.1 i = i) Finset.univ).card) - (hn2 : N = (Finset.filter (fun i => f2.1 i = i) Finset.univ).card): - HEq ((involutionAddEquiv f1).symm (Option.map (finCongr hn1) n)) ((involutionAddEquiv f2).symm (Option.map (finCongr hn2) n)) := by + {N : ℕ} (hf : f1 = f2) (n : Option (Fin N)) + (hn1 : N = (Finset.filter (fun i => f1.1 i = i) Finset.univ).card) + (hn2 : N = (Finset.filter (fun i => f2.1 i = i) Finset.univ).card) : + HEq ((involutionAddEquiv f1).symm (Option.map (finCongr hn1) n)) + ((involutionAddEquiv f2).symm (Option.map (finCongr hn2) n)) := by subst hf rfl lemma involutionAddEquiv_none_succ {n : ℕ} {f : {f : Fin n.succ → Fin n.succ // Function.Involutive f}} - (h : involutionAddEquiv (involutionCons n f).1 (involutionCons n f).2 = none) - (x : Fin n) : f.1 x.succ = x.succ ↔ (involutionCons n f).1.1 x = x := by + (h : involutionAddEquiv (involutionCons n f).1 (involutionCons n f).2 = none) + (x : Fin n) : f.1 x.succ = x.succ ↔ (involutionCons n f).1.1 x = x := by simp only [succ_eq_add_one, involutionCons, Fin.cons_update, Equiv.coe_fn_mk, dite_eq_left_iff] have hn' := involutionAddEquiv_none_image_zero h - have hx : ¬ f.1 x.succ = ⟨0, Nat.zero_lt_succ n⟩:= by + have hx : ¬ f.1 x.succ = ⟨0, Nat.zero_lt_succ n⟩:= by rw [← hn'] exact fun hn => Fin.succ_ne_zero x (Function.Involutive.injective f.2 hn) apply Iff.intro @@ -286,7 +284,7 @@ lemma involutionAddEquiv_isSome_image_zero {n : ℕ} : def involutionNoFixedEquivSum {n : ℕ} : {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ ∀ i, f i ≠ i} ≃ Σ (k : Fin (2 * n + 1)), - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} where toFun f := ⟨(f.1 0).pred (f.2.2 0), ⟨f.1, f.2.1, by simpa using f.2.2⟩⟩ invFun f := ⟨f.2.1, ⟨f.2.2.1, f.2.2.2.1⟩⟩ @@ -310,7 +308,7 @@ The main aim of thes equivalences is to define `involutionNoFixedZeroEquivProd`. def involutionNoFixedZeroSetEquivEquiv {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive (e.symm ∘ f ∘ e) ∧ (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} where @@ -332,11 +330,12 @@ def involutionNoFixedZeroSetEquivEquiv {n : ℕ} ext i simp -def involutionNoFixedZeroSetEquivSetEquiv {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : +def involutionNoFixedZeroSetEquivSetEquiv {n : ℕ} (k : Fin (2 * n + 1)) + (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive (e.symm ∘ f ∘ e) ∧ (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} := by + (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} := by refine Equiv.subtypeEquivRight ?_ intro f have h1 : Function.Involutive (⇑e.symm ∘ f ∘ ⇑e) ↔ Function.Involutive f := by @@ -360,12 +359,12 @@ def involutionNoFixedZeroSetEquivSetEquiv {n : ℕ} (k : Fin (2 * n + 1)) (e : F nth_rewrite 2 [← hn] at hi simp at hi - -def involutionNoFixedZeroSetEquivEquiv' {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : +def involutionNoFixedZeroSetEquivEquiv' {n : ℕ} (k : Fin (2 * n + 1)) + (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} + (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f (e 0) = e k.succ} := by + (∀ i, f i ≠ i) ∧ f (e 0) = e k.succ} := by refine Equiv.subtypeEquivRight ?_ simp only [succ_eq_add_one, ne_eq, Function.comp_apply, and_congr_right_iff] intro f hi h1 @@ -373,13 +372,13 @@ def involutionNoFixedZeroSetEquivEquiv' {n : ℕ} (k : Fin (2 * n + 1)) (e : Fin def involutionNoFixedZeroSetEquivSetOne {n : ℕ} (k : Fin (2 * n + 1)) : {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f 0 = k.succ} + (∀ i, f i ≠ i) ∧ f 0 = k.succ} ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f 0 = 1} := by + (∀ i, f i ≠ i) ∧ f 0 = 1} := by refine Equiv.trans (involutionNoFixedZeroSetEquivEquiv k (Equiv.swap k.succ 1)) ?_ refine Equiv.trans (involutionNoFixedZeroSetEquivSetEquiv k (Equiv.swap k.succ 1)) ?_ refine Equiv.trans (involutionNoFixedZeroSetEquivEquiv' k (Equiv.swap k.succ 1)) ?_ - refine Equiv.subtypeEquivRight ?_ + refine Equiv.subtypeEquivRight ?_ simp only [succ_eq_add_one, ne_eq, Equiv.swap_apply_left, and_congr_right_iff] intro f hi h1 rw [Equiv.swap_apply_of_ne_of_ne] @@ -387,10 +386,9 @@ def involutionNoFixedZeroSetEquivSetOne {n : ℕ} (k : Fin (2 * n + 1)) : · exact Fin.zero_ne_one def involutionNoFixedSetOne {n : ℕ} : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f 0 = 1} ≃ - {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ - (∀ i, f i ≠ i)} where + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f 0 = 1} ≃ {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ + (∀ i, f i ≠ i)} where toFun f := by have hf1 : f.1 1 = 0 := by have hf := f.2.2.2 @@ -444,7 +442,7 @@ def involutionNoFixedSetOne {n : ℕ} : rename_i x r simp_all only [succ_eq_add_one, Fin.ext_iff, Fin.val_succ, add_left_inj] have hfn {a b : ℕ} {ha : a < 2 * n} {hb : b < 2 * n} - (hab : ↑(f.1 ⟨a, ha⟩) = b): ↑(f.1 ⟨b, hb⟩) = a := by + (hab : ↑(f.1 ⟨a, ha⟩) = b) : ↑(f.1 ⟨b, hb⟩) = a := by have ht : f.1 ⟨a, ha⟩ = ⟨b, hb⟩ := by simp [hab, Fin.ext_iff] rw [← ht, f.2.1] @@ -504,7 +502,7 @@ def involutionNoFixedSetOne {n : ℕ} : simp only [Fin.coe_pred] split · rename_i h - simp [Fin.ext_iff] at h + simp [Fin.ext_iff] at h · rename_i h simp [Fin.ext_iff] at h · rename_i h @@ -513,22 +511,23 @@ def involutionNoFixedSetOne {n : ℕ} : apply congrArg simp_all [Fin.ext_iff] - def involutionNoFixedZeroSetEquiv {n : ℕ} (k : Fin (2 * n + 1)) : {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f 0 = k.succ} + (∀ i, f i ≠ i) ∧ f 0 = k.succ} ≃ {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by refine Equiv.trans (involutionNoFixedZeroSetEquivSetOne k) involutionNoFixedSetOne def involutionNoFixedEquivSumSame {n : ℕ} : {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} - ≃ Σ (_ : Fin (2 * n + 1)), {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + ≃ Σ (_ : Fin (2 * n + 1)), + {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by refine Equiv.trans involutionNoFixedEquivSum ?_ refine Equiv.sigmaCongrRight involutionNoFixedZeroSetEquiv def involutionNoFixedZeroEquivProd {n : ℕ} : {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} - ≃ Fin (2 * n + 1) × {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + ≃ Fin (2 * n + 1) × + {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by refine Equiv.trans involutionNoFixedEquivSumSame ?_ exact Equiv.sigmaEquivProd (Fin (2 * n + 1)) { f // Function.Involutive f ∧ ∀ (i : Fin (2 * n)), f i ≠ i} @@ -539,27 +538,24 @@ def involutionNoFixedZeroEquivProd {n : ℕ} : -/ - -instance {n : ℕ} : Fintype { f // Function.Involutive f ∧ ∀ (i : Fin n), f i ≠ i } := by - haveI : DecidablePred fun x => Function.Involutive x := by - intro f - apply Fintype.decidableForallFintype (α := Fin n) - haveI : DecidablePred fun x => Function.Involutive x ∧ ∀ (i : Fin n), x i ≠ i := by - intro x - apply instDecidableAnd +instance {n : ℕ} : Fintype { f // Function.Involutive f ∧ ∀ (i : Fin n), f i ≠ i } := by + haveI : DecidablePred fun x => Function.Involutive x := + fun f => Fintype.decidableForallFintype (α := Fin n) + haveI : DecidablePred fun x => Function.Involutive x ∧ ∀ (i : Fin n), x i ≠ i := + fun x => instDecidableAnd apply Subtype.fintype lemma involutionNoFixed_card_succ {n : ℕ} : - Fintype.card {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} - = (2 * n + 1) * Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by - rw [Fintype.card_congr (involutionNoFixedZeroEquivProd)] - rw [Fintype.card_prod ] + Fintype.card + {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + = (2 * n + 1) * + Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + rw [Fintype.card_congr (involutionNoFixedZeroEquivProd), Fintype.card_prod] congr exact Fintype.card_fin (2 * n + 1) - lemma involutionNoFixed_card_mul_two : (n : ℕ) → - Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} = (2 * n - 1)‼ | 0 => rfl | Nat.succ n => by @@ -568,7 +564,7 @@ lemma involutionNoFixed_card_mul_two : (n : ℕ) → exact Eq.symm (Nat.doubleFactorial_add_one (Nat.mul 2 n)) lemma involutionNoFixed_card_even : (n : ℕ) → (he : Even n) → - Fintype.card {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} = (n - 1)‼ := by + Fintype.card {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} = (n - 1)‼ := by intro n he obtain ⟨r, hr⟩ := he have hr' : n = 2 * r := by omega diff --git a/HepLean/Mathematics/List.lean b/HepLean/Mathematics/List.lean index ef5ed81b..5966221d 100644 --- a/HepLean/Mathematics/List.lean +++ b/HepLean/Mathematics/List.lean @@ -662,7 +662,6 @@ def optionErase {I : Type} (l : List I) (i : Option (Fin l.length)) : List I := | none => l | some i => List.eraseIdx l i - lemma eraseIdx_length {I : Type} (l : List I) (i : Fin l.length) : (List.eraseIdx l i).length + 1 = l.length := by simp only [List.length_eraseIdx, Fin.is_lt, ↓reduceIte] @@ -744,7 +743,7 @@ lemma optionEraseZ_some_length {I : Type} (l : List I) (a : I) (i : (Fin l.lengt lemma optionEraseZ_ext {I : Type} {l l' : List I} {a a' : I} {i : Option (Fin l.length)} {i' : Option (Fin l'.length)} (hl : l = l') (ha : a = a') (hi : Option.map (Fin.cast (by rw [hl])) i = i') : - optionEraseZ l a i = optionEraseZ l' a' i' := by + optionEraseZ l a i = optionEraseZ l' a' i' := by subst hl subst ha cases hi diff --git a/HepLean/PerturbationTheory/Contractions/Basic.lean b/HepLean/PerturbationTheory/Contractions/Basic.lean index f4af2b31..57faf247 100644 --- a/HepLean/PerturbationTheory/Contractions/Basic.lean +++ b/HepLean/PerturbationTheory/Contractions/Basic.lean @@ -33,12 +33,12 @@ namespace Contractions variable {l : List 𝓕} (c : Contractions l) -def auxCongr : {φs: List 𝓕} → {φsᵤₙ φsᵤₙ' : List 𝓕} → (h : φsᵤₙ = φsᵤₙ') → +def auxCongr : {φs: List 𝓕} → {φsᵤₙ φsᵤₙ' : List 𝓕} → (h : φsᵤₙ = φsᵤₙ') → ContractionsAux φs φsᵤₙ ≃ ContractionsAux φs φsᵤₙ' | _, _, _, Eq.refl _ => Equiv.refl _ lemma auxCongr_ext {φs: List 𝓕} {c c2 : Contractions φs} (h : c.1 = c2.1) - (hx : c.2 = auxCongr h.symm c2.2) : c = c2 := by + (hx : c.2 = auxCongr h.symm c2.2) : c = c2 := by cases c cases c2 simp only at h @@ -50,7 +50,7 @@ lemma auxCongr_ext {φs: List 𝓕} {c c2 : Contractions φs} (h : c.1 = c2.1) /-- The list of uncontracted fields. -/ def uncontracted : List 𝓕 := c.1 -lemma uncontracted_length_even_iff : {l : List 𝓕} → (c : Contractions l) → +lemma uncontracted_length_even_iff : {l : List 𝓕} → (c : Contractions l) → Even l.length ↔ Even c.uncontracted.length | [], ⟨[], ContractionsAux.nil⟩ => by simp [uncontracted] @@ -76,7 +76,8 @@ lemma contractions_nil (a : Contractions ([] : List 𝓕)) : a = ⟨[], Contract cases c rfl -def embedUncontracted {l : List 𝓕} (c : Contractions l) : Fin c.uncontracted.length → Fin l.length := +def embedUncontracted {l : List 𝓕} (c : Contractions l) : + Fin c.uncontracted.length → Fin l.length := match l, c with | [], ⟨[], ContractionsAux.nil⟩ => Fin.elim0 | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => @@ -93,7 +94,7 @@ def embedUncontracted {l : List 𝓕} (c : Contractions l) : Fin c.uncontracted. exact Fin.elim0 n omega -lemma embedUncontracted_injective {l : List 𝓕} (c : Contractions l) : +lemma embedUncontracted_injective {l : List 𝓕} (c : Contractions l) : Function.Injective c.embedUncontracted := by match l, c with | [], ⟨[], ContractionsAux.nil⟩ => @@ -108,12 +109,12 @@ lemma embedUncontracted_injective {l : List 𝓕} (c : Contractions l) : exact fun x => Fin.succ_ne_zero (embedUncontracted ⟨aux, c⟩ x) · exact Function.Injective.comp (Fin.succ_injective φs.length) (embedUncontracted_injective ⟨aux, c⟩) - | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some i) c⟩ => - dsimp only [List.length_cons, embedUncontracted] - refine Function.Injective.comp (Fin.succ_injective φs.length) ?hf - refine Function.Injective.comp (embedUncontracted_injective ⟨aux, c⟩) ?hf.hf - refine Function.Injective.comp (Fin.cast_injective (embedUncontracted.proof_5 φ φs aux i c)) - Fin.succAbove_right_injective + | φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some i) c⟩ => + dsimp only [List.length_cons, embedUncontracted] + refine Function.Injective.comp (Fin.succ_injective φs.length) ?hf + refine Function.Injective.comp (embedUncontracted_injective ⟨aux, c⟩) ?hf.hf + refine Function.Injective.comp (Fin.cast_injective (embedUncontracted.proof_5 φ φs aux i c)) + Fin.succAbove_right_injective /-- Establishes uniqueness of contractions for a single field, showing that any contraction of a single field must be equivalent to the trivial contraction with no pairing. -/ @@ -161,7 +162,7 @@ def consEquiv {φ : 𝓕} {φs : List 𝓕} : lemma consEquiv_ext {φs : List 𝓕} {c1 c2 : Contractions φs} {n1 : Option (Fin c1.uncontracted.length)} {n2 : Option (Fin c2.uncontracted.length)} (hc : c1 = c2) (hn : Option.map (finCongr (by rw [hc])) n1 = n2) : - (⟨c1, n1⟩ : (c : Contractions φs) × Option (Fin c.uncontracted.length)) = ⟨c2, n2⟩ := by + (⟨c1, n1⟩ : (c : Contractions φs) × Option (Fin c.uncontracted.length)) = ⟨c2, n2⟩ := by subst hc subst hn simp diff --git a/HepLean/PerturbationTheory/Contractions/Card.lean b/HepLean/PerturbationTheory/Contractions/Card.lean index 0edb7f97..03a9ef3f 100644 --- a/HepLean/PerturbationTheory/Contractions/Card.lean +++ b/HepLean/PerturbationTheory/Contractions/Card.lean @@ -17,13 +17,13 @@ open HepLean.Fin open Nat /-- There are `(φs.length - 1)‼` full contractions of a list `φs` with an even number of fields. -/ -lemma card_of_full_contractions_even {φs : List 𝓕} (he : Even φs.length ) : +lemma card_of_full_contractions_even {φs : List 𝓕} (he : Even φs.length) : Fintype.card {c : Contractions φs // IsFull c} = (φs.length - 1)‼ := by rw [Fintype.card_congr (isFullInvolutionEquiv (φs := φs))] exact involutionNoFixed_card_even φs.length he /-- There are no full contractions of a list with an odd number of fields. -/ -lemma card_of_full_contractions_odd {φs : List 𝓕} (ho : Odd φs.length ) : +lemma card_of_full_contractions_odd {φs : List 𝓕} (ho : Odd φs.length) : Fintype.card {c : Contractions φs // IsFull c} = 0 := by rw [Fintype.card_eq_zero_iff, isEmpty_subtype] intro c diff --git a/HepLean/PerturbationTheory/Contractions/Involutions.lean b/HepLean/PerturbationTheory/Contractions/Involutions.lean index d624ad12..65b4d4b2 100644 --- a/HepLean/PerturbationTheory/Contractions/Involutions.lean +++ b/HepLean/PerturbationTheory/Contractions/Involutions.lean @@ -38,7 +38,7 @@ variable {l : List 𝓕} /-- Given an involution the uncontracted fields associated with it (corresponding to the fixed points of that involution). -/ -def uncontractedFromInvolution : {φs : List 𝓕} → +def uncontractedFromInvolution : {φs : List 𝓕} → (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) → {l : List 𝓕 // l.length = (Finset.univ.filter fun i => f.1 i = i).card} | [], _ => ⟨[], by simp⟩ @@ -46,7 +46,7 @@ def uncontractedFromInvolution : {φs : List 𝓕} → let luc := uncontractedFromInvolution (involutionCons φs.length f).fst let n' := involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2 let np : Option (Fin luc.1.length) := Option.map (finCongr luc.2.symm) n' - if hn : n' = none then + if hn : n' = none then have hn' := involutionAddEquiv_none_image_zero (n := φs.length) (f := f) hn ⟨optionEraseZ luc φ none, by simp only [optionEraseZ, Nat.succ_eq_add_one, List.length_cons, Mathlib.Vector.length_val] @@ -55,7 +55,7 @@ def uncontractedFromInvolution : {φs : List 𝓕} → rw [Fin.sum_univ_succ] conv_rhs => erw [if_pos hn'] ring_nf - simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, Nat.cast_id, + simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, Nat.cast_id, add_right_inj] rw [Finset.card_filter] apply congrArg @@ -72,17 +72,18 @@ def uncontractedFromInvolution : {φs : List 𝓕} → Option.isSome_none, Equiv.trans_apply, Equiv.coe_fn_mk, Equiv.optionCongr_apply, Equiv.coe_trans, RelIso.coe_fn_toEquiv, Option.map_eq_none', n'] at hn split at hn - · simp_all only [reduceCtorEq, not_false_eq_true, Nat.succ_eq_add_one, Option.isSome_some, k'] + · simp_all only [reduceCtorEq, not_false_eq_true, Nat.succ_eq_add_one, + Option.isSome_some, k'] · simp_all only [not_true_eq_false] let k := k'.1.get hkIsSome rw [optionEraseZ_some_length] have hksucc : k.succ = f.1 ⟨0, Nat.zero_lt_succ φs.length⟩ := by simp [k, k', involutionCons] - have hzero : ⟨0, Nat.zero_lt_succ φs.length⟩ = f.1 k.succ := by + have hzero : ⟨0, Nat.zero_lt_succ φs.length⟩ = f.1 k.succ := by rw [hksucc, f.2] have hksuccNe : f.1 k.succ ≠ k.succ := by conv_rhs => rw [hksucc] - exact fun hn => Fin.succ_ne_zero k (Function.Involutive.injective f.2 hn ) + exact fun hn => Fin.succ_ne_zero k (Function.Involutive.injective f.2 hn) have hluc : 1 ≤ luc.1.length := by simp only [Nat.succ_eq_add_one, Mathlib.Vector.length_val, Finset.one_le_card] use k @@ -131,8 +132,8 @@ lemma uncontractedFromInvolution_cons {φs : List 𝓕} {φ : 𝓕} (involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2)) := by let luc := uncontractedFromInvolution (involutionCons φs.length f).fst let n' := involutionAddEquiv (involutionCons φs.length f).1 (involutionCons φs.length f).2 - change _ = optionEraseZ luc φ - (Option.map (finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm)) n') + change _ = optionEraseZ luc φ (Option.map + (finCongr ((uncontractedFromInvolution (involutionCons φs.length f).fst).2.symm)) n') dsimp only [List.length_cons, uncontractedFromInvolution, Nat.succ_eq_add_one, Fin.zero_eta] by_cases hn : n' = none · have hn' := hn @@ -165,10 +166,10 @@ lemma uncontractedFromInvolution_cons {φs : List 𝓕} {φ : 𝓕} simp_all only [Option.get_some] /-- The `ContractionsAux` associated to an involution. -/ -def fromInvolutionAux : {l : List 𝓕} → +def fromInvolutionAux : {l : List 𝓕} → (f : {f : Fin l.length → Fin l.length // Function.Involutive f}) → ContractionsAux l (uncontractedFromInvolution f) - | [] => fun _ => ContractionsAux.nil + | [] => fun _ => ContractionsAux.nil | _ :: φs => fun f => let f' := involutionCons φs.length f let c' := fromInvolutionAux f'.1 @@ -177,8 +178,9 @@ def fromInvolutionAux : {l : List 𝓕} → auxCongr (uncontractedFromInvolution_cons f).symm (ContractionsAux.cons n' c') /-- The contraction associated with an involution. -/ -def fromInvolution {φs : List 𝓕} (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) : - Contractions φs := ⟨uncontractedFromInvolution f, fromInvolutionAux f⟩ +def fromInvolution {φs : List 𝓕} + (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) : Contractions φs := + ⟨uncontractedFromInvolution f, fromInvolutionAux f⟩ lemma fromInvolution_cons {φs : List 𝓕} {φ : 𝓕} (f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) : @@ -195,8 +197,8 @@ lemma fromInvolution_cons {φs : List 𝓕} {φ : 𝓕} rfl lemma fromInvolution_of_involutionCons {φs : List 𝓕} {φ : 𝓕} - (f : {f : Fin (φs ).length → Fin (φs).length // Function.Involutive f}) - (n : { i : Option (Fin φs.length) // ∀ (h : i.isSome = true), f.1 (i.get h) = i.get h }): + (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) + (n : { i : Option (Fin φs.length) // ∀ (h : i.isSome = true), f.1 (i.get h) = i.get h }) : fromInvolution (φs := φ :: φs) ((involutionCons φs.length).symm ⟨f, n⟩) = consEquiv.symm ⟨fromInvolution f, Option.map (finCongr ((uncontractedFromInvolution f).2.symm)) (involutionAddEquiv f n)⟩ := by @@ -212,7 +214,7 @@ lemma fromInvolution_of_involutionCons {φs : List 𝓕} {φ : 𝓕} -/ /-- The involution associated with a contraction. -/ -def toInvolution : {φs : List 𝓕} → (c : Contractions φs) → +def toInvolution : {φs : List 𝓕} → (c : Contractions φs) → {f : {f : Fin φs.length → Fin φs.length // Function.Involutive f} // uncontractedFromInvolution f = c.1} | [], ⟨[], ContractionsAux.nil⟩ => ⟨⟨fun i => i, by @@ -223,22 +225,23 @@ def toInvolution : {φs : List 𝓕} → (c : Contractions φs) → let n' : Option (Fin (uncontractedFromInvolution ⟨f', hf1⟩).1.length) := Option.map (finCongr (by rw [hf2])) n let F := (involutionCons φs.length).symm ⟨⟨f', hf1⟩, - (involutionAddEquiv ⟨f', hf1⟩).symm - (Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ + (involutionAddEquiv ⟨f', hf1⟩).symm + (Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ refine ⟨F, ?_⟩ have hF0 : ((involutionCons φs.length) F) = ⟨⟨f', hf1⟩, - (involutionAddEquiv ⟨f', hf1⟩).symm - (Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ := by + (involutionAddEquiv ⟨f', hf1⟩).symm + (Option.map (finCongr ((uncontractedFromInvolution ⟨f', hf1⟩).2)) n')⟩ := by simp [F] have hF1 : ((involutionCons φs.length) F).fst = ⟨f', hf1⟩ := by rw [hF0] have hF2L : ((uncontractedFromInvolution ⟨f', hf1⟩)).1.length = (Finset.filter (fun i => ((involutionCons φs.length) F).1.1 i = i) Finset.univ).card := by - apply Eq.trans ((uncontractedFromInvolution ⟨f', hf1⟩)).2 + apply Eq.trans ((uncontractedFromInvolution ⟨f', hf1⟩)).2 congr rw [hF1] - have hF2 : ((involutionCons φs.length) F).snd = (involutionAddEquiv ((involutionCons φs.length) F).fst).symm - (Option.map (finCongr hF2L) n') := by + have hF2 : ((involutionCons φs.length) F).snd = + (involutionAddEquiv ((involutionCons φs.length) F).fst).symm + (Option.map (finCongr hF2L) n') := by rw [@Sigma.subtype_ext_iff] at hF0 ext1 rw [hF0.2] @@ -314,8 +317,9 @@ lemma toInvolution_fromInvolution : {φs : List 𝓕} → (c : Contractions φs) simp only [Option.mem_def, Option.map_eq_some', Function.comp_apply, Fin.cast_trans, Fin.cast_eq_self, exists_eq_right] -lemma fromInvolution_toInvolution : {φs : List 𝓕} → (f : {f : Fin (φs ).length → Fin (φs).length // Function.Involutive f}) - → toInvolution (fromInvolution f) = f +lemma fromInvolution_toInvolution : {φs : List 𝓕} → + (f : {f : Fin φs.length → Fin φs.length // Function.Involutive f}) → + toInvolution (fromInvolution f) = f | [], _ => by ext x exact Fin.elim0 x @@ -331,7 +335,7 @@ lemma fromInvolution_toInvolution : {φs : List 𝓕} → (f : {f : Fin (φs ). rw [Equiv.symm_apply_eq] conv_rhs => lhs - rw [involutionAddEquiv_cast hx] + rw [involutionAddEquiv_cast hx] simp only [Nat.succ_eq_add_one, Equiv.trans_apply] rfl @@ -339,7 +343,7 @@ lemma fromInvolution_toInvolution : {φs : List 𝓕} → (f : {f : Fin (φs ). Note: This shows that the type of contractions only depends on the length of the list `φs`. -/ def equivInvolutions {φs : List 𝓕} : Contractions φs ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f} where - toFun := fun c => toInvolution c + toFun := fun c => toInvolution c invFun := fromInvolution left_inv := toInvolution_fromInvolution right_inv := fromInvolution_toInvolution @@ -355,7 +359,7 @@ lemma isFull_iff_uncontractedFromInvolution_empty {φs : List 𝓕} (c : Contrac erw [l.2] rfl -lemma isFull_iff_filter_card_involution_zero {φs : List 𝓕} (c : Contractions φs) : +lemma isFull_iff_filter_card_involution_zero {φs : List 𝓕} (c : Contractions φs) : IsFull c ↔ (Finset.univ.filter fun i => (equivInvolutions c).1 i = i).card = 0 := by rw [isFull_iff_uncontractedFromInvolution_empty, List.ext_get_iff] simp @@ -372,21 +376,18 @@ lemma isFull_iff_involution_no_fixed_points {φs : List 𝓕} (c : Contractions · intro i h exact fun a => i h - /-- The equivalence between full contractions and fixed-point free involutions. -/ def isFullInvolutionEquiv {φs : List 𝓕} : {c : Contractions φs // IsFull c} ≃ {f : Fin φs.length → Fin φs.length // Function.Involutive f ∧ (∀ i, f i ≠ i)} where toFun c := ⟨equivInvolutions c.1, by apply And.intro (equivInvolutions c.1).2 rw [← isFull_iff_involution_no_fixed_points] - exact c.2 - ⟩ + exact c.2⟩ invFun f := ⟨equivInvolutions.symm ⟨f.1, f.2.1⟩, by rw [isFull_iff_involution_no_fixed_points] simpa using f.2.2⟩ left_inv c := by simp right_inv f := by simp - end Contractions end Wick diff --git a/HepLean/PerturbationTheory/Wick/CreateAnnihilateSection.lean b/HepLean/PerturbationTheory/Wick/CreateAnnihilateSection.lean index 1d43b95b..68ffbcf6 100644 --- a/HepLean/PerturbationTheory/Wick/CreateAnnihilateSection.lean +++ b/HepLean/PerturbationTheory/Wick/CreateAnnihilateSection.lean @@ -26,7 +26,8 @@ namespace CreateAnnihilateSect section basic_defs -variable {𝓕 : Type} {f : 𝓕 → Type} [∀ i, Fintype (f i)] {φs : List 𝓕} (a : CreateAnnihilateSect f φs) +variable {𝓕 : Type} {f : 𝓕 → Type} [∀ i, Fintype (f i)] {φs : List 𝓕} + (a : CreateAnnihilateSect f φs) /-- The type `CreateAnnihilateSect f φs` is finite. -/ instance fintype : Fintype (CreateAnnihilateSect f φs) := Pi.fintype diff --git a/HepLean/PerturbationTheory/Wick/Signs/InsertSign.lean b/HepLean/PerturbationTheory/Wick/Signs/InsertSign.lean index c6d21dd4..4a7b818c 100644 --- a/HepLean/PerturbationTheory/Wick/Signs/InsertSign.lean +++ b/HepLean/PerturbationTheory/Wick/Signs/InsertSign.lean @@ -84,7 +84,7 @@ def insertSign (n : ℕ) (φ : 𝓕) (φs : List 𝓕) : ℂ := superCommuteCoef q [φ] (List.take n φs) /-- If `φ` is bosonic, there is no sign associated with inserting it into a list of fields. -/ -lemma insertSign_bosonic (n : ℕ) (φ : 𝓕) (φs : List 𝓕) (hφ : q φ = bosonic) : +lemma insertSign_bosonic (n : ℕ) (φ : 𝓕) (φs : List 𝓕) (hφ : q φ = bosonic) : insertSign q n φ φs = 1 := by simp only [insertSign, superCommuteCoef, ofList_singleton, hφ, reduceCtorEq, false_and, ↓reduceIte] diff --git a/HepLean/PerturbationTheory/Wick/Signs/KoszulSign.lean b/HepLean/PerturbationTheory/Wick/Signs/KoszulSign.lean index dfc2a2e2..31ce0bcf 100644 --- a/HepLean/PerturbationTheory/Wick/Signs/KoszulSign.lean +++ b/HepLean/PerturbationTheory/Wick/Signs/KoszulSign.lean @@ -160,10 +160,10 @@ lemma koszulSign_insertIdx [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) : rw [← insertionSortEquiv_get] simp only [Function.comp_apply, Equiv.symm_apply_apply, List.get_eq_getElem, ni] simp_all only [List.length_cons, add_le_add_iff_right, List.getElem_insertIdx_self] - have hc1 (hninro : ni.castSucc < nro) : ¬ le φ1 φ := by + have hc1 (hninro : ni.castSucc < nro) : ¬ le φ1 φ := by rw [← hns] exact lt_orderedInsertPos_rel le φ1 rs ni hninro - have hc2 (hninro : ¬ ni.castSucc < nro) : le φ1 φ := by + have hc2 (hninro : ¬ ni.castSucc < nro) : le φ1 φ := by rw [← hns] refine gt_orderedInsertPos_rel le φ1 rs ?_ ni hninro exact List.sorted_insertionSort le (List.insertIdx n φ φs) diff --git a/HepLean/PerturbationTheory/Wick/StaticTheorem.lean b/HepLean/PerturbationTheory/Wick/StaticTheorem.lean index 16bad3ad..ab50d867 100644 --- a/HepLean/PerturbationTheory/Wick/StaticTheorem.lean +++ b/HepLean/PerturbationTheory/Wick/StaticTheorem.lean @@ -29,7 +29,7 @@ lemma static_wick_nil {A : Type} [Semiring A] [Algebra ℂ A] rw [← Contractions.nilEquiv.symm.sum_comp] simp only [Finset.univ_unique, PUnit.default_eq_unit, Contractions.nilEquiv, Equiv.coe_fn_symm_mk, Finset.sum_const, Finset.card_singleton, one_smul] - dsimp [Contractions.uncontracted, Contractions.toCenterTerm] + dsimp only [Contractions.toCenterTerm, Contractions.uncontracted] simp [ofListLift_empty] lemma static_wick_cons [IsTrans ((i : 𝓕) × f i) le] [IsTotal ((i : 𝓕) × f i) le] From 83908c6d0df511546d960a8ff6bb913ba3c93778 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Mon, 6 Jan 2025 05:35:35 +0000 Subject: [PATCH 8/8] refactor: Lint and update involutions of Fin n --- HepLean/Mathematics/Fin/Involutions.lean | 150 +++++++++++------- .../Contractions/Basic.lean | 7 +- .../PerturbationTheory/Contractions/Card.lean | 11 +- .../Contractions/Involutions.lean | 2 +- 4 files changed, 103 insertions(+), 67 deletions(-) diff --git a/HepLean/Mathematics/Fin/Involutions.lean b/HepLean/Mathematics/Fin/Involutions.lean index b5d76894..addb453c 100644 --- a/HepLean/Mathematics/Fin/Involutions.lean +++ b/HepLean/Mathematics/Fin/Involutions.lean @@ -19,6 +19,9 @@ namespace HepLean.Fin open Nat +/-- There is an equivalence between involutions of `Fin n.succ` and involutions of + `Fin n` and an optional valid choice of an element in `Fin n` (which is where `0` + in `Fin n.succ` will be sent). -/ def involutionCons (n : ℕ) : {f : Fin n.succ → Fin n.succ // Function.Involutive f } ≃ (f : {f : Fin n → Fin n // Function.Involutive f}) × {i : Option (Fin n) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} where @@ -180,6 +183,9 @@ lemma involutionCons_ext {n : ℕ} {f1 f2 : (f : {f : Fin n → Fin n // Functio simp_all only rfl +/-- Given an involution of `Fin n`, the optional choice of an element in `Fin n` which + maps to itself is equivalent to the optional choice of an element in + `Fin (Finset.univ.filter fun i => f.1 i = i).card`. -/ def involutionAddEquiv {n : ℕ} (f : {f : Fin n → Fin n // Function.Involutive f}) : {i : Option (Fin n) // ∀ (h : i.isSome), f.1 (Option.get i h) = (Option.get i h)} ≃ Option (Fin (Finset.univ.filter fun i => f.1 i = i).card) := by @@ -240,9 +246,9 @@ lemma involutionAddEquiv_cast {n : ℕ} {f1 f2 : {f : Fin n → Fin n // Functio rfl lemma involutionAddEquiv_cast' {m : ℕ} {f1 f2 : {f : Fin m → Fin m // Function.Involutive f}} - {N : ℕ} (hf : f1 = f2) (n : Option (Fin N)) - (hn1 : N = (Finset.filter (fun i => f1.1 i = i) Finset.univ).card) - (hn2 : N = (Finset.filter (fun i => f2.1 i = i) Finset.univ).card) : + {N : ℕ} (hf : f1 = f2) (n : Option (Fin N)) + (hn1 : N = (Finset.filter (fun i => f1.1 i = i) Finset.univ).card) + (hn2 : N = (Finset.filter (fun i => f2.1 i = i) Finset.univ).card) : HEq ((involutionAddEquiv f1).symm (Option.map (finCongr hn1) n)) ((involutionAddEquiv f2).symm (Option.map (finCongr hn2) n)) := by subst hf @@ -281,10 +287,19 @@ lemma involutionAddEquiv_isSome_image_zero {n : ℕ} : intro a simp_all only [↓reduceDIte, Option.isSome_none, Bool.false_eq_true] +/-! + +## Equivalences of involutions with no fixed points. + +The main aim of thes equivalences is to define `involutionNoFixedZeroEquivProd`. + +-/ + +/-- Fixed point free involutions of `Fin n.succ` can be seperated based on where they sent + `0`. -/ def involutionNoFixedEquivSum {n : ℕ} : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f - ∧ ∀ i, f i ≠ i} ≃ Σ (k : Fin (2 * n + 1)), - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f + {f : Fin n.succ → Fin n.succ // Function.Involutive f + ∧ ∀ i, f i ≠ i} ≃ Σ (k : Fin n), {f : Fin n.succ → Fin n.succ // Function.Involutive f ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} where toFun f := ⟨(f.1 0).pred (f.2.2 0), ⟨f.1, f.2.1, by simpa using f.2.2⟩⟩ invFun f := ⟨f.2.1, ⟨f.2.2.1, f.2.2.2.1⟩⟩ @@ -298,19 +313,12 @@ def involutionNoFixedEquivSum {n : ℕ} : simp · simp -/-! - -## Equivalences of involutions with no fixed points. - -The main aim of thes equivalences is to define `involutionNoFixedZeroEquivProd`. - --/ - +/-- The condition on fixed point free involutions of `Fin n.succ` for a fixed value of `f 0`, + can be modified by conjugation with an equivalence. -/ def involutionNoFixedZeroSetEquivEquiv {n : ℕ} - (k : Fin (2 * n + 1)) (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f - ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} ≃ - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive (e.symm ∘ f ∘ e) ∧ + (k : Fin n) (e : Fin n.succ ≃ Fin n.succ) : + {f : Fin n.succ → Fin n.succ // Function.Involutive f ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} ≃ + {f : Fin n.succ → Fin n.succ // Function.Involutive (e.symm ∘ f ∘ e) ∧ (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} where toFun f := ⟨e ∘ f.1 ∘ e.symm, by intro i @@ -330,11 +338,14 @@ def involutionNoFixedZeroSetEquivEquiv {n : ℕ} ext i simp -def involutionNoFixedZeroSetEquivSetEquiv {n : ℕ} (k : Fin (2 * n + 1)) - (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive (e.symm ∘ f ∘ e) ∧ - (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} ≃ - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ +/-- The condition on fixed point free involutions of `Fin n.succ` for a fixed value of `f 0` + given an equivalence `e`, + can be modified so that only the condition on `f 0` is up-to the equivalence `e`. -/ +def involutionNoFixedZeroSetEquivSetEquiv {n : ℕ} (k : Fin n) + (e : Fin n.succ ≃ Fin n.succ) : + {f : Fin n.succ → Fin n.succ // Function.Involutive (e.symm ∘ f ∘ e) ∧ + (∀ i, (e.symm ∘ f ∘ e) i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} ≃ + {f : Fin n.succ → Fin n.succ // Function.Involutive f ∧ (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} := by refine Equiv.subtypeEquivRight ?_ intro f @@ -359,21 +370,24 @@ def involutionNoFixedZeroSetEquivSetEquiv {n : ℕ} (k : Fin (2 * n + 1)) nth_rewrite 2 [← hn] at hi simp at hi -def involutionNoFixedZeroSetEquivEquiv' {n : ℕ} (k : Fin (2 * n + 1)) - (e : Fin (2 * n.succ) ≃ Fin (2 * n.succ)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} - ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f (e 0) = e k.succ} := by +/-- Fixed point free involutions of `Fin n.succ` fixing `(e.symm ∘ f ∘ e) = k.succ` for a given `e` + are equivalent to fixing `f (e 0) = e k.succ`. -/ +def involutionNoFixedZeroSetEquivEquiv' {n : ℕ} (k : Fin n) (e : Fin n.succ ≃ Fin n.succ) : + {f : Fin n.succ → Fin n.succ // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ (e.symm ∘ f ∘ e) 0 = k.succ} ≃ + {f : Fin n.succ → Fin n.succ // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f (e 0) = e k.succ} := by refine Equiv.subtypeEquivRight ?_ simp only [succ_eq_add_one, ne_eq, Function.comp_apply, and_congr_right_iff] intro f hi h1 exact Equiv.symm_apply_eq e -def involutionNoFixedZeroSetEquivSetOne {n : ℕ} (k : Fin (2 * n + 1)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ +/-- Fixed point involutions of `Fin n.succ.succ` with `f 0 = k.succ` are equivalent + to fixed point involutions with `f 0 = 1`. -/ +def involutionNoFixedZeroSetEquivSetOne {n : ℕ} (k : Fin n.succ) : + {f : Fin n.succ.succ → Fin n.succ.succ // Function.Involutive f ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} - ≃ {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ + ≃ {f : Fin n.succ.succ → Fin n.succ.succ // Function.Involutive f ∧ (∀ i, f i ≠ i) ∧ f 0 = 1} := by refine Equiv.trans (involutionNoFixedZeroSetEquivEquiv k (Equiv.swap k.succ 1)) ?_ refine Equiv.trans (involutionNoFixedZeroSetEquivSetEquiv k (Equiv.swap k.succ 1)) ?_ @@ -385,9 +399,11 @@ def involutionNoFixedZeroSetEquivSetOne {n : ℕ} (k : Fin (2 * n + 1)) : · exact Ne.symm (Fin.succ_ne_zero k) · exact Fin.zero_ne_one +/-- Fixed point involutions of `Fin n.succ.succ` fixing `f 0 = 1` are equivalent to + fixed point involutions of `Fin n`. -/ def involutionNoFixedSetOne {n : ℕ} : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ - (∀ i, f i ≠ i) ∧ f 0 = 1} ≃ {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ + {f : Fin n.succ.succ → Fin n.succ.succ // Function.Involutive f ∧ + (∀ i, f i ≠ i) ∧ f 0 = 1} ≃ {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} where toFun f := by have hf1 : f.1 1 = 0 := by @@ -395,14 +411,14 @@ def involutionNoFixedSetOne {n : ℕ} : simp only [succ_eq_add_one, ne_eq, ← hf] rw [f.2.1] let f' := f.1 ∘ Fin.succ ∘ Fin.succ - have hf' (i : Fin (2 * n)) : f' i ≠ 0 := by + have hf' (i : Fin n) : f' i ≠ 0 := by simp only [succ_eq_add_one, mul_eq, ne_eq, Function.comp_apply, f'] simp only [← hf1, succ_eq_add_one, ne_eq] by_contra hn have hn' := Function.Involutive.injective f.2.1 hn simp [Fin.ext_iff] at hn' let f'' := fun i => (f' i).pred (hf' i) - have hf'' (i : Fin (2 * n)) : f'' i ≠ 0 := by + have hf'' (i : Fin n) : f'' i ≠ 0 := by simp only [mul_eq, ne_eq, f''] rw [@Fin.pred_eq_iff_eq_succ] simp only [mul_eq, succ_eq_add_one, ne_eq, Function.comp_apply, Fin.succ_zero_eq_one, f'] @@ -420,7 +436,7 @@ def involutionNoFixedSetOne {n : ℕ} : rw [Fin.pred_eq_iff_eq_succ, Fin.pred_eq_iff_eq_succ] exact f.2.2.1 i.succ.succ invFun f := by - let f' := fun (i : Fin (2 * n.succ))=> + let f' := fun (i : Fin n.succ.succ)=> match i with | ⟨0, h⟩ => 1 | ⟨1, h⟩ => 0 @@ -441,7 +457,7 @@ def involutionNoFixedSetOne {n : ℕ} : · rename_i h rename_i x r simp_all only [succ_eq_add_one, Fin.ext_iff, Fin.val_succ, add_left_inj] - have hfn {a b : ℕ} {ha : a < 2 * n} {hb : b < 2 * n} + have hfn {a b : ℕ} {ha : a < n} {hb : b < n} (hab : ↑(f.1 ⟨a, ha⟩) = b) : ↑(f.1 ⟨b, hb⟩) = a := by have ht : f.1 ⟨a, ha⟩ = ⟨b, hb⟩ := by simp [hab, Fin.ext_iff] @@ -511,26 +527,33 @@ def involutionNoFixedSetOne {n : ℕ} : apply congrArg simp_all [Fin.ext_iff] -def involutionNoFixedZeroSetEquiv {n : ℕ} (k : Fin (2 * n + 1)) : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ +/-- Fixed point involutions of `Fin n.succ.succ` for fixed `f 0 = k.succ` are + equivalent to fixed point involutions of `Fin n`. -/ +def involutionNoFixedZeroSetEquiv {n : ℕ} (k : Fin n.succ) : + {f : Fin n.succ.succ → Fin n.succ.succ // Function.Involutive f ∧ (∀ i, f i ≠ i) ∧ f 0 = k.succ} - ≃ {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + ≃ {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by refine Equiv.trans (involutionNoFixedZeroSetEquivSetOne k) involutionNoFixedSetOne +/-- The type of fixed point free involutions of `Fin n.succ.succ` is equivalent to the sum + of `Fin n.succ` copies of fixed point involutions of `Fin n`. -/ def involutionNoFixedEquivSumSame {n : ℕ} : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} - ≃ Σ (_ : Fin (2 * n + 1)), - {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + {f : Fin n.succ.succ → Fin n.succ.succ // Function.Involutive f ∧ (∀ i, f i ≠ i)} + ≃ Σ (_ : Fin n.succ), + {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by refine Equiv.trans involutionNoFixedEquivSum ?_ refine Equiv.sigmaCongrRight involutionNoFixedZeroSetEquiv +/-- Ever fixed-point free involutions of `Fin n.succ.succ` can be decomponsed into a + element of `Fin n.succ` (where `0` is sent) and a fixed-point free involution of + `Fin n`. -/ def involutionNoFixedZeroEquivProd {n : ℕ} : - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} - ≃ Fin (2 * n + 1) × - {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + {f : Fin n.succ.succ → Fin n.succ.succ // Function.Involutive f ∧ (∀ i, f i ≠ i)} + ≃ Fin n.succ × + {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by refine Equiv.trans involutionNoFixedEquivSumSame ?_ - exact Equiv.sigmaEquivProd (Fin (2 * n + 1)) - { f // Function.Involutive f ∧ ∀ (i : Fin (2 * n)), f i ≠ i} + exact Equiv.sigmaEquivProd (Fin n.succ) + { f // Function.Involutive f ∧ ∀ (i : Fin n), f i ≠ i} /-! @@ -538,6 +561,7 @@ def involutionNoFixedZeroEquivProd {n : ℕ} : -/ +/-- The type of fixed-point free involutions of `Fin n` is finite. -/ instance {n : ℕ} : Fintype { f // Function.Involutive f ∧ ∀ (i : Fin n), f i ≠ i } := by haveI : DecidablePred fun x => Function.Involutive x := fun f => Fintype.decidableForallFintype (α := Fin n) @@ -547,22 +571,31 @@ instance {n : ℕ} : Fintype { f // Function.Involutive f ∧ ∀ (i : Fin n), f lemma involutionNoFixed_card_succ {n : ℕ} : Fintype.card - {f : Fin (2 * n.succ) → Fin (2 * n.succ) // Function.Involutive f ∧ (∀ i, f i ≠ i)} - = (2 * n + 1) * - Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by + {f : Fin n.succ.succ → Fin n.succ.succ // Function.Involutive f ∧ (∀ i, f i ≠ i)} + = n.succ * + Fintype.card {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} := by rw [Fintype.card_congr (involutionNoFixedZeroEquivProd), Fintype.card_prod] congr - exact Fintype.card_fin (2 * n + 1) + exact Fintype.card_fin n.succ lemma involutionNoFixed_card_mul_two : (n : ℕ) → Fintype.card {f : Fin (2 * n) → Fin (2 * n) // Function.Involutive f ∧ (∀ i, f i ≠ i)} = (2 * n - 1)‼ | 0 => rfl | Nat.succ n => by - rw [involutionNoFixed_card_succ] - rw [involutionNoFixed_card_mul_two n] + erw [involutionNoFixed_card_succ] + erw [involutionNoFixed_card_mul_two n] exact Eq.symm (Nat.doubleFactorial_add_one (Nat.mul 2 n)) +lemma involutionNoFixed_card_mul_two_plus_one : (n : ℕ) → + Fintype.card {f : Fin (2 * n + 1) → Fin (2 * n + 1) // Function.Involutive f ∧ (∀ i, f i ≠ i)} + = 0 + | 0 => rfl + | Nat.succ n => by + erw [involutionNoFixed_card_succ] + erw [involutionNoFixed_card_mul_two_plus_one n] + exact rfl + lemma involutionNoFixed_card_even : (n : ℕ) → (he : Even n) → Fintype.card {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} = (n - 1)‼ := by intro n he @@ -571,4 +604,11 @@ lemma involutionNoFixed_card_even : (n : ℕ) → (he : Even n) → subst hr' exact involutionNoFixed_card_mul_two r +lemma involutionNoFixed_card_odd : (n : ℕ) → (ho : Odd n) → + Fintype.card {f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} = 0 := by + intro n ho + obtain ⟨r, hr⟩ := ho + subst hr + exact involutionNoFixed_card_mul_two_plus_one r + end HepLean.Fin diff --git a/HepLean/PerturbationTheory/Contractions/Basic.lean b/HepLean/PerturbationTheory/Contractions/Basic.lean index 57faf247..d591a969 100644 --- a/HepLean/PerturbationTheory/Contractions/Basic.lean +++ b/HepLean/PerturbationTheory/Contractions/Basic.lean @@ -33,11 +33,13 @@ namespace Contractions variable {l : List 𝓕} (c : Contractions l) -def auxCongr : {φs: List 𝓕} → {φsᵤₙ φsᵤₙ' : List 𝓕} → (h : φsᵤₙ = φsᵤₙ') → +/-- The equivalence between `ContractionsAux` based on the propositionally equivalent + uncontracted lists. -/ +def auxCongr : {φs : List 𝓕} → {φsᵤₙ φsᵤₙ' : List 𝓕} → (h : φsᵤₙ = φsᵤₙ') → ContractionsAux φs φsᵤₙ ≃ ContractionsAux φs φsᵤₙ' | _, _, _, Eq.refl _ => Equiv.refl _ -lemma auxCongr_ext {φs: List 𝓕} {c c2 : Contractions φs} (h : c.1 = c2.1) +lemma auxCongr_ext {φs : List 𝓕} {c c2 : Contractions φs} (h : c.1 = c2.1) (hx : c.2 = auxCongr h.symm c2.2) : c = c2 := by cases c cases c2 @@ -76,6 +78,7 @@ lemma contractions_nil (a : Contractions ([] : List 𝓕)) : a = ⟨[], Contract cases c rfl +/-- The embedding of the uncontracted fields into all fields. -/ def embedUncontracted {l : List 𝓕} (c : Contractions l) : Fin c.uncontracted.length → Fin l.length := match l, c with diff --git a/HepLean/PerturbationTheory/Contractions/Card.lean b/HepLean/PerturbationTheory/Contractions/Card.lean index 03a9ef3f..cb0ea0b4 100644 --- a/HepLean/PerturbationTheory/Contractions/Card.lean +++ b/HepLean/PerturbationTheory/Contractions/Card.lean @@ -25,15 +25,8 @@ lemma card_of_full_contractions_even {φs : List 𝓕} (he : Even φs.length) : /-- There are no full contractions of a list with an odd number of fields. -/ lemma card_of_full_contractions_odd {φs : List 𝓕} (ho : Odd φs.length) : Fintype.card {c : Contractions φs // IsFull c} = 0 := by - rw [Fintype.card_eq_zero_iff, isEmpty_subtype] - intro c - simp only [IsFull] - by_contra hn - have hc := uncontracted_length_even_iff c - rw [hn] at hc - simp only [List.length_nil, even_zero, iff_true] at hc - rw [← Nat.not_odd_iff_even] at hc - exact hc ho + rw [Fintype.card_congr (isFullInvolutionEquiv (φs := φs))] + exact involutionNoFixed_card_odd φs.length ho end Contractions diff --git a/HepLean/PerturbationTheory/Contractions/Involutions.lean b/HepLean/PerturbationTheory/Contractions/Involutions.lean index 65b4d4b2..ac7ebf09 100644 --- a/HepLean/PerturbationTheory/Contractions/Involutions.lean +++ b/HepLean/PerturbationTheory/Contractions/Involutions.lean @@ -183,7 +183,7 @@ def fromInvolution {φs : List 𝓕} ⟨uncontractedFromInvolution f, fromInvolutionAux f⟩ lemma fromInvolution_cons {φs : List 𝓕} {φ : 𝓕} - (f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) : + (f : {f : Fin (φ :: φs).length → Fin (φ :: φs).length // Function.Involutive f}) : let f' := involutionCons φs.length f fromInvolution f = consEquiv.symm ⟨fromInvolution f'.1, Option.map (finCongr ((uncontractedFromInvolution f'.fst).2.symm))