Skip to content

Commit 4f51227

Browse files
author
M Hidalgo
committed
feat(Combinatorics/SimpleGraph/Walk): add lemma support_injective (#27472)
This contribution was created as part of the Utrecht Summerschool "Formalizing Mathematics in Lean" in July 2025.
1 parent f0b4174 commit 4f51227

File tree

1 file changed

+3
-0
lines changed

1 file changed

+3
-0
lines changed

Mathlib/Combinatorics/SimpleGraph/Walk.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1136,6 +1136,9 @@ lemma ext_support {u v} {p q : G.Walk u v} (h : p.support = q.support) :
11361136
have : (p.copy h rfl).support = q.support := by simpa
11371137
simp [← ih this]
11381138

1139+
lemma support_injective {u v : V} : (support (G := G) (u := u) (v := v)).Injective :=
1140+
fun _ _ ↦ ext_support
1141+
11391142
lemma ext_getVert_le_length {u v} {p q : G.Walk u v} (hl : p.length = q.length)
11401143
(h : ∀ k ≤ p.length, p.getVert k = q.getVert k) :
11411144
p = q := by

0 commit comments

Comments
 (0)