refactor: countId (partial)

This commit is contained in:
jstoobysmith 2024-08-23 15:37:14 -04:00
parent 21bbad1e19
commit b4c6c3faac
7 changed files with 173 additions and 166 deletions

View file

@ -411,13 +411,24 @@ end append
/-!
## countP on id
## countId
-/
lemma countP_id_neq_zero (i : Fin l.length) :
l.val.countP (fun J => (l.val.get i).id = J.id) ≠ 0 := by
def countId (I : Index X) : :=
l.val.countP (fun J => I.id = J.id)
@[simp]
lemma countId_append (I : Index X) : (l ++ l2).countId I = l.countId I + l2.countId I := by
simp [countId]
lemma countId_eq_length_filter (I : Index X) :
l.countId I = (l.val.filter (fun J => I.id = J.id)).length := by
simp [countId]
rw [List.countP_eq_length_filter]
lemma countId_index_neq_zero (i : Fin l.length) : l.countId (l.val.get i) ≠ 0 := by
rw [countId_eq_length_filter]
by_contra hn
rw [List.length_eq_zero] at hn
have hm : l.val.get i ∈ List.filter (fun J => decide ((l.val.get i).id = J.id)) l.val := by
@ -425,6 +436,42 @@ lemma countP_id_neq_zero (i : Fin l.length) :
rw [hn] at hm
simp at hm
lemma countId_append_symm (I : Index X) : (l ++ l2).countId I = (l2 ++ l).countId I := by
simp
omega
lemma countId_eq_one_append_mem_right_self_eq_one {I : Index X} (hI : I ∈ l2.val)
(h : (l ++ l2).countId I = 1) : l2.countId I = 1 := by
simp at h
have hmem : I ∈ l2.val.filter (fun J => I.id = J.id) := by
simp [List.mem_filter, decide_True, and_true, hI]
have h1 : l2.countId I ≠ 0 := by
rw [countId_eq_length_filter]
by_contra hn
rw [@List.length_eq_zero] at hn
rw [hn] at hmem
simp at hmem
omega
lemma countId_eq_one_append_mem_right_other_eq_zero {I : Index X} (hI : I ∈ l2.val)
(h : (l ++ l2).countId I = 1) : l.countId I = 0 := by
simp at h
have hmem : I ∈ l2.val.filter (fun J => I.id = J.id) := by
simp [List.mem_filter, decide_True, and_true, hI]
have h1 : l2.countId I ≠ 0 := by
rw [countId_eq_length_filter]
by_contra hn
rw [@List.length_eq_zero] at hn
rw [hn] at hmem
simp at hmem
omega
/-!
## Filter id
-/
/-! TODO: Replace with Mathlib lemma. -/
lemma filter_sort_comm {n : } (s : Finset (Fin n)) (p : Fin n → Prop) [DecidablePred p] :
List.filter p (Finset.sort (fun i j => i ≤ j) s) =

View file

@ -76,8 +76,8 @@ lemma color_quot_filter_of_countP_two (hl : l.withUniqueDual = l.withDual) (i :
(l.colorMap i = l.colorMap ((l.getDual? i).get hi)
l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))) := by
have hi1 := hi
rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countP] at hi1
rcases l.filter_id_of_countP_two hi1 with hf | hf
rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi1
rcases l.filter_id_of_countId_eq_two hi1 with hf | hf
all_goals
erw [hf]
simp [List.countP, List.countP.go]
@ -103,8 +103,8 @@ lemma color_dual_eq_self_filter_of_countP_two (hl : l.withUniqueDual = l.withDua
↔ l.colorMap i = 𝓒.τ (l.colorMap ((l.getDual? i).get hi))
(l.colorMap i) = 𝓒.τ (l.colorMap i) := by
have hi1 := hi
rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countP] at hi1
rcases l.filter_id_of_countP_two hi1 with hf | hf
rw [← mem_withDual_iff_isSome, ← hl, mem_withUniqueDual_iff_countId_eq_two] at hi1
rcases l.filter_id_of_countId_eq_two hi1 with hf | hf
all_goals
erw [hf]
simp [List.countP, List.countP.go]
@ -216,7 +216,7 @@ lemma iff_countP (hl : l.withUniqueDual = l.withDual) :
rw [iff_countP_isSome hl]
refine Iff.intro (fun h i hi => ?_) (fun h i hi => ?_)
· rw [← List.countP_eq_length_filter] at hi
rw [← @mem_withUniqueDual_iff_countP] at hi
rw [← mem_withUniqueDual_iff_countId_eq_two] at hi
exact h i (mem_withUniqueDual_isSome l i hi)
· rw [← @mem_withDual_iff_isSome, ← hl] at hi
rw [mem_withUniqueDual_iff_countP, List.countP_eq_length_filter] at hi

View file

@ -100,6 +100,7 @@ lemma equivalence : Equivalence (@Reindexing 𝓒 _) where
symm := symm
trans := trans
/-! TODO: Prove that `Reindexing l l2` implies `Reindexing l.contr l2.contr`. -/
end Reindexing
/-!

View file

