refactor: free simps

This commit is contained in:
jstoobysmith 2025-01-05 16:46:15 +00:00
parent 9184f6087c
commit 7026d8ce66
7 changed files with 150 additions and 146 deletions

View file

@ -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

View file

@ -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]

View file

@ -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

View file

@ -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

View file

@ -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

View file

@ -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

View file

@ -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