refactor: more multiple-goals

This commit is contained in:
jstoobysmith 2024-08-20 15:27:45 -04:00
parent 01f9f7da8b
commit b9479c904d
4 changed files with 166 additions and 180 deletions

View file

@ -179,36 +179,36 @@ lemma basis_on_δ₁_other {k j : Fin n} (h : k ≠ j) :
simp [basisAsCharges] simp [basisAsCharges]
simp [δ₁, δ₂] simp [δ₁, δ₂]
split split
rename_i h1 · rename_i h1
rw [Fin.ext_iff] at h1 rw [Fin.ext_iff] at h1
simp_all simp_all
rw [Fin.ext_iff] at h rw [Fin.ext_iff] at h
simp_all simp_all
split · split
rename_i h1 h2 · rename_i h1 h2
simp_all simp_all
rw [Fin.ext_iff] at h2 rw [Fin.ext_iff] at h2
simp at h2 simp at h2
omega omega
rfl · rfl
lemma basis!_on_δ!₁_other {k j : Fin n} (h : k ≠ j) : lemma basis!_on_δ!₁_other {k j : Fin n} (h : k ≠ j) :
basis!AsCharges k (δ!₁ j) = 0 := by basis!AsCharges k (δ!₁ j) = 0 := by
simp [basis!AsCharges] simp [basis!AsCharges]
simp [δ!₁, δ!₂] simp [δ!₁, δ!₂]
split split
rename_i h1 · rename_i h1
rw [Fin.ext_iff] at h1 rw [Fin.ext_iff] at h1
simp_all simp_all
rw [Fin.ext_iff] at h rw [Fin.ext_iff] at h
simp_all simp_all
split · split
rename_i h1 h2 · rename_i h1 h2
simp_all simp_all
rw [Fin.ext_iff] at h2 rw [Fin.ext_iff] at h2
simp at h2 simp at h2
omega omega
rfl rfl
lemma basis_on_other {k : Fin n} {j : Fin (2 * n + 1)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) : lemma basis_on_other {k : Fin n} {j : Fin (2 * n + 1)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) :
basisAsCharges k j = 0 := by basisAsCharges k j = 0 := by
@ -245,8 +245,8 @@ lemma basis!_δ!₂_eq_minus_δ!₁ (j i : Fin n) :
all_goals rename_i h1 h2 all_goals rename_i h1 h2
all_goals rw [Fin.ext_iff] at h1 h2 all_goals rw [Fin.ext_iff] at h1 h2
all_goals simp_all all_goals simp_all
subst h1 · subst h1
exact Fin.elim0 i exact Fin.elim0 i
all_goals rename_i h3 all_goals rename_i h3
all_goals rw [Fin.ext_iff] at h3 all_goals rw [Fin.ext_iff] at h3
all_goals simp_all all_goals simp_all
@ -269,26 +269,26 @@ lemma basis!_on_δ!₂_other {k j : Fin n} (h : k ≠ j) : basis!AsCharges k (δ
lemma basis_on_δ₃ (j : Fin n) : basisAsCharges j δ₃ = 0 := by lemma basis_on_δ₃ (j : Fin n) : basisAsCharges j δ₃ = 0 := by
simp [basisAsCharges] simp [basisAsCharges]
split <;> rename_i h split <;> rename_i h
rw [Fin.ext_iff] at h · rw [Fin.ext_iff] at h
simp [δ₃, δ₁] at h simp [δ₃, δ₁] at h
omega omega
split <;> rename_i h2 · split <;> rename_i h2
rw [Fin.ext_iff] at h2 · rw [Fin.ext_iff] at h2
simp [δ₃, δ₂] at h2 simp [δ₃, δ₂] at h2
omega omega
rfl · rfl
lemma basis!_on_δ!₃ (j : Fin n) : basis!AsCharges j δ!₃ = 0 := by lemma basis!_on_δ!₃ (j : Fin n) : basis!AsCharges j δ!₃ = 0 := by
simp [basis!AsCharges] simp [basis!AsCharges]
split <;> rename_i h split <;> rename_i h
rw [Fin.ext_iff] at h · rw [Fin.ext_iff] at h
simp [δ!₃, δ!₁] at h simp [δ!₃, δ!₁] at h
omega omega
split <;> rename_i h2 · split <;> rename_i h2
rw [Fin.ext_iff] at h2 · rw [Fin.ext_iff] at h2
simp [δ!₃, δ!₂] at h2 simp [δ!₃, δ!₂] at h2
omega omega
rfl · rfl
lemma basis_linearACC (j : Fin n) : (accGrav (2 * n + 1)) (basisAsCharges j) = 0 := by lemma basis_linearACC (j : Fin n) : (accGrav (2 * n + 1)) (basisAsCharges j) = 0 := by
rw [accGrav] rw [accGrav]
@ -338,16 +338,16 @@ lemma swap!_as_add {S S' : (PureU1 (2 * n + 1)).LinSols} (j : Fin n)
funext i funext i
rw [← hS, FamilyPermutations_anomalyFreeLinear_apply] rw [← hS, FamilyPermutations_anomalyFreeLinear_apply]
by_cases hi : i = δ!₁ j by_cases hi : i = δ!₁ j
subst hi · subst hi
simp [HSMul.hSMul, basis!_on_δ!₁_self, pairSwap_inv_fst] simp [HSMul.hSMul, basis!_on_δ!₁_self, pairSwap_inv_fst]
by_cases hi2 : i = δ!₂ j · by_cases hi2 : i = δ!₂ j
subst hi2 · subst hi2
simp [HSMul.hSMul,basis!_on_δ!₂_self, pairSwap_inv_snd] simp [HSMul.hSMul,basis!_on_δ!₂_self, pairSwap_inv_snd]
simp [HSMul.hSMul] · simp [HSMul.hSMul]
rw [basis!_on_other hi hi2] rw [basis!_on_other hi hi2]
change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_ change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_
erw [pairSwap_inv_other (Ne.symm hi) (Ne.symm hi2)] erw [pairSwap_inv_other (Ne.symm hi) (Ne.symm hi2)]
simp simp
/-- A point in the span of the first part of the basis as a charge. -/ /-- A point in the span of the first part of the basis as a charge. -/
def P (f : Fin n → ) : (PureU1 (2 * n + 1)).Charges := ∑ i, f i • basisAsCharges i def P (f : Fin n → ) : (PureU1 (2 * n + 1)).Charges := ∑ i, f i • basisAsCharges i
@ -362,45 +362,45 @@ lemma P_δ₁ (f : Fin n → ) (j : Fin n) : P f (δ₁ j) = f j := by
rw [P, sum_of_charges] rw [P, sum_of_charges]
simp [HSMul.hSMul, SMul.smul] simp [HSMul.hSMul, SMul.smul]
rw [Finset.sum_eq_single j] rw [Finset.sum_eq_single j]
rw [basis_on_δ₁_self] · rw [basis_on_δ₁_self]
simp only [mul_one] simp only [mul_one]
intro k _ hkj · intro k _ hkj
rw [basis_on_δ₁_other hkj] rw [basis_on_δ₁_other hkj]
simp only [mul_zero] simp only [mul_zero]
simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff] · simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff]
lemma P!_δ!₁ (f : Fin n → ) (j : Fin n) : P! f (δ!₁ j) = f j := by lemma P!_δ!₁ (f : Fin n → ) (j : Fin n) : P! f (δ!₁ j) = f j := by
rw [P!, sum_of_charges] rw [P!, sum_of_charges]
simp [HSMul.hSMul, SMul.smul] simp [HSMul.hSMul, SMul.smul]
rw [Finset.sum_eq_single j] rw [Finset.sum_eq_single j]
rw [basis!_on_δ!₁_self] · rw [basis!_on_δ!₁_self]
simp only [mul_one] simp only [mul_one]
intro k _ hkj · intro k _ hkj
rw [basis!_on_δ!₁_other hkj] rw [basis!_on_δ!₁_other hkj]
simp only [mul_zero] simp only [mul_zero]
simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff] · simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff]
lemma P_δ₂ (f : Fin n → ) (j : Fin n) : P f (δ₂ j) = - f j := by lemma P_δ₂ (f : Fin n → ) (j : Fin n) : P f (δ₂ j) = - f j := by
rw [P, sum_of_charges] rw [P, sum_of_charges]
simp [HSMul.hSMul, SMul.smul] simp [HSMul.hSMul, SMul.smul]
rw [Finset.sum_eq_single j] rw [Finset.sum_eq_single j]
rw [basis_on_δ₂_self] · rw [basis_on_δ₂_self]
simp only [mul_neg, mul_one] simp only [mul_neg, mul_one]
intro k _ hkj · intro k _ hkj
rw [basis_on_δ₂_other hkj] rw [basis_on_δ₂_other hkj]
simp only [mul_zero] simp only [mul_zero]
simp · simp
lemma P!_δ!₂ (f : Fin n → ) (j : Fin n) : P! f (δ!₂ j) = - f j := by lemma P!_δ!₂ (f : Fin n → ) (j : Fin n) : P! f (δ!₂ j) = - f j := by
rw [P!, sum_of_charges] rw [P!, sum_of_charges]
simp [HSMul.hSMul, SMul.smul] simp [HSMul.hSMul, SMul.smul]
rw [Finset.sum_eq_single j] rw [Finset.sum_eq_single j]
rw [basis!_on_δ!₂_self] · rw [basis!_on_δ!₂_self]
simp only [mul_neg, mul_one] simp only [mul_neg, mul_one]
intro k _ hkj · intro k _ hkj
rw [basis!_on_δ!₂_other hkj] rw [basis!_on_δ!₂_other hkj]
simp only [mul_zero] simp only [mul_zero]
simp · simp
lemma P_δ₃ (f : Fin n → ) : P f (δ₃) = 0 := by lemma P_δ₃ (f : Fin n → ) : P f (δ₃) = 0 := by
rw [P, sum_of_charges] rw [P, sum_of_charges]
@ -476,12 +476,12 @@ lemma P_P_P!_accCube (g : Fin n → ) (j : Fin n) :
rw [sum_δ!, basis!_on_δ!₃] rw [sum_δ!, basis!_on_δ!₃]
simp only [mul_zero, Function.comp_apply, zero_add] simp only [mul_zero, Function.comp_apply, zero_add]
rw [Finset.sum_eq_single j, basis!_on_δ!₁_self, basis!_on_δ!₂_self] rw [Finset.sum_eq_single j, basis!_on_δ!₁_self, basis!_on_δ!₂_self]
rw [← δ₂_δ!₂, P_δ₂] · rw [← δ₂_δ!₂, P_δ₂]
ring ring
intro k _ hkj · intro k _ hkj
erw [basis!_on_δ!₁_other hkj.symm, basis!_on_δ!₂_other hkj.symm] erw [basis!_on_δ!₁_other hkj.symm, basis!_on_δ!₂_other hkj.symm]
simp only [mul_zero, add_zero] simp only [mul_zero, add_zero]
simp · simp
lemma P_zero (f : Fin n → ) (h : P f = 0) : ∀ i, f i = 0 := by lemma P_zero (f : Fin n → ) (h : P f = 0) : ∀ i, f i = 0 := by
intro i intro i
@ -584,29 +584,27 @@ theorem basisa_linear_independent : LinearIndependent (@basisa n.succ) := by
have hf := Pa_zero (f ∘ Sum.inl) (f ∘ Sum.inr) h1 have hf := Pa_zero (f ∘ Sum.inl) (f ∘ Sum.inr) h1
have hg := Pa_zero! (f ∘ Sum.inl) (f ∘ Sum.inr) h1 have hg := Pa_zero! (f ∘ Sum.inl) (f ∘ Sum.inr) h1
intro i intro i
simp_all simp_all only [succ_eq_add_one, Function.comp_apply]
cases i cases i
simp_all · simp_all
simp_all · simp_all
lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n.succ) → ) : Pa' f = Pa' f' ↔ f = f' := by lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n.succ) → ) : Pa' f = Pa' f' ↔ f = f' := by
apply Iff.intro refine Iff.intro (fun h => ?_) (fun h => ?_)
intro h · funext i
funext i rw [Pa', Pa'] at h
rw [Pa', Pa'] at h have h1 : ∑ i : Fin n.succ ⊕ Fin n.succ, (f i + (- f' i)) • basisa i = 0 := by
have h1 : ∑ i : Fin n.succ ⊕ Fin n.succ, (f i + (- f' i)) • basisa i = 0 := by simp only [add_smul, neg_smul]
simp only [add_smul, neg_smul] rw [Finset.sum_add_distrib]
rw [Finset.sum_add_distrib] rw [h]
rw [h] rw [← Finset.sum_add_distrib]
rw [← Finset.sum_add_distrib] simp
simp have h2 : ∀ i, (f i + (- f' i)) = 0 := by
have h2 : ∀ i, (f i + (- f' i)) = 0 := by exact Fintype.linearIndependent_iff.mp (@basisa_linear_independent n)
exact Fintype.linearIndependent_iff.mp (@basisa_linear_independent n) (fun i => f i + -f' i) h1
(fun i => f i + -f' i) h1 have h2i := h2 i
have h2i := h2 i linarith
linarith · rw [h]
intro h
rw [h]
/-! TODO: Replace the definition of `join` with a Mathlib definition, most likely `Sum.elim`. -/ /-! TODO: Replace the definition of `join` with a Mathlib definition, most likely `Sum.elim`. -/
/-- A helper function for what follows. -/ /-- A helper function for what follows. -/
@ -617,28 +615,23 @@ def join (g f : Fin n → ) : Fin n ⊕ Fin n → := fun i =>
lemma join_ext (g g' : Fin n → ) (f f' : Fin n → ) : lemma join_ext (g g' : Fin n → ) (f f' : Fin n → ) :
join g f = join g' f' ↔ g = g' ∧ f = f' := by join g f = join g' f' ↔ g = g' ∧ f = f' := by
apply Iff.intro refine Iff.intro (fun h => ?_) (fun h => ?_)
intro h · apply And.intro
apply And.intro · funext i
funext i exact congr_fun h (Sum.inl i)
exact congr_fun h (Sum.inl i) · funext i
funext i exact congr_fun h (Sum.inr i)
exact congr_fun h (Sum.inr i) · rw [h.left, h.right]
intro h
rw [h.left, h.right]
lemma join_Pa (g g' : Fin n.succ → ) (f f' : Fin n.succ → ) : lemma join_Pa (g g' : Fin n.succ → ) (f f' : Fin n.succ → ) :
Pa' (join g f) = Pa' (join g' f') ↔ Pa g f = Pa g' f' := by Pa' (join g f) = Pa' (join g' f') ↔ Pa g f = Pa g' f' := by
apply Iff.intro refine Iff.intro (fun h => ?_) (fun h => ?_)
intro h · rw [Pa'_eq, join_ext] at h
rw [Pa'_eq] at h rw [h.left, h.right]
rw [join_ext] at h · apply ACCSystemLinear.LinSols.ext
rw [h.left, h.right] rw [Pa'_P'_P!', Pa'_P'_P!']
intro h simp only [succ_eq_add_one, ACCSystemLinear.linSolsAddCommMonoid_add_val, P'_val, P!'_val]
apply ACCSystemLinear.LinSols.ext exact h
rw [Pa'_P'_P!', Pa'_P'_P!']
simp [P'_val, P!'_val]
exact h
lemma Pa_eq (g g' : Fin n.succ → ) (f f' : Fin n.succ → ) : lemma Pa_eq (g g' : Fin n.succ → ) (f f' : Fin n.succ → ) :
Pa g f = Pa g' f' ↔ g = g' ∧ f = f' := by Pa g f = Pa g' f' ↔ g = g' ∧ f = f' := by

View file

@ -114,8 +114,8 @@ lemma generic_or_special (S : (PureU1 (2 * n.succ + 1)).Sols) :
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
exact ne_or_eq _ _ exact ne_or_eq _ _
cases h1 <;> rename_i h1 cases h1 <;> rename_i h1
exact Or.inl (genericCase_exists S ⟨g, f, h, h1⟩) · exact Or.inl (genericCase_exists S ⟨g, f, h, h1⟩)
exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩) · exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩)
theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : GenericCase S) : theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : GenericCase S) :
∃ g f a, S = parameterization g f a := by ∃ g f a, S = parameterization g f a := by
@ -127,11 +127,11 @@ theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : GenericCase S) :
change S.val = _ • (_ • P g + _• P! f) change S.val = _ • (_ • P g + _• P! f)
rw [anomalyFree_param _ _ hS] rw [anomalyFree_param _ _ hS]
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul] rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul]
exact hS · exact hS
have h := h g f hS · have h := h g f hS
rw [anomalyFree_param _ _ hS] at h rw [anomalyFree_param _ _ hS] at h
simp at h simp at h
exact h exact h
lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols} lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
(h : SpecialCase S) : (h : SpecialCase S) :

View file

@ -107,10 +107,10 @@ lemma cubic_zero_E'_zero (S : linearParameters) (hc : accCube (S.asCharges) = 0)
rw [h1] at hc rw [h1] at hc
simp at hc simp at hc
cases' hc with hc hc cases' hc with hc hc
have h2 := (add_eq_zero_iff' (by nlinarith) (sq_nonneg S.Y)).mp hc · have h2 := (add_eq_zero_iff' (by nlinarith) (sq_nonneg S.Y)).mp hc
simp at h2 simp at h2
exact h2.1 exact h2.1
exact hc · exact hc
/-- The bijection between the type of linear parameters and `(SMNoGrav 1).LinSols`. -/ /-- The bijection between the type of linear parameters and `(SMNoGrav 1).LinSols`. -/
def bijection : linearParameters ≃ (SMNoGrav 1).LinSols where def bijection : linearParameters ≃ (SMNoGrav 1).LinSols where
@ -120,13 +120,13 @@ def bijection : linearParameters ≃ (SMNoGrav 1).LinSols where
SMCharges.E S.val (0 : Fin 1)⟩ SMCharges.E S.val (0 : Fin 1)⟩
left_inv S := by left_inv S := by
apply linearParameters.ext apply linearParameters.ext
rfl · rfl
simp only [Fin.isValue] · simp only [Fin.isValue]
repeat erw [speciesVal] repeat erw [speciesVal]
simp only [asCharges, neg_add_rev] simp only [asCharges, neg_add_rev]
ring ring
simp only [toSpecies_apply] · simp only [toSpecies_apply]
rfl rfl
right_inv S := by right_inv S := by
simp only [Fin.isValue, toSpecies_apply] simp only [Fin.isValue, toSpecies_apply]
apply ACCSystemLinear.LinSols.ext apply ACCSystemLinear.LinSols.ext
@ -239,25 +239,25 @@ def bijectionLinearParameters :
have hvw := S.hvw have hvw := S.hvw
have hQ := S.hx have hQ := S.hx
apply linearParametersQENeqZero.ext apply linearParametersQENeqZero.ext
rfl · rfl
field_simp · field_simp
ring ring
simp only [tolinearParametersQNeqZero_w, toLinearParameters_coe_Y, toLinearParameters_coe_Q', · simp only [tolinearParametersQNeqZero_w, toLinearParameters_coe_Y, toLinearParameters_coe_Q',
toLinearParameters_coe_E'] toLinearParameters_coe_E']
field_simp field_simp
ring ring
right_inv S := by right_inv S := by
apply Subtype.ext apply Subtype.ext
have hQ := S.2.1 have hQ := S.2.1
have hE := S.2.2 have hE := S.2.2
apply linearParameters.ext apply linearParameters.ext
rfl · rfl
field_simp · field_simp
ring_nf ring_nf
field_simp [hQ, hE] field_simp [hQ, hE]
field_simp · field_simp
ring_nf ring_nf
field_simp [hQ, hE] field_simp [hQ, hE]
/-- The bijection between `linearParametersQENeqZero` and `LinSols` with `Q` and `E` non-zero. -/ /-- The bijection between `linearParametersQENeqZero` and `LinSols` with `Q` and `E` non-zero. -/
def bijection : linearParametersQENeqZero ≃ def bijection : linearParametersQENeqZero ≃
@ -332,23 +332,21 @@ lemma cube_w_v (S : linearParametersQENeqZero) (h : accCube (bijection S).1.val
(S.v = -1 ∧ S.w = 0) (S.v = 0 ∧ S.w = -1) := by (S.v = -1 ∧ S.w = 0) (S.v = 0 ∧ S.w = -1) := by
have h' := cubic_v_or_w_zero S h FLTThree have h' := cubic_v_or_w_zero S h FLTThree
cases' h' with hx hx cases' h' with hx hx
simp [hx] · simp [hx]
exact cubic_v_zero S h hx exact cubic_v_zero S h hx
simp [hx] · simp [hx]
exact cube_w_zero S h hx exact cube_w_zero S h hx
lemma grav (S : linearParametersQENeqZero) : accGrav (bijection S).1.val = 0 ↔ S.v + S.w = -1 := by lemma grav (S : linearParametersQENeqZero) : accGrav (bijection S).1.val = 0 ↔ S.v + S.w = -1 := by
erw [linearParameters.grav] erw [linearParameters.grav]
have hvw := S.hvw have hvw := S.hvw
have hQ := S.hx have hQ := S.hx
field_simp field_simp
apply Iff.intro refine Iff.intro (fun h => ?_) (fun h => ?_)
intro h · apply (mul_right_inj' hQ).mp
apply (mul_right_inj' hQ).mp linear_combination -1 * h / 6
linear_combination -1 * h / 6 · rw [h]
intro h exact Eq.symm (mul_neg_one (6 * S.x))
rw [h]
exact Eq.symm (mul_neg_one (6 * S.x))
lemma grav_of_cubic (S : linearParametersQENeqZero) (h : accCube (bijection S).1.val = 0) lemma grav_of_cubic (S : linearParametersQENeqZero) (h : accCube (bijection S).1.val = 0)
(FLTThree : FermatLastTheoremWith 3) : (FLTThree : FermatLastTheoremWith 3) :
@ -356,10 +354,10 @@ lemma grav_of_cubic (S : linearParametersQENeqZero) (h : accCube (bijection S).1
rw [grav] rw [grav]
have h' := cube_w_v S h FLTThree have h' := cube_w_v S h FLTThree
cases' h' with h h cases' h' with h h
rw [h.1, h.2] · rw [h.1, h.2]
exact Rat.add_zero (-1) exact Rat.add_zero (-1)
rw [h.1, h.2] · rw [h.1, h.2]
exact Rat.zero_add (-1) exact Rat.zero_add (-1)
end linearParametersQENeqZero end linearParametersQENeqZero

View file

@ -115,9 +115,7 @@ lemma all_dual_eq_getDual?_of_mem_withUniqueDual (i : Fin l.length) (h : i ∈ l
lemma some_eq_getDual?_of_withUniqueDual_iff (i j : Fin l.length) (h : i ∈ l.withUniqueDual) : lemma some_eq_getDual?_of_withUniqueDual_iff (i j : Fin l.length) (h : i ∈ l.withUniqueDual) :
l.AreDualInSelf i j ↔ some j = l.getDual? i := by l.AreDualInSelf i j ↔ some j = l.getDual? i := by
apply Iff.intro refine Iff.intro (fun h' => all_dual_eq_getDual?_of_mem_withUniqueDual l i h j h') (fun h' => ?_)
exact fun h' => all_dual_eq_getDual?_of_mem_withUniqueDual l i h j h'
intro h'
have hj : j = (l.getDual? i).get (mem_withUniqueDual_isSome l i h) := have hj : j = (l.getDual? i).get (mem_withUniqueDual_isSome l i h) :=
Eq.symm (Option.get_of_mem (mem_withUniqueDual_isSome l i h) (id (Eq.symm h'))) Eq.symm (Option.get_of_mem (mem_withUniqueDual_isSome l i h) (id (Eq.symm h')))
subst hj subst hj
@ -128,8 +126,8 @@ lemma eq_getDual?_get_of_withUniqueDual_iff (i j : Fin l.length) (h : i ∈ l.wi
l.AreDualInSelf i j ↔ j = (l.getDual? i).get (mem_withUniqueDual_isSome l i h) := by l.AreDualInSelf i j ↔ j = (l.getDual? i).get (mem_withUniqueDual_isSome l i h) := by
rw [l.some_eq_getDual?_of_withUniqueDual_iff i j h] rw [l.some_eq_getDual?_of_withUniqueDual_iff i j h]
refine Iff.intro (fun h' => ?_) (fun h' => ?_) refine Iff.intro (fun h' => ?_) (fun h' => ?_)
exact Eq.symm (Option.get_of_mem (mem_withUniqueDual_isSome l i h) (id (Eq.symm h'))) · exact Eq.symm (Option.get_of_mem (mem_withUniqueDual_isSome l i h) (id (Eq.symm h')))
simp [h'] · simp [h']
lemma eq_of_areDualInSelf_withUniqueDual {j k : Fin l.length} (i : l.withUniqueDual) lemma eq_of_areDualInSelf_withUniqueDual {j k : Fin l.length} (i : l.withUniqueDual)
(hj : l.AreDualInSelf i j) (hk : l.AreDualInSelf i k) : j = k := by (hj : l.AreDualInSelf i j) (hk : l.AreDualInSelf i k) : j = k := by
@ -198,9 +196,7 @@ lemma all_dual_eq_of_withUniqueDualInOther (i : Fin l.length)
lemma some_eq_getDualInOther?_of_withUniqueDualInOther_iff (i : Fin l.length) (j : Fin l2.length) lemma some_eq_getDualInOther?_of_withUniqueDualInOther_iff (i : Fin l.length) (j : Fin l2.length)
(h : i ∈ l.withUniqueDualInOther l2) : (h : i ∈ l.withUniqueDualInOther l2) :
l.AreDualInOther l2 i j ↔ some j = l.getDualInOther? l2 i := by l.AreDualInOther l2 i j ↔ some j = l.getDualInOther? l2 i := by
apply Iff.intro refine Iff.intro (fun h' => l.all_dual_eq_of_withUniqueDualInOther l2 i h j h') (fun h' => ?_)
exact fun h' => l.all_dual_eq_of_withUniqueDualInOther l2 i h j h'
intro h'
have hj : j = (l.getDualInOther? l2 i).get (mem_withUniqueDualInOther_isSome l l2 i h) := have hj : j = (l.getDualInOther? l2 i).get (mem_withUniqueDualInOther_isSome l l2 i h) :=
Eq.symm (Option.get_of_mem (mem_withUniqueDualInOther_isSome l l2 i h) (id (Eq.symm h'))) Eq.symm (Option.get_of_mem (mem_withUniqueDualInOther_isSome l l2 i h) (id (Eq.symm h')))
subst hj subst hj
@ -213,8 +209,8 @@ lemma eq_getDualInOther?_get_of_withUniqueDualInOther_iff (i : Fin l.length) (j
(mem_withUniqueDualInOther_isSome l l2 i h) := by (mem_withUniqueDualInOther_isSome l l2 i h) := by
rw [l.some_eq_getDualInOther?_of_withUniqueDualInOther_iff l2 i j h] rw [l.some_eq_getDualInOther?_of_withUniqueDualInOther_iff l2 i j h]
refine Iff.intro (fun h' => ?_) (fun h' => ?_) refine Iff.intro (fun h' => ?_) (fun h' => ?_)
exact Eq.symm (Option.get_of_mem (mem_withUniqueDualInOther_isSome l l2 i h) (id (Eq.symm h'))) · exact (Option.get_of_mem (mem_withUniqueDualInOther_isSome l l2 i h) (id (Eq.symm h'))).symm
simp [h'] · simp [h']
@[simp, nolint simpNF] @[simp, nolint simpNF]
lemma getDualInOther?_get_getDualInOther?_get_of_withUniqueDualInOther lemma getDualInOther?_get_getDualInOther?_get_of_withUniqueDualInOther
@ -284,8 +280,7 @@ lemma getDualInOther?_get_of_mem_withUniqueInOther_mem (i : Fin l.length)
simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome, simp only [withUniqueDualInOther, mem_withDual_iff_isSome, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.mem_filter, Finset.mem_univ, Option.isNone_iff_eq_none, mem_withInDualOther_iff_isSome, Finset.mem_filter, Finset.mem_univ,
getDualInOther?_getDualInOther?_get_isSome, true_and] getDualInOther?_getDualInOther?_get_isSome, true_and]
apply And.intro apply And.intro (getDual?_of_getDualInOther?_of_mem_withUniqueInOther_eq_none l l2 i h)
exact getDual?_of_getDualInOther?_of_mem_withUniqueInOther_eq_none l l2 i h
intro j hj intro j hj
simp only [h, getDualInOther?_getDualInOther?_get_of_withUniqueDualInOther, Option.some.injEq] simp only [h, getDualInOther?_getDualInOther?_get_of_withUniqueDualInOther, Option.some.injEq]
by_contra hn by_contra hn