@ -69,15 +69,15 @@ lemma withDual_union_withoutDual : l.withDual l.withoutDual = Finset.univ :=
· simp at h
simp [withoutDual, Finset.mem_filter, Finset.mem_univ, h]
lemma mem_withoutDual_iff_countP (i : Fin l.length) :
i ∈ l.withoutDual ↔ l.val.countP (fun j => (l.val.get i).id = j.id) = 1 := by
lemma mem_withoutDual_iff_countId_eq_one (i : Fin l.length) :
i ∈ l.withoutDual ↔ l.countId (l.val.get i) = 1 := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· exact countP_of_not_mem_withDual l i (l.not_mem_withDual_of_mem_withoutDual i h)
· exact countId_of_not_mem_withDual l i (l.not_mem_withDual_of_mem_withoutDual i h)
· by_contra hn
have h : i ∈ l.withDual := by
simp [withoutDual] at hn
simpa using Option.ne_none_iff_isSome.mp hn
rw [mem_withDual_iff_countP] at h
rw [mem_withDual_iff_countId_gt_one] at h
omega
/-- An equivalence from `Fin l.withoutDual.card` to `l.withoutDual` determined by
@ -110,40 +110,6 @@ def dualEquiv : l.withDual ⊕ Fin l.withoutDual.card ≃ Fin l.length :=
/-!
## Some lemmas involvoing countP = 1
-/
lemma countP_eq_one_append_mem_right_self_eq_one {I : Index X} (hI : I ∈ l2.val)
(h : (l ++ l2).val.countP (fun J => I.id = J.id) = 1) :
l2.val.countP (fun J => I.id = J.id) = 1 := by
simp at h
have hmem : I ∈ l2.val.filter (fun J => I.id = J.id) := by
simp [List.mem_filter, decide_True, and_true, hI]
have h1 : l2.val.countP (fun J => I.id = J.id) ≠ 0 := by
rw [List.countP_eq_length_filter]
by_contra hn
rw [@List.length_eq_zero] at hn
rw [hn] at hmem
simp at hmem
omega
lemma countP_eq_one_append_mem_right_other_eq_zero {I : Index X} (hI : I ∈ l2.val)
(h : (l ++ l2).val.countP (fun J => I.id = J.id) = 1) :
l.val.countP (fun J => I.id = J.id) = 0 := by
simp at h
have hmem : I ∈ l2.val.filter (fun J => I.id = J.id) := by
simp [List.mem_filter, decide_True, and_true, hI]
have h1 : l2.val.countP (fun J => I.id = J.id) ≠ 0 := by
rw [List.countP_eq_length_filter]
by_contra hn
rw [@List.length_eq_zero] at hn
rw [hn] at hmem
simp at hmem
omega
/-!
## The contraction list
-/
@ -151,7 +117,7 @@ lemma countP_eq_one_append_mem_right_other_eq_zero {I : Index X} (hI : I ∈ l2.
/-- The index list formed from `l` by selecting only those indices in `l` which
do not have a dual. -/
def contrIndexList : IndexList X where
val := l.val.filter (fun i => l.val.countP (fun j => i.id = j.id) = 1)
val := l.val.filter (fun I => l.countId I = 1)
@[simp]
lemma contrIndexList_empty : (⟨[]⟩ : IndexList X).contrIndexList = { val := [] } := by
@ -159,7 +125,7 @@ lemma contrIndexList_empty : (⟨[]⟩ : IndexList X).contrIndexList = { val :=
simp [contrIndexList]
lemma contrIndexList_val : l.contrIndexList.val =
l.val.filter (fun i => l.val.countP (fun j => i.id = j.id) = 1) := by
l.val.filter (fun I => l.countId I = 1) := by
rfl
/-- An alternative form of the contracted index list. -/
@ -173,12 +139,13 @@ lemma contrIndexList_eq_contrIndexList' : l.contrIndexList = l.contrIndexList' :
trans List.map l.val.get (List.ofFn (Subtype.val ∘ ⇑l.withoutDualEquiv))
· rw [list_ofFn_withoutDualEquiv_eq_sort, withoutDual_sort_eq_filter]
simp only [contrIndexList]
let f1 : Index X → Bool := fun (i : Index X) => l.val.countP (fun j => i.id = j.id) = 1
let f1 : Index X → Bool := fun (I : Index X) => l.countId I = 1
let f2 : (Fin l.length) → Bool := fun i => i ∈ l.withoutDual
change List.filter f1 l.val = List.map l.val.get (List.filter f2 (List.finRange l.length))
have hf : f2 = f1 ∘ l.val.get := by
funext i
simp only [mem_withoutDual_iff_countP l, List.get_eq_getElem, Function.comp_apply, f2, f1]
simp only [mem_withoutDual_iff_countId_eq_one l, List.get_eq_getElem, Function.comp_apply, f2,
f1]
rw [hf, ← List.filter_map]
apply congrArg
simp [length]
@ -277,8 +244,8 @@ lemma contrIndexList_getDualInOther?_self (i : Fin l.contrIndexList.length) :
exact (contrIndexList_areDualInSelf_false l i j).mp h
exact Fin.ge_of_eq (id (Eq.symm h1))
lemma mem_contrIndexList_count {I : Index X} (h : I ∈ l.contrIndexList.val) :
l.val.countP (fun J => I.id = J.id) = 1 := by
lemma mem_contrIndexList_countId {I : Index X} (h : I ∈ l.contrIndexList.val) :
l.countId I = 1 := by
simp only [contrIndexList, List.mem_filter, decide_eq_true_eq] at h
exact h.2
@ -286,7 +253,7 @@ lemma mem_contrIndexList_filter {I : Index X} (h : I ∈ l.contrIndexList.val) :
l.val.filter (fun J => I.id = J.id) = [I] := by
have h1 : (l.val.filter (fun J => I.id = J.id)).length = 1 := by
rw [← List.countP_eq_length_filter]
exact l.mem_contrIndexList_count h
exact l.mem_contrIndexList_countId h
have h2 : I ∈ l.val.filter (fun J => I.id = J.id) := by
simp only [List.mem_filter, decide_True, and_true]
simp only [contrIndexList, List.mem_filter, decide_eq_true_eq] at h
@ -298,17 +265,17 @@ lemma mem_contrIndexList_filter {I : Index X} (h : I ∈ l.contrIndexList.val) :
subst h2
exact hJ
lemma cons_contrIndexList_of_countP_eq_zero {I : Index X}
(h : l.val.countP (fun J => I.id = J.id) = 0) :
lemma cons_contrIndexList_of_countId_eq_zero {I : Index X}
(h : l.countId I = 0) :
(l.cons I).contrIndexList = l.contrIndexList.cons I := by
apply ext
simp [contrIndexList]
simp [contrIndexList, countId]
rw [List.filter_cons_of_pos]
· apply congrArg
apply List.filter_congr
intro J hJ
simp only [decide_eq_decide]
rw [List.countP_eq_zero] at h
rw [countId, List.countP_eq_zero] at h
simp only [decide_eq_true_eq] at h
rw [List.countP_cons_of_neg]
simp only [decide_eq_true_eq]
@ -316,18 +283,17 @@ lemma cons_contrIndexList_of_countP_eq_zero {I : Index X}
· simp only [decide_True, List.countP_cons_of_pos, add_left_eq_self, decide_eq_true_eq]
exact h
lemma cons_contrIndexList_of_countP_neq_zero {I : Index X}
(h : l.val.countP (fun J => I.id = J.id) ≠ 0) :
lemma cons_contrIndexList_of_countId_neq_zero {I : Index X} (h : l.countId I ≠ 0) :
(l.cons I).contrIndexList.val = l.contrIndexList.val.filter (fun J => ¬ I.id = J.id) := by
simp only [contrIndexList, cons_val, decide_not, List.filter_filter, Bool.not_eq_true',
simp only [contrIndexList, countId, cons_val, decide_not, List.filter_filter, Bool.not_eq_true',
decide_eq_false_iff_not, decide_eq_true_eq, Bool.decide_and]
rw [List.filter_cons_of_neg]
· apply List.filter_congr
intro J hJ
by_cases hI : I.id = J.id
· simp only [hI, decide_True, List.countP_cons_of_pos, add_left_eq_self, Bool.not_true,
Bool.false_and, decide_eq_false_iff_not]
rw [hI] at h
Bool.false_and, decide_eq_false_iff_not, countId]
rw [countId, hI] at h
simp only [h, not_false_eq_true]
· simp only [hI, decide_False, Bool.not_false, Bool.true_and, decide_eq_decide]
rw [List.countP_cons_of_neg]
@ -336,12 +302,12 @@ lemma cons_contrIndexList_of_countP_neq_zero {I : Index X}
· simp only [decide_True, List.countP_cons_of_pos, add_left_eq_self, decide_eq_true_eq]
exact h
lemma mem_contrIndexList_count_contrIndexList {I : Index X} (h : I ∈ l.contrIndexList.val) :
l.contrIndexList.val.countP (fun J => I.id = J.id) = 1 := by
lemma mem_contrIndexList_countId_contrIndexList {I : Index X} (h : I ∈ l.contrIndexList.val) :
l.contrIndexList.countId I = 1 := by
trans (l.val.filter (fun J => I.id = J.id)).countP
(fun i => l.val.countP (fun j => i.id = j.id) = 1)
· rw [contrIndexList]
simp only
simp only [countId]
rw [List.countP_filter, List.countP_filter]
simp only [decide_eq_true_eq, Bool.decide_and]
congr
@ -350,22 +316,21 @@ lemma mem_contrIndexList_count_contrIndexList {I : Index X} (h : I ∈ l.contrIn
· rw [l.mem_contrIndexList_filter h, List.countP_cons_of_pos]
· rfl
· simp only [decide_eq_true_eq]
exact mem_contrIndexList_count l h
exact mem_contrIndexList_countId l h
lemma countP_contrIndexList_zero_of_countP (I : Index X)
(h : List.countP (fun J => I.id = J.id) l.val = 0) :
l.contrIndexList.val.countP (fun J => I.id = J.id) = 0 := by
lemma countId_contrIndexList_zero_of_countId (I : Index X)
(h : l.countId I = 0) :
l.contrIndexList.countId I = 0 := by
trans (l.val.filter (fun J => I.id = J.id)).countP
(fun i => l.val.countP (fun j => i.id = j.id) = 1)
· rw [contrIndexList]
simp only
simp only [countId]
rw [List.countP_filter, List.countP_filter]
simp only [decide_eq_true_eq, Bool.decide_and]
congr
funext a
rw [Bool.and_comm]
· rw [List.countP_eq_length_filter] at h
rw [List.length_eq_zero] at h
· rw [countId_eq_length_filter, List.length_eq_zero] at h
rw [h]
simp only [List.countP_nil]

View file

@ -139,53 +139,43 @@ lemma withUniqueDual_eq_withDual_iff_sort_eq :
-/
lemma withUniqueDual_eq_withDual_iff_countP :
l.withUniqueDual = l.withDual ↔
∀ i, l.val.countP (fun J => (l.val.get i).id = J.id) ≤ 2 := by
lemma withUniqueDual_eq_withDual_iff_countId_leq_two :
l.withUniqueDual = l.withDual ↔ ∀ i, l.countId (l.val.get i) ≤ 2 := by
refine Iff.intro (fun h i => ?_) (fun h => ?_)
· by_cases hi : i ∈ l.withDual
· rw [← h] at hi
rw [mem_withUniqueDual_iff_countP] at hi
rw [mem_withUniqueDual_iff_countId_eq_two] at hi
rw [hi]
· rw [mem_withDual_iff_countP] at hi
· rw [mem_withDual_iff_countId_gt_one] at hi
simp at hi
exact Nat.le_succ_of_le hi
· refine Finset.ext (fun i => ?_)
rw [mem_withUniqueDual_iff_countP, mem_withDual_iff_countP]
rw [mem_withUniqueDual_iff_countId_eq_two, mem_withDual_iff_countId_gt_one]
have hi := h i
omega
lemma withUniqueDual_eq_withDual_filter_id_fin (h : l.withUniqueDual = l.withDual)
(i : Fin l.length) :
l.val.countP (fun J => (l.val.get i).id = J.id) = 0
l.val.countP (fun J => (l.val.get i).id = J.id) = 1
l.val.countP (fun J => (l.val.get i).id = J.id) = 2 := by
by_cases h0 : l.val.countP (fun J => (l.val.get i).id = J.id) = 0
lemma withUniqueDual_eq_withDual_countId_cases (h : l.withUniqueDual = l.withDual)
(i : Fin l.length) : l.countId (l.val.get i) = 0
l.countId (l.val.get i) = 1 l.countId (l.val.get i) = 2 := by
by_cases h0 : l.countId (l.val.get i)= 0
· exact Or.inl h0
· by_cases h1 : l.val.countP (fun J => (l.val.get i).id = J.id) = 1
· by_cases h1 : l.countId (l.val.get i) = 1
· exact Or.inr (Or.inl h1)
· have h2 : l.val.countP (fun J => (l.val.get i).id = J.id) ≤ 2 := by
rw [withUniqueDual_eq_withDual_iff_countP] at h
· have h2 : l.countId (l.val.get i) ≤ 2 := by
rw [withUniqueDual_eq_withDual_iff_countId_leq_two] at h
exact h i
omega
section filterID
lemma filter_id_of_countP_zero {i : Fin l.length}
(h1 : l.val.countP (fun J => (l.val.get i).id = J.id) = 0) :
lemma filter_id_of_countId_eq_zero {i : Fin l.length} (h1 : l.countId (l.val.get i) = 0) :
l.val.filter (fun J => (l.val.get i).id = J.id) = [] := by
rw [List.countP_eq_length_filter, List.length_eq_zero] at h1
rw [countId_eq_length_filter, List.length_eq_zero] at h1
exact h1
lemma filter_of_countP_zero {I : Index X} (h1 : l.val.countP (fun J => I.id = J.id) = 0) :
l.val.filter (fun J => I.id = J.id) = [] := by
rw [List.countP_eq_length_filter, List.length_eq_zero] at h1
exact h1
lemma filter_id_of_countP_one {i : Fin l.length}
(h1 : l.val.countP (fun J => (l.val.get i).id = J.id) = 1) :
lemma filter_id_of_countId_eq_one {i : Fin l.length} (h1 : l.countId (l.val.get i) = 1) :
l.val.filter (fun J => (l.val.get i).id = J.id) = [l.val.get i] := by
rw [List.countP_eq_length_filter, List.length_eq_one] at h1
rw [countId_eq_length_filter, List.length_eq_one] at h1
obtain ⟨J, hJ⟩ := h1
have hme : (l.val.get i) ∈ List.filter (fun J => decide ((l.val.get i).id = J.id)) l.val := by
simp only [List.get_eq_getElem, List.mem_filter, decide_True, and_true]
@ -196,30 +186,31 @@ lemma filter_id_of_countP_one {i : Fin l.length}
simp only [List.get_eq_getElem, List.cons.injEq, and_true]
exact id (Eq.symm hme)
lemma filter_id_of_countP_two {i : Fin l.length}
(h : l.val.countP (fun J => (l.val.get i).id = J.id) = 2) :
lemma filter_id_of_countId_eq_two {i : Fin l.length}
(h : l.countId (l.val.get i) = 2) :
l.val.filter (fun J => (l.val.get i).id = J.id) =
[l.val.get i, l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countP_two h))]
[l.val.get i, l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))]
l.val.filter (fun J => (l.val.get i).id = J.id) =
[l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countP_two h)), l.val.get i] := by
have hc := l.countP_of_mem_withUniqueDual i ((mem_withUniqueDual_iff_countP l i).mpr h)
rw [List.countP_eq_length_filter] at hc
by_cases hi : i < ((l.getDual? i).get (l.getDual?_isSome_of_countP_two h))
[l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)), l.val.get i] := by
have hc := l.countId_eq_two_of_mem_withUniqueDual i
((mem_withUniqueDual_iff_countId_eq_two l i).mpr h)
rw [countId_eq_length_filter] at hc
by_cases hi : i < ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))
· apply Or.inl
symm
apply List.Sublist.eq_of_length
· have h1 : [l.val.get i, l.val.get
((l.getDual? i).get (l.getDual?_isSome_of_countP_two h))].filter
((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))].filter
(fun J => (l.val.get i).id = J.id) = [l.val.get i, l.val.get
((l.getDual? i).get (l.getDual?_isSome_of_countP_two h))] := by
((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))] := by
simp only [List.get_eq_getElem, decide_True, List.filter_cons_of_pos, List.cons.injEq,
List.filter_eq_self, List.mem_singleton, decide_eq_true_eq, forall_eq, true_and]
change l.idMap i = l.idMap ((l.getDual? i).get (l.getDual?_isSome_of_countP_two h))
change l.idMap i = l.idMap ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))
simp
rw [← h1]
refine List.Sublist.filter (fun (J : Index X) => ((l.val.get i).id = J.id)) ?_
rw [List.sublist_iff_exists_fin_orderEmbedding_get_eq]
refine ⟨⟨⟨![i, (l.getDual? i).get (l.getDual?_isSome_of_countP_two h)], ?_⟩, ?_⟩, ?_⟩
refine ⟨⟨⟨![i, (l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)], ?_⟩, ?_⟩, ?_⟩
· refine List.nodup_ofFn.mp ?_
simpa using Fin.ne_of_lt hi
· intro a b
@ -230,27 +221,27 @@ lemma filter_id_of_countP_two {i : Fin l.length}
fin_cases a <;> rfl
· rw [hc]
simp
· have hi' : ((l.getDual? i).get (l.getDual?_isSome_of_countP_two h)) < i := by
have h1 := l.getDual?_get_areDualInSelf i (getDual?_isSome_of_countP_two l h)
· have hi' : ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)) < i := by
have h1 := l.getDual?_get_areDualInSelf i (getDual?_isSome_of_countId_eq_two l h)
simp only [AreDualInSelf] at h1
have h2 := h1.1
omega
apply Or.inr
symm
apply List.Sublist.eq_of_length
· have h1 : [l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countP_two h)),
· have h1 : [l.val.get ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)),
l.val.get i].filter (fun J => (l.val.get i).id = J.id) = [l.val.get
((l.getDual? i).get (l.getDual?_isSome_of_countP_two h)), l.val.get i,] := by
((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h)), l.val.get i,] := by
simp only [List.get_eq_getElem, List.filter_eq_self, List.mem_cons, List.mem_singleton,
decide_eq_true_eq, forall_eq_or_imp, forall_eq, and_true]
apply And.intro
· change l.idMap i = l.idMap ((l.getDual? i).get (l.getDual?_isSome_of_countP_two h))
· change l.idMap i = l.idMap ((l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h))
simp
· simp only [List.not_mem_nil, false_implies, implies_true, and_self]
rw [← h1]
refine List.Sublist.filter (fun (J : Index X) => ((l.val.get i).id = J.id)) ?_
rw [List.sublist_iff_exists_fin_orderEmbedding_get_eq]
refine ⟨⟨⟨![(l.getDual? i).get (l.getDual?_isSome_of_countP_two h), i], ?_⟩, ?_⟩, ?_⟩
refine ⟨⟨⟨![(l.getDual? i).get (l.getDual?_isSome_of_countId_eq_two h), i], ?_⟩, ?_⟩, ?_⟩
· refine List.nodup_ofFn.mp ?_
simp only [List.get_eq_getElem, List.length_singleton, Nat.succ_eq_add_one, Nat.reduceAdd,
List.length_nil, List.ofFn_succ, Fin.isValue, Matrix.cons_val_zero, Matrix.cons_val_succ,
@ -268,30 +259,32 @@ lemma filter_id_of_countP_two {i : Fin l.length}
/-- Given an index `I` such that there is one index in `l` with the same `id` as `I`.
This is that index. -/
def consDual {I : Index X} (hI : l.val.countP (fun J => I.id = J.id) = 1) :
Index X := (l.cons I).val.get (((l.cons I).getDual? ⟨0, by simp⟩).get
((l.cons I).getDual?_isSome_of_countP_two (by simpa using hI)))
def consDual {I : Index X} (hI : l.countId I = 1) : Index X :=
(l.cons I).val.get (((l.cons I).getDual? ⟨0, by simp⟩).get
((l.cons I).getDual?_isSome_of_countId_eq_two (by simpa [countId] using hI)))
/-! TODO: Relate `consDual` to `getDualInOther?`. -/
@[simp]
lemma consDual_id {I : Index X} (hI : l.val.countP (fun J => I.id = J.id) = 1) :
lemma consDual_id {I : Index X} (hI : l.countId I = 1) :
(l.consDual hI).id = I.id := by
change (l.cons I).idMap ((((l.cons I).getDual? ⟨0, by simp⟩).get
((l.cons I).getDual?_isSome_of_countP_two (by simpa using hI)))) = _
((l.cons I).getDual?_isSome_of_countId_eq_two (by simpa [countId] using hI)))) = _
simp only [cons, Fin.zero_eta, getDual?_get_id]
simp only [idMap, List.get_eq_getElem, List.length_cons, Fin.val_zero, List.getElem_cons_zero]
@[simp]
lemma consDual_mem {I : Index X} (hI : l.val.countP (fun J => I.id = J.id) = 1) :
lemma consDual_mem {I : Index X} (hI : l.countId I = 1) :
l.consDual hI ∈ l.val := by
let Di := (((l.cons I).getDual? ⟨0, by simp⟩).get (by
rw [← zero_mem_withUniqueDual_of_cons_iff] at hI; exact
rw [← zero_mem_withUniqueDual_of_cons_iff_countId_one] at hI; exact
mem_withUniqueDual_isSome (l.cons I) ⟨0, _⟩ hI))
have hDiz : Di ≠ ⟨0, by simp⟩ := by
have hd : (l.cons I).AreDualInSelf ⟨0, by simp⟩ Di := by
simp [Di]
symm
exact (l.cons I).getDual?_get_areDualInSelf ⟨0, by simp⟩ (by
rw [← zero_mem_withUniqueDual_of_cons_iff] at hI;
rw [← zero_mem_withUniqueDual_of_cons_iff_countId_one] at hI;
exact mem_withUniqueDual_isSome (l.cons I) ⟨0, _⟩ hI)
simp [AreDualInSelf] at hd
have hd2 := hd.1
@ -313,7 +306,7 @@ lemma consDual_mem {I : Index X} (hI : l.val.countP (fun J => I.id = J.id) = 1)
rw [h1]
exact List.getElem_mem l.val _ _
lemma consDual_filter {I : Index X} (hI : l.val.countP (fun J => I.id = J.id) = 1) :
lemma consDual_filter {I : Index X} (hI : l.countId I = 1) :
l.val.filter (fun J => I.id = J.id) = [l.consDual hI] := by
have h1 : (l.val.filter (fun J => I.id = J.id)).length = 1 := by
rw [← List.countP_eq_length_filter]
@ -328,7 +321,7 @@ lemma consDual_filter {I : Index X} (hI : l.val.countP (fun J => I.id = J.id) =
simp at haI
exact haI.symm
lemma consDual_iff {I : Index X} (hI : l.val.countP (fun J => I.id = J.id) = 1)
lemma consDual_iff {I : Index X} (hI : l.countId I = 1)
(I' : Index X) : I' = l.consDual hI ↔ I' ∈ l.val ∧ I'.id = I.id := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· subst h
@ -338,17 +331,17 @@ lemma consDual_iff {I : Index X} (hI : l.val.countP (fun J => I.id = J.id) = 1)
rw [l.consDual_filter hI] at hI'
simpa using hI'
lemma filter_of_constDual {I : Index X} (hI : l.val.countP (fun J => I.id = J.id) = 1) :
lemma filter_of_constDual {I : Index X} (hI : l.countId I = 1) :
(l.cons I).val.filter (fun J => (l.consDual hI).id = J.id) = [I, l.consDual hI] := by
simp only [consDual_id, cons_val, decide_True, List.filter_cons_of_pos, List.cons.injEq, true_and]
exact consDual_filter l hI
end filterID
lemma withUniqueDual_eq_withDual_iff_countP_mem_le_two :
lemma withUniqueDual_eq_withDual_iff_countId_mem_le_two :
l.withUniqueDual = l.withDual ↔
∀ I (_ : I ∈ l.val), l.val.countP (fun J => I.id = J.id) ≤ 2 := by
rw [withUniqueDual_eq_withDual_iff_countP]
∀ I (_ : I ∈ l.val), l.countId I ≤ 2 := by
rw [withUniqueDual_eq_withDual_iff_countId_leq_two]
refine Iff.intro (fun h I hI => ?_) (fun h i => ?_)
· let i := l.val.indexOf I
have hi : i < l.length := List.indexOf_lt_length.mpr hI
@ -357,10 +350,10 @@ lemma withUniqueDual_eq_withDual_iff_countP_mem_le_two :
exact h ⟨i, hi⟩
· exact h (l.val.get i) (List.getElem_mem l.val (↑i) i.isLt)
lemma withUniqueDual_eq_withDual_iff_all_countP_le_two :
lemma withUniqueDual_eq_withDual_iff_all_countId_le_two :
l.withUniqueDual = l.withDual ↔
l.val.all (fun I => l.val.countP (fun J => I.id = J.id) ≤ 2) := by
rw [withUniqueDual_eq_withDual_iff_countP_mem_le_two]
l.val.all (fun I => l.countId I ≤ 2) := by
rw [withUniqueDual_eq_withDual_iff_countId_mem_le_two]
simp only [List.all_eq_true, decide_eq_true_eq]
/-!
@ -370,11 +363,10 @@ lemma withUniqueDual_eq_withDual_iff_all_countP_le_two :
-/
lemma withUniqueDual_eq_withDual_cons_iff (I : Index X) (hl : l.withUniqueDual = l.withDual) :
(l.cons I).withUniqueDual = (l.cons I).withDual
↔ l.val.countP (fun J => I.id = J.id) ≤ 1 := by
rw [withUniqueDual_eq_withDual_iff_all_countP_le_two]
simp only [cons_val, List.all_cons, decide_True, List.countP_cons_of_pos, Nat.reduceLeDiff,
Bool.and_eq_true, decide_eq_true_eq, List.all_eq_true, and_iff_left_iff_imp]
(l.cons I).withUniqueDual = (l.cons I).withDual ↔ l.countId I ≤ 1 := by
rw [withUniqueDual_eq_withDual_iff_all_countId_le_two]
simp only [cons_val, countId, List.all_cons, decide_True, List.countP_cons_of_pos,
Nat.reduceLeDiff, Bool.and_eq_true, decide_eq_true_eq, List.all_eq_true, and_iff_left_iff_imp]
intro h I' hI'
by_cases hII' : I'.id = I.id
· rw [List.countP_cons_of_pos]
@ -382,14 +374,14 @@ lemma withUniqueDual_eq_withDual_cons_iff (I : Index X) (hl : l.withUniqueDual =
omega
· simpa using hII'
· rw [List.countP_cons_of_neg]
· rw [withUniqueDual_eq_withDual_iff_countP_mem_le_two] at hl
· rw [withUniqueDual_eq_withDual_iff_countId_mem_le_two] at hl
exact hl I' hI'
· simpa using hII'
lemma withUniqueDual_eq_withDual_of_cons {I : Index X}
(hl : (l.cons I).withUniqueDual = (l.cons I).withDual) :
l.withUniqueDual = l.withDual := by
rw [withUniqueDual_eq_withDual_iff_countP_mem_le_two] at hl ⊢
rw [withUniqueDual_eq_withDual_iff_countId_mem_le_two] at hl ⊢
intro I' hI'
have hImem : I' ∈ (l.cons I).val := by
simp [hI']
@ -405,7 +397,7 @@ lemma withUniqueDual_eq_withDual_of_cons {I : Index X}
lemma withUniqueDual_eq_withDual_cons_iff' (I : Index X) :
(l.cons I).withUniqueDual = (l.cons I).withDual ↔
l.withUniqueDual = l.withDual ∧ l.val.countP (fun J => I.id = J.id) ≤ 1 := by
l.withUniqueDual = l.withDual ∧ l.countId I ≤ 1 := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· have h1 : l.withUniqueDual = l.withDual := by
exact withUniqueDual_eq_withDual_of_cons l h

View file

@ -77,13 +77,14 @@ lemma mem_withDual_iff_exists : i ∈ l.withDual ↔ ∃ j, l.AreDualInSelf i j
-/
lemma countP_of_mem_withDual (i : Fin l.length) (h : i ∈ l.withDual) :
1 < l.val.countP (fun J => (l.val.get i).id = J.id) := by
lemma countId_gt_zero_of_mem_withDual (i : Fin l.length) (h : i ∈ l.withDual) :
1 < l.countId (l.val.get i) := by
rw [countId]
rw [mem_withDual_iff_exists] at h
obtain ⟨j, hj⟩ := h
simp [AreDualInSelf, idMap] at hj
by_contra hn
have hn' := l.countP_id_neq_zero i
have hn' := l.countId_index_neq_zero i
have hl : 2 ≤ l.val.countP (fun J => (l.val.get i).id = J.id) := by
by_cases hij : i < j
· have hsub : List.Sublist [l.val.get i, l.val.get j] l.val := by
@ -115,8 +116,9 @@ lemma countP_of_mem_withDual (i : Fin l.length) (h : i ∈ l.withDual) :
(fun (j : Index X) => decide (l.val[i].id = j.id)) hsub
omega
lemma countP_of_not_mem_withDual (i : Fin l.length)(h : i ∉ l.withDual) :
l.val.countP (fun J => (l.val.get i).id = J.id) = 1 := by
lemma countId_of_not_mem_withDual (i : Fin l.length)(h : i ∉ l.withDual) :
l.countId (l.val.get i) = 1 := by
rw [countId]
rw [mem_withDual_iff_exists] at h
simp [AreDualInSelf] at h
have h1 : ¬ l.val.Duplicate (l.val.get i) := by
@ -156,11 +158,11 @@ lemma countP_of_not_mem_withDual (i : Fin l.length)(h : i ∉ l.withDual) :
· rw [h']
rfl
lemma mem_withDual_iff_countP (i : Fin l.length) :
i ∈ l.withDual ↔ 1 < l.val.countP (fun J => (l.val.get i).id = J.id) := by
refine Iff.intro (fun h => countP_of_mem_withDual l i h) (fun h => ?_)
lemma mem_withDual_iff_countId_gt_one (i : Fin l.length) :
i ∈ l.withDual ↔ 1 < l.countId (l.val.get i) := by
refine Iff.intro (fun h => countId_gt_zero_of_mem_withDual l i h) (fun h => ?_)
by_contra hn
have hn' := countP_of_not_mem_withDual l i hn
have hn' := countId_of_not_mem_withDual l i hn
omega
/-!

View file

@ -968,20 +968,20 @@ lemma card_finset_self_dual (i : Fin l.length) (h : i ∈ l.withDual) :
simp
exact h1.1
lemma countP_of_mem_withUniqueDual (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
l.val.countP (fun J => (l.val.get i).id = J.id) = 2 := by
rw [List.countP_eq_length_filter, filter_id_eq_sort]
lemma countId_eq_two_of_mem_withUniqueDual (i : Fin l.length) (h : i ∈ l.withUniqueDual) :
l.countId (l.val.get i) = 2 := by
rw [countId_eq_length_filter, filter_id_eq_sort]
simp only [List.length_map, Finset.length_sort]
erw [l.finset_filter_id_mem_withUniqueDual i h]
refine l.card_finset_self_dual i (mem_withDual_of_mem_withUniqueDual l i h)
lemma mem_withUniqueDual_of_countP (i : Fin l.length)
(h : l.val.countP (fun J => (l.val.get i).id = J.id) = 2) : i ∈ l.withUniqueDual := by
lemma mem_withUniqueDual_of_countId_eq_two (i : Fin l.length)
(h : l.countId (l.val.get i) = 2) : i ∈ l.withUniqueDual := by
have hw : i ∈ l.withDual := by
rw [mem_withDual_iff_countP, h]
rw [mem_withDual_iff_countId_gt_one, h]
exact Nat.one_lt_two
rw [l.mem__withUniqueDual_iff_finset_filter i hw]
rw [List.countP_eq_length_filter, filter_id_eq_sort] at h
rw [countId_eq_length_filter, filter_id_eq_sort] at h
simp at h
have hsub : {i, (l.getDual? i).get ((mem_withDual_iff_isSome l i).mp hw)} ⊆
Finset.filter (fun j => l.idMap i = l.idMap j) Finset.univ := by
@ -991,15 +991,15 @@ lemma mem_withUniqueDual_of_countP (i : Fin l.length)
erw [h]
rw [l.card_finset_self_dual i hw]
lemma mem_withUniqueDual_iff_countP (i : Fin l.length) :
i ∈ l.withUniqueDual ↔ l.val.countP (fun J => (l.val.get i).id = J.id) = 2 :=
Iff.intro (fun h => l.countP_of_mem_withUniqueDual i h)
(fun h => l.mem_withUniqueDual_of_countP i h)
lemma mem_withUniqueDual_iff_countId_eq_two (i : Fin l.length) :
i ∈ l.withUniqueDual ↔ l.countId (l.val.get i) = 2 :=
Iff.intro (fun h => l.countId_eq_two_of_mem_withUniqueDual i h)
(fun h => l.mem_withUniqueDual_of_countId_eq_two i h)
lemma getDual?_isSome_of_countP_two {i : Fin l.length}
(h : l.val.countP (fun J => (l.val.get i).id = J.id) = 2) :
lemma getDual?_isSome_of_countId_eq_two {i : Fin l.length}
(h : l.countId (l.val.get i) = 2) :
(l.getDual? i).isSome := by
rw [← l.mem_withUniqueDual_iff_countP] at h
rw [← l.mem_withUniqueDual_iff_countId_eq_two] at h
exact mem_withUniqueDual_isSome l i h
section cons
@ -1008,9 +1008,9 @@ section cons
## Cons and withUniqueDual
-/
lemma zero_mem_withUniqueDual_of_cons_iff (I : Index X) :
⟨0, by simp⟩ ∈ (l.cons I).withUniqueDual ↔ l.val.countP (fun J => I.id = J.id) = 1 := by
rw [mem_withUniqueDual_iff_countP]
lemma zero_mem_withUniqueDual_of_cons_iff_countId_one (I : Index X) :
⟨0, by simp⟩ ∈ (l.cons I).withUniqueDual ↔ l.countId I = 1 := by
rw [mem_withUniqueDual_iff_countId_eq_two, countId, countId]
simp
end cons