![]() |
Metamath
Proof Explorer Theorem List (p. 263 of 429) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > MPE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Color key: | ![]() (1-27903) |
![]() (27904-29428) |
![]() (29429-42879) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | griedg0prc 26201* | The class of empty graphs (represented as ordered pairs) is a proper class. (Contributed by AV, 27-Dec-2020.) |
⊢ 𝑈 = {〈𝑣, 𝑒〉 ∣ 𝑒:∅⟶∅} ⇒ ⊢ 𝑈 ∉ V | ||
Theorem | griedg0ssusgr 26202* | The class of all simple graphs is a superclass of the class of empty graphs represented as ordered pairs. (Contributed by AV, 27-Dec-2020.) |
⊢ 𝑈 = {〈𝑣, 𝑒〉 ∣ 𝑒:∅⟶∅} ⇒ ⊢ 𝑈 ⊆ USGraph | ||
Theorem | usgrprc 26203 | The class of simple graphs is a proper class (and therefore, because of prcssprc 4839, the classes of multigraphs, pseudographs and hypergraphs are proper classes, too). (Contributed by AV, 27-Dec-2020.) |
⊢ USGraph ∉ V | ||
Syntax | csubgr 26204 | Extend class notation with subgraphs. |
class SubGraph | ||
Definition | df-subgr 26205* | Define the class of the subgraph relation. A class 𝑠 is a subgraph of a class 𝑔 (the supergraph of 𝑠) if its vertices are also vertices of 𝑔, and its edges are also edges of 𝑔, connecting vertices of 𝑠 only (see section I.1 in [Bollobas] p. 2 or section 1.1 in [Diestel] p. 4). The second condition is ensured by the requirement that the edge function of 𝑠 is a restriction of the edge function of 𝑔 having only vertices of 𝑠 in its range. Note that the domains of the edge functions of the subgraph and the supergraph should be compatible. (Contributed by AV, 16-Nov-2020.) |
⊢ SubGraph = {〈𝑠, 𝑔〉 ∣ ((Vtx‘𝑠) ⊆ (Vtx‘𝑔) ∧ (iEdg‘𝑠) = ((iEdg‘𝑔) ↾ dom (iEdg‘𝑠)) ∧ (Edg‘𝑠) ⊆ 𝒫 (Vtx‘𝑠))} | ||
Theorem | relsubgr 26206 | The class of the subgraph relation is a relation. (Contributed by AV, 16-Nov-2020.) |
⊢ Rel SubGraph | ||
Theorem | subgrv 26207 | If a class is a subgraph of another class, both classes are sets. (Contributed by AV, 16-Nov-2020.) |
⊢ (𝑆 SubGraph 𝐺 → (𝑆 ∈ V ∧ 𝐺 ∈ V)) | ||
Theorem | issubgr 26208 | The property of a set to be a subgraph of another set. (Contributed by AV, 16-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝑆) & ⊢ 𝐴 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝑆) & ⊢ 𝐵 = (iEdg‘𝐺) & ⊢ 𝐸 = (Edg‘𝑆) ⇒ ⊢ ((𝐺 ∈ 𝑊 ∧ 𝑆 ∈ 𝑈) → (𝑆 SubGraph 𝐺 ↔ (𝑉 ⊆ 𝐴 ∧ 𝐼 = (𝐵 ↾ dom 𝐼) ∧ 𝐸 ⊆ 𝒫 𝑉))) | ||
Theorem | issubgr2 26209 | The property of a set to be a subgraph of a set whose edge function is actually a function. (Contributed by AV, 20-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝑆) & ⊢ 𝐴 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝑆) & ⊢ 𝐵 = (iEdg‘𝐺) & ⊢ 𝐸 = (Edg‘𝑆) ⇒ ⊢ ((𝐺 ∈ 𝑊 ∧ Fun 𝐵 ∧ 𝑆 ∈ 𝑈) → (𝑆 SubGraph 𝐺 ↔ (𝑉 ⊆ 𝐴 ∧ 𝐼 ⊆ 𝐵 ∧ 𝐸 ⊆ 𝒫 𝑉))) | ||
Theorem | subgrprop 26210 | The properties of a subgraph. (Contributed by AV, 19-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝑆) & ⊢ 𝐴 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝑆) & ⊢ 𝐵 = (iEdg‘𝐺) & ⊢ 𝐸 = (Edg‘𝑆) ⇒ ⊢ (𝑆 SubGraph 𝐺 → (𝑉 ⊆ 𝐴 ∧ 𝐼 = (𝐵 ↾ dom 𝐼) ∧ 𝐸 ⊆ 𝒫 𝑉)) | ||
Theorem | subgrprop2 26211 | The properties of a subgraph: If 𝑆 is a subgraph of 𝐺, its vertices are also vertices of 𝐺, and its edges are also edges of 𝐺, connecting vertices of the subgraph only. (Contributed by AV, 19-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝑆) & ⊢ 𝐴 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝑆) & ⊢ 𝐵 = (iEdg‘𝐺) & ⊢ 𝐸 = (Edg‘𝑆) ⇒ ⊢ (𝑆 SubGraph 𝐺 → (𝑉 ⊆ 𝐴 ∧ 𝐼 ⊆ 𝐵 ∧ 𝐸 ⊆ 𝒫 𝑉)) | ||
Theorem | uhgrissubgr 26212 | The property of a hypergraph to be a subgraph. (Contributed by AV, 19-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝑆) & ⊢ 𝐴 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝑆) & ⊢ 𝐵 = (iEdg‘𝐺) ⇒ ⊢ ((𝐺 ∈ 𝑊 ∧ Fun 𝐵 ∧ 𝑆 ∈ UHGraph) → (𝑆 SubGraph 𝐺 ↔ (𝑉 ⊆ 𝐴 ∧ 𝐼 ⊆ 𝐵))) | ||
Theorem | subgrprop3 26213 | The properties of a subgraph: If 𝑆 is a subgraph of 𝐺, its vertices are also vertices of 𝐺, and its edges are also edges of 𝐺. (Contributed by AV, 19-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝑆) & ⊢ 𝐴 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝑆) & ⊢ 𝐵 = (Edg‘𝐺) ⇒ ⊢ (𝑆 SubGraph 𝐺 → (𝑉 ⊆ 𝐴 ∧ 𝐸 ⊆ 𝐵)) | ||
Theorem | egrsubgr 26214 | An empty graph consisting of a subset of vertices of a graph (and having no edges) is a subgraph of the graph. (Contributed by AV, 17-Nov-2020.) (Proof shortened by AV, 17-Dec-2020.) |
⊢ (((𝐺 ∈ 𝑊 ∧ 𝑆 ∈ 𝑈) ∧ (Vtx‘𝑆) ⊆ (Vtx‘𝐺) ∧ (Fun (iEdg‘𝑆) ∧ (Edg‘𝑆) = ∅)) → 𝑆 SubGraph 𝐺) | ||
Theorem | 0grsubgr 26215 | The null graph (represented by an empty set) is a subgraph of all graphs. (Contributed by AV, 17-Nov-2020.) |
⊢ (𝐺 ∈ 𝑊 → ∅ SubGraph 𝐺) | ||
Theorem | 0uhgrsubgr 26216 | The null graph (as hypergraph) is a subgraph of all graphs. (Contributed by AV, 17-Nov-2020.) (Proof shortened by AV, 28-Nov-2020.) |
⊢ ((𝐺 ∈ 𝑊 ∧ 𝑆 ∈ UHGraph ∧ (Vtx‘𝑆) = ∅) → 𝑆 SubGraph 𝐺) | ||
Theorem | uhgrsubgrself 26217 | A hypergraph is a subgraph of itself. (Contributed by AV, 17-Nov-2020.) (Proof shortened by AV, 21-Nov-2020.) |
⊢ (𝐺 ∈ UHGraph → 𝐺 SubGraph 𝐺) | ||
Theorem | subgrfun 26218 | The edge function of a subgraph of a graph whose edge function is actually a function is a function. (Contributed by AV, 20-Nov-2020.) |
⊢ ((Fun (iEdg‘𝐺) ∧ 𝑆 SubGraph 𝐺) → Fun (iEdg‘𝑆)) | ||
Theorem | subgruhgrfun 26219 | The edge function of a subgraph of a hypergraph is a function. (Contributed by AV, 16-Nov-2020.) (Proof shortened by AV, 20-Nov-2020.) |
⊢ ((𝐺 ∈ UHGraph ∧ 𝑆 SubGraph 𝐺) → Fun (iEdg‘𝑆)) | ||
Theorem | subgreldmiedg 26220 | An element of the domain of the edge function of a subgraph is an element of the domain of the edge function of the supergraph. (Contributed by AV, 20-Nov-2020.) |
⊢ ((𝑆 SubGraph 𝐺 ∧ 𝑋 ∈ dom (iEdg‘𝑆)) → 𝑋 ∈ dom (iEdg‘𝐺)) | ||
Theorem | subgruhgredgd 26221 | An edge of a subgraph of a hypergraph is a nonempty subset of its vertices. (Contributed by AV, 17-Nov-2020.) (Revised by AV, 21-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝑆) & ⊢ 𝐼 = (iEdg‘𝑆) & ⊢ (𝜑 → 𝐺 ∈ UHGraph) & ⊢ (𝜑 → 𝑆 SubGraph 𝐺) & ⊢ (𝜑 → 𝑋 ∈ dom 𝐼) ⇒ ⊢ (𝜑 → (𝐼‘𝑋) ∈ (𝒫 𝑉 ∖ {∅})) | ||
Theorem | subumgredg2 26222* | An edge of a subgraph of a multigraph connects exactly two different vertices. (Contributed by AV, 26-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝑆) & ⊢ 𝐼 = (iEdg‘𝑆) ⇒ ⊢ ((𝑆 SubGraph 𝐺 ∧ 𝐺 ∈ UMGraph ∧ 𝑋 ∈ dom 𝐼) → (𝐼‘𝑋) ∈ {𝑒 ∈ 𝒫 𝑉 ∣ (#‘𝑒) = 2}) | ||
Theorem | subuhgr 26223 | A subgraph of a hypergraph is a hypergraph. (Contributed by AV, 16-Nov-2020.) (Proof shortened by AV, 21-Nov-2020.) |
⊢ ((𝐺 ∈ UHGraph ∧ 𝑆 SubGraph 𝐺) → 𝑆 ∈ UHGraph) | ||
Theorem | subupgr 26224 | A subgraph of a pseudograph is a pseudograph. (Contributed by AV, 16-Nov-2020.) (Proof shortened by AV, 21-Nov-2020.) |
⊢ ((𝐺 ∈ UPGraph ∧ 𝑆 SubGraph 𝐺) → 𝑆 ∈ UPGraph) | ||
Theorem | subumgr 26225 | A subgraph of a multigraph is a multigraph. (Contributed by AV, 26-Nov-2020.) |
⊢ ((𝐺 ∈ UMGraph ∧ 𝑆 SubGraph 𝐺) → 𝑆 ∈ UMGraph) | ||
Theorem | subusgr 26226 | A subgraph of a simple graph is a simple graph. (Contributed by AV, 16-Nov-2020.) (Proof shortened by AV, 27-Nov-2020.) |
⊢ ((𝐺 ∈ USGraph ∧ 𝑆 SubGraph 𝐺) → 𝑆 ∈ USGraph) | ||
Theorem | uhgrspansubgrlem 26227 | Lemma for uhgrspansubgr 26228: The edges of the graph 𝑆 obtained by removing some edges of a hypergraph 𝐺 are subsets of its vertices (a spanning subgraph, see comment for uhgrspansubgr 26228. (Contributed by AV, 18-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ (𝜑 → 𝑆 ∈ 𝑊) & ⊢ (𝜑 → (Vtx‘𝑆) = 𝑉) & ⊢ (𝜑 → (iEdg‘𝑆) = (𝐸 ↾ 𝐴)) & ⊢ (𝜑 → 𝐺 ∈ UHGraph) ⇒ ⊢ (𝜑 → (Edg‘𝑆) ⊆ 𝒫 (Vtx‘𝑆)) | ||
Theorem | uhgrspansubgr 26228 | A spanning subgraph 𝑆 of a hypergraph 𝐺 is actually a subgraph of 𝐺. A subgraph 𝑆 of a graph 𝐺 which has the same vertices as 𝐺 and is obtained by removing some edges of 𝐺 is called a spanning subgraph (see section I.1 in [Bollobas] p. 2 and section 1.1 in [Diestel] p. 4). Formally, the edges are "removed" by restricting the edge function of the original graph by an arbitrary class (which actually needs not to be a subset of the domain of the edge function). (Contributed by AV, 18-Nov-2020.) (Proof shortened by AV, 21-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ (𝜑 → 𝑆 ∈ 𝑊) & ⊢ (𝜑 → (Vtx‘𝑆) = 𝑉) & ⊢ (𝜑 → (iEdg‘𝑆) = (𝐸 ↾ 𝐴)) & ⊢ (𝜑 → 𝐺 ∈ UHGraph) ⇒ ⊢ (𝜑 → 𝑆 SubGraph 𝐺) | ||
Theorem | uhgrspan 26229 | A spanning subgraph 𝑆 of a hypergraph 𝐺 is a hypergraph. (Contributed by AV, 11-Oct-2020.) (Proof shortened by AV, 18-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ (𝜑 → 𝑆 ∈ 𝑊) & ⊢ (𝜑 → (Vtx‘𝑆) = 𝑉) & ⊢ (𝜑 → (iEdg‘𝑆) = (𝐸 ↾ 𝐴)) & ⊢ (𝜑 → 𝐺 ∈ UHGraph) ⇒ ⊢ (𝜑 → 𝑆 ∈ UHGraph) | ||
Theorem | upgrspan 26230 | A spanning subgraph 𝑆 of a pseudograph 𝐺 is a pseudograph. (Contributed by AV, 11-Oct-2020.) (Proof shortened by AV, 18-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ (𝜑 → 𝑆 ∈ 𝑊) & ⊢ (𝜑 → (Vtx‘𝑆) = 𝑉) & ⊢ (𝜑 → (iEdg‘𝑆) = (𝐸 ↾ 𝐴)) & ⊢ (𝜑 → 𝐺 ∈ UPGraph) ⇒ ⊢ (𝜑 → 𝑆 ∈ UPGraph) | ||
Theorem | umgrspan 26231 | A spanning subgraph 𝑆 of a multigraph 𝐺 is a multigraph. (Contributed by AV, 27-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ (𝜑 → 𝑆 ∈ 𝑊) & ⊢ (𝜑 → (Vtx‘𝑆) = 𝑉) & ⊢ (𝜑 → (iEdg‘𝑆) = (𝐸 ↾ 𝐴)) & ⊢ (𝜑 → 𝐺 ∈ UMGraph) ⇒ ⊢ (𝜑 → 𝑆 ∈ UMGraph) | ||
Theorem | usgrspan 26232 | A spanning subgraph 𝑆 of a simple graph 𝐺 is a simple graph. (Contributed by AV, 15-Oct-2020.) (Revised by AV, 16-Oct-2020.) (Proof shortened by AV, 18-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ (𝜑 → 𝑆 ∈ 𝑊) & ⊢ (𝜑 → (Vtx‘𝑆) = 𝑉) & ⊢ (𝜑 → (iEdg‘𝑆) = (𝐸 ↾ 𝐴)) & ⊢ (𝜑 → 𝐺 ∈ USGraph) ⇒ ⊢ (𝜑 → 𝑆 ∈ USGraph) | ||
Theorem | uhgrspanop 26233 | A spanning subgraph of a hypergraph represented by an ordered pair is a hypergraph. (Contributed by Alexander van der Vekens, 27-Dec-2017.) (Revised by AV, 11-Oct-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) ⇒ ⊢ (𝐺 ∈ UHGraph → 〈𝑉, (𝐸 ↾ 𝐴)〉 ∈ UHGraph) | ||
Theorem | upgrspanop 26234 | A spanning subgraph of a pseudograph represented by an ordered pair is a pseudograph. (Contributed by Mario Carneiro, 12-Mar-2015.) (Revised by AV, 13-Oct-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) ⇒ ⊢ (𝐺 ∈ UPGraph → 〈𝑉, (𝐸 ↾ 𝐴)〉 ∈ UPGraph) | ||
Theorem | umgrspanop 26235 | A spanning subgraph of a multigraph represented by an ordered pair is a multigraph. (Contributed by AV, 27-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) ⇒ ⊢ (𝐺 ∈ UMGraph → 〈𝑉, (𝐸 ↾ 𝐴)〉 ∈ UMGraph) | ||
Theorem | usgrspanop 26236 | A spanning subgraph of a simple graph represented by an ordered pair is a simple graph. (Contributed by Alexander van der Vekens, 10-Aug-2017.) (Revised by AV, 16-Oct-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) ⇒ ⊢ (𝐺 ∈ USGraph → 〈𝑉, (𝐸 ↾ 𝐴)〉 ∈ USGraph) | ||
Theorem | uhgrspan1lem1 26237 | Lemma 1 for uhgrspan1 26240. (Contributed by AV, 19-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝐺) & ⊢ 𝐹 = {𝑖 ∈ dom 𝐼 ∣ 𝑁 ∉ (𝐼‘𝑖)} ⇒ ⊢ ((𝑉 ∖ {𝑁}) ∈ V ∧ (𝐼 ↾ 𝐹) ∈ V) | ||
Theorem | uhgrspan1lem2 26238 | Lemma 2 for uhgrspan1 26240. (Contributed by AV, 19-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝐺) & ⊢ 𝐹 = {𝑖 ∈ dom 𝐼 ∣ 𝑁 ∉ (𝐼‘𝑖)} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), (𝐼 ↾ 𝐹)〉 ⇒ ⊢ (Vtx‘𝑆) = (𝑉 ∖ {𝑁}) | ||
Theorem | uhgrspan1lem3 26239 | Lemma 3 for uhgrspan1 26240. (Contributed by AV, 19-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝐺) & ⊢ 𝐹 = {𝑖 ∈ dom 𝐼 ∣ 𝑁 ∉ (𝐼‘𝑖)} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), (𝐼 ↾ 𝐹)〉 ⇒ ⊢ (iEdg‘𝑆) = (𝐼 ↾ 𝐹) | ||
Theorem | uhgrspan1 26240* | The induced subgraph 𝑆 of a hypergraph 𝐺 obtained by removing one vertex is actually a subgraph of 𝐺. A subgraph is called induced or spanned by a subset of vertices of a graph if it contains all edges of the original graph that join two vertices of the subgraph (see section I.1 in [Bollobas] p. 2 and section 1.1 in [Diestel] p. 4). (Contributed by AV, 19-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝐺) & ⊢ 𝐹 = {𝑖 ∈ dom 𝐼 ∣ 𝑁 ∉ (𝐼‘𝑖)} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), (𝐼 ↾ 𝐹)〉 ⇒ ⊢ ((𝐺 ∈ UHGraph ∧ 𝑁 ∈ 𝑉) → 𝑆 SubGraph 𝐺) | ||
Theorem | upgrreslem 26241* | Lemma for upgrres 26243. (Contributed by AV, 27-Nov-2020.) (Revised by AV, 19-Dec-2021.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ 𝐹 = {𝑖 ∈ dom 𝐸 ∣ 𝑁 ∉ (𝐸‘𝑖)} ⇒ ⊢ ((𝐺 ∈ UPGraph ∧ 𝑁 ∈ 𝑉) → ran (𝐸 ↾ 𝐹) ⊆ {𝑝 ∈ (𝒫 (𝑉 ∖ {𝑁}) ∖ {∅}) ∣ (#‘𝑝) ≤ 2}) | ||
Theorem | umgrreslem 26242* | Lemma for umgrres 26244 and usgrres 26245. (Contributed by AV, 27-Nov-2020.) (Revised by AV, 19-Dec-2021.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ 𝐹 = {𝑖 ∈ dom 𝐸 ∣ 𝑁 ∉ (𝐸‘𝑖)} ⇒ ⊢ ((𝐺 ∈ UMGraph ∧ 𝑁 ∈ 𝑉) → ran (𝐸 ↾ 𝐹) ⊆ {𝑝 ∈ 𝒫 (𝑉 ∖ {𝑁}) ∣ (#‘𝑝) = 2}) | ||
Theorem | upgrres 26243* | A subgraph obtained by removing one vertex and all edges incident with this vertex from a pseudograph (see uhgrspan1 26240) is a pseudograph. (Contributed by AV, 8-Nov-2020.) (Revised by AV, 19-Dec-2021.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ 𝐹 = {𝑖 ∈ dom 𝐸 ∣ 𝑁 ∉ (𝐸‘𝑖)} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), (𝐸 ↾ 𝐹)〉 ⇒ ⊢ ((𝐺 ∈ UPGraph ∧ 𝑁 ∈ 𝑉) → 𝑆 ∈ UPGraph) | ||
Theorem | umgrres 26244* | A subgraph obtained by removing one vertex and all edges incident with this vertex from a multigraph (see uhgrspan1 26240) is a multigraph. (Contributed by AV, 27-Nov-2020.) (Revised by AV, 19-Dec-2021.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ 𝐹 = {𝑖 ∈ dom 𝐸 ∣ 𝑁 ∉ (𝐸‘𝑖)} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), (𝐸 ↾ 𝐹)〉 ⇒ ⊢ ((𝐺 ∈ UMGraph ∧ 𝑁 ∈ 𝑉) → 𝑆 ∈ UMGraph) | ||
Theorem | usgrres 26245* | A subgraph obtained by removing one vertex and all edges incident with this vertex from a simple graph (see uhgrspan1 26240) is a simple graph. (Contributed by Alexander van der Vekens, 2-Jan-2018.) (Revised by AV, 19-Dec-2021.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (iEdg‘𝐺) & ⊢ 𝐹 = {𝑖 ∈ dom 𝐸 ∣ 𝑁 ∉ (𝐸‘𝑖)} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), (𝐸 ↾ 𝐹)〉 ⇒ ⊢ ((𝐺 ∈ USGraph ∧ 𝑁 ∈ 𝑉) → 𝑆 ∈ USGraph) | ||
Theorem | upgrres1lem1 26246* | Lemma 1 for upgrres1 26250. (Contributed by AV, 7-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) & ⊢ 𝐹 = {𝑒 ∈ 𝐸 ∣ 𝑁 ∉ 𝑒} ⇒ ⊢ ((𝑉 ∖ {𝑁}) ∈ V ∧ ( I ↾ 𝐹) ∈ V) | ||
Theorem | umgrres1lem 26247* | Lemma for umgrres1 26251. (Contributed by AV, 27-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) & ⊢ 𝐹 = {𝑒 ∈ 𝐸 ∣ 𝑁 ∉ 𝑒} ⇒ ⊢ ((𝐺 ∈ UMGraph ∧ 𝑁 ∈ 𝑉) → ran ( I ↾ 𝐹) ⊆ {𝑝 ∈ 𝒫 (𝑉 ∖ {𝑁}) ∣ (#‘𝑝) = 2}) | ||
Theorem | upgrres1lem2 26248* | Lemma 2 for upgrres1 26250. (Contributed by AV, 7-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) & ⊢ 𝐹 = {𝑒 ∈ 𝐸 ∣ 𝑁 ∉ 𝑒} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), ( I ↾ 𝐹)〉 ⇒ ⊢ (Vtx‘𝑆) = (𝑉 ∖ {𝑁}) | ||
Theorem | upgrres1lem3 26249* | Lemma 3 for upgrres1 26250. (Contributed by AV, 7-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) & ⊢ 𝐹 = {𝑒 ∈ 𝐸 ∣ 𝑁 ∉ 𝑒} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), ( I ↾ 𝐹)〉 ⇒ ⊢ (iEdg‘𝑆) = ( I ↾ 𝐹) | ||
Theorem | upgrres1 26250* | A pseudograph obtained by removing one vertex and all edges incident with this vertex is a pseudograph. Remark: This graph is not a subgraph of the original graph in the sense of df-subgr 26205 since the domains of the edge functions may not be compatible. (Contributed by AV, 8-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) & ⊢ 𝐹 = {𝑒 ∈ 𝐸 ∣ 𝑁 ∉ 𝑒} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), ( I ↾ 𝐹)〉 ⇒ ⊢ ((𝐺 ∈ UPGraph ∧ 𝑁 ∈ 𝑉) → 𝑆 ∈ UPGraph) | ||
Theorem | umgrres1 26251* | A multigraph obtained by removing one vertex and all edges incident with this vertex is a multigraph. Remark: This graph is not a subgraph of the original graph in the sense of df-subgr 26205 since the domains of the edge functions may not be compatible. (Contributed by AV, 27-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) & ⊢ 𝐹 = {𝑒 ∈ 𝐸 ∣ 𝑁 ∉ 𝑒} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), ( I ↾ 𝐹)〉 ⇒ ⊢ ((𝐺 ∈ UMGraph ∧ 𝑁 ∈ 𝑉) → 𝑆 ∈ UMGraph) | ||
Theorem | usgrres1 26252* | Restricting a simple graph by removing one vertex results in a simple graph. Remark: This restricted graph is not a subgraph of the original graph in the sense of df-subgr 26205 since the domains of the edge functions may not be compatible. (Contributed by Alexander van der Vekens, 2-Jan-2018.) (Revised by AV, 10-Jan-2020.) (Revised by AV, 23-Oct-2020.) (Proof shortened by AV, 27-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) & ⊢ 𝐹 = {𝑒 ∈ 𝐸 ∣ 𝑁 ∉ 𝑒} & ⊢ 𝑆 = 〈(𝑉 ∖ {𝑁}), ( I ↾ 𝐹)〉 ⇒ ⊢ ((𝐺 ∈ USGraph ∧ 𝑁 ∈ 𝑉) → 𝑆 ∈ USGraph) | ||
Syntax | cfusgr 26253 | Extend class notation with finite simple graphs. |
class FinUSGraph | ||
Definition | df-fusgr 26254 | Define the class of all finite undirected simple graphs without loops (called "finite simple graphs" in the following). A finite simple graph is an undirected simple graph of finite order, i.e. with a finite set of vertices. (Contributed by AV, 3-Jan-2020.) (Revised by AV, 21-Oct-2020.) |
⊢ FinUSGraph = {𝑔 ∈ USGraph ∣ (Vtx‘𝑔) ∈ Fin} | ||
Theorem | isfusgr 26255 | The property of being a finite simple graph. (Contributed by AV, 3-Jan-2020.) (Revised by AV, 21-Oct-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) ⇒ ⊢ (𝐺 ∈ FinUSGraph ↔ (𝐺 ∈ USGraph ∧ 𝑉 ∈ Fin)) | ||
Theorem | fusgrvtxfi 26256 | A finite simple graph has a finite set of vertices. (Contributed by AV, 16-Dec-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) ⇒ ⊢ (𝐺 ∈ FinUSGraph → 𝑉 ∈ Fin) | ||
Theorem | isfusgrf1 26257* | The property of being a finite simple graph. (Contributed by AV, 3-Jan-2020.) (Revised by AV, 21-Oct-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝐺) ⇒ ⊢ (𝐺 ∈ 𝑊 → (𝐺 ∈ FinUSGraph ↔ (𝐼:dom 𝐼–1-1→{𝑥 ∈ 𝒫 𝑉 ∣ (#‘𝑥) = 2} ∧ 𝑉 ∈ Fin))) | ||
Theorem | isfusgrcl 26258 | The property of being a finite simple graph. (Contributed by AV, 3-Jan-2020.) (Revised by AV, 9-Jan-2020.) |
⊢ (𝐺 ∈ FinUSGraph ↔ (𝐺 ∈ USGraph ∧ (#‘(Vtx‘𝐺)) ∈ ℕ0)) | ||
Theorem | fusgrusgr 26259 | A finite simple graph is a simple graph. (Contributed by AV, 16-Jan-2020.) (Revised by AV, 21-Oct-2020.) |
⊢ (𝐺 ∈ FinUSGraph → 𝐺 ∈ USGraph) | ||
Theorem | opfusgr 26260 | A finite simple graph represented as ordered pair. (Contributed by AV, 23-Oct-2020.) |
⊢ ((𝑉 ∈ 𝑋 ∧ 𝐸 ∈ 𝑌) → (〈𝑉, 𝐸〉 ∈ FinUSGraph ↔ (〈𝑉, 𝐸〉 ∈ USGraph ∧ 𝑉 ∈ Fin))) | ||
Theorem | usgredgffibi 26261 | The number of edges in a simple graph is finite iff its edge function is finite. (Contributed by AV, 10-Jan-2020.) (Revised by AV, 22-Oct-2020.) |
⊢ 𝐼 = (iEdg‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ (𝐺 ∈ USGraph → (𝐸 ∈ Fin ↔ 𝐼 ∈ Fin)) | ||
Theorem | fusgredgfi 26262* | In a finite simple graph the number of edges which contain a given vertex is also finite. (Contributed by Alexander van der Vekens, 4-Jan-2018.) (Revised by AV, 21-Oct-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ ((𝐺 ∈ FinUSGraph ∧ 𝑁 ∈ 𝑉) → {𝑒 ∈ 𝐸 ∣ 𝑁 ∈ 𝑒} ∈ Fin) | ||
Theorem | usgr1v0e 26263 | The size of a (finite) simple graph with 1 vertex is 0. (Contributed by Alexander van der Vekens, 5-Jan-2018.) (Revised by AV, 22-Oct-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ ((𝐺 ∈ USGraph ∧ (#‘𝑉) = 1) → (#‘𝐸) = 0) | ||
Theorem | usgrfilem 26264* | In a finite simple graph, the number of edges is finite iff the number of edges not containing one of the vertices is finite. (Contributed by Alexander van der Vekens, 4-Jan-2018.) (Revised by AV, 9-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) & ⊢ 𝐹 = {𝑒 ∈ 𝐸 ∣ 𝑁 ∉ 𝑒} ⇒ ⊢ ((𝐺 ∈ FinUSGraph ∧ 𝑁 ∈ 𝑉) → (𝐸 ∈ Fin ↔ 𝐹 ∈ Fin)) | ||
Theorem | fusgrfisbase 26265 | Induction base for fusgrfis 26267. Main work is done in uhgr0v0e 26175. (Contributed by Alexander van der Vekens, 5-Jan-2018.) (Revised by AV, 23-Oct-2020.) |
⊢ (((𝑉 ∈ 𝑋 ∧ 𝐸 ∈ 𝑌) ∧ 〈𝑉, 𝐸〉 ∈ USGraph ∧ (#‘𝑉) = 0) → 𝐸 ∈ Fin) | ||
Theorem | fusgrfisstep 26266* | Induction step in fusgrfis 26267: In a finite simple graph, the number of edges is finite if the number of edges not containing one of the vertices is finite. (Contributed by Alexander van der Vekens, 5-Jan-2018.) (Revised by AV, 23-Oct-2020.) |
⊢ (((𝑉 ∈ 𝑋 ∧ 𝐸 ∈ 𝑌) ∧ 〈𝑉, 𝐸〉 ∈ FinUSGraph ∧ 𝑁 ∈ 𝑉) → (( I ↾ {𝑝 ∈ (Edg‘〈𝑉, 𝐸〉) ∣ 𝑁 ∉ 𝑝}) ∈ Fin → 𝐸 ∈ Fin)) | ||
Theorem | fusgrfis 26267 | A finite simple graph is of finite size, i.e. has a finite number of edges. (Contributed by Alexander van der Vekens, 6-Jan-2018.) (Revised by AV, 8-Nov-2020.) |
⊢ (𝐺 ∈ FinUSGraph → (Edg‘𝐺) ∈ Fin) | ||
Theorem | fusgrfupgrfs 26268 | A finite simple graph is a finite pseudograph of finite size. (Contributed by AV, 27-Dec-2021.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝐺) ⇒ ⊢ (𝐺 ∈ FinUSGraph → (𝐺 ∈ UPGraph ∧ 𝑉 ∈ Fin ∧ 𝐼 ∈ Fin)) | ||
Syntax | cnbgr 26269 | Extend class notation with neighbors (of a vertex in a graph). |
class NeighbVtx | ||
Definition | df-nbgr 26270* |
Define the (open) neighborhood resp. the class of all neighbors of a
vertex (in a graph), see definition in section I.1 of [Bollobas] p. 3 or
definition in section 1.1 of [Diestel]
p. 3. The neighborhood/neighbors
of a vertex are all (other) vertices which are connected with this
vertex by an edge. In contrast to a closed neighborhood, a vertex is
not a neighbor of itself. This definition is applicable even for
arbitrary hypergraphs.
Remark: To distinguish this definition from other definitions for neighborhoods resp. neighbors (e.g., nei in Topology, see df-nei 20950), the suffix Vtx is added to the class constant NeighbVtx. (Contributed by Alexander van der Vekens and Mario Carneiro, 7-Oct-2017.) (Revised by AV, 24-Oct-2020.) |
⊢ NeighbVtx = (𝑔 ∈ V, 𝑣 ∈ (Vtx‘𝑔) ↦ {𝑛 ∈ ((Vtx‘𝑔) ∖ {𝑣}) ∣ ∃𝑒 ∈ (Edg‘𝑔){𝑣, 𝑛} ⊆ 𝑒}) | ||
Theorem | nbgrprc0 26271 | The set of neighbors is empty if the graph 𝐺 or the vertex 𝑁 are proper classes. (Contributed by AV, 26-Oct-2020.) |
⊢ (¬ (𝐺 ∈ V ∧ 𝑁 ∈ V) → (𝐺 NeighbVtx 𝑁) = ∅) | ||
Theorem | nbgrcl 26272 | If a class 𝑋 has at least one neighbor, this class must be a vertex. (Contributed by AV, 6-Jun-2021.) (Revised by AV, 12-Feb-2022.) |
⊢ 𝑉 = (Vtx‘𝐺) ⇒ ⊢ (𝑁 ∈ (𝐺 NeighbVtx 𝑋) → 𝑋 ∈ 𝑉) | ||
Theorem | nbgrclOLD 26273 | Obsolete version of nbgrcl 26272 as of 12-Feb-2022. (Contributed by AV, 6-Jun-2021.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (𝑁 ∈ (𝐺 NeighbVtx 𝑋) → 𝑋 ∈ (Vtx‘𝐺)) | ||
Theorem | nbgrval 26274* | The set of neighbors of a vertex 𝑉 in a graph 𝐺. (Contributed by Alexander van der Vekens, 7-Oct-2017.) (Revised by AV, 24-Oct-2020.) (Revised by AV, 21-Mar-2021.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ (𝑁 ∈ 𝑉 → (𝐺 NeighbVtx 𝑁) = {𝑛 ∈ (𝑉 ∖ {𝑁}) ∣ ∃𝑒 ∈ 𝐸 {𝑁, 𝑛} ⊆ 𝑒}) | ||
Theorem | dfnbgr2 26275* | Alternate definition of the neighbors of a vertex breaking up the subset relationship of an unordered pair. (Contributed by AV, 15-Nov-2020.) (Revised by AV, 21-Mar-2021.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ (𝑁 ∈ 𝑉 → (𝐺 NeighbVtx 𝑁) = {𝑛 ∈ (𝑉 ∖ {𝑁}) ∣ ∃𝑒 ∈ 𝐸 (𝑁 ∈ 𝑒 ∧ 𝑛 ∈ 𝑒)}) | ||
Theorem | dfnbgr3 26276* | Alternate definition of the neighbors of a vertex using the edge function instead of the edges themselves (see also nbgrval 26274). (Contributed by Alexander van der Vekens, 17-Dec-2017.) (Revised by AV, 25-Oct-2020.) (Revised by AV, 21-Mar-2021.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐼 = (iEdg‘𝐺) ⇒ ⊢ ((𝑁 ∈ 𝑉 ∧ Fun 𝐼) → (𝐺 NeighbVtx 𝑁) = {𝑛 ∈ (𝑉 ∖ {𝑁}) ∣ ∃𝑖 ∈ dom 𝐼{𝑁, 𝑛} ⊆ (𝐼‘𝑖)}) | ||
Theorem | nbgrnvtx0 26277 | If a class 𝑋 is not a vertex of a graph 𝐺, then it has no neighbors in 𝐺. (Contributed by Alexander van der Vekens, 12-Oct-2017.) (Revised by AV, 26-Oct-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) ⇒ ⊢ (𝑋 ∉ 𝑉 → (𝐺 NeighbVtx 𝑋) = ∅) | ||
Theorem | nbgrel 26278* | Characterization of a neighbor 𝑁 of a vertex 𝑋 in a graph 𝐺. (Contributed by Alexander van der Vekens and Mario Carneiro, 9-Oct-2017.) (Revised by AV, 26-Oct-2020.) (Revised by AV, 12-Feb-2022.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ (𝑁 ∈ (𝐺 NeighbVtx 𝑋) ↔ ((𝑁 ∈ 𝑉 ∧ 𝑋 ∈ 𝑉) ∧ 𝑁 ≠ 𝑋 ∧ ∃𝑒 ∈ 𝐸 {𝑋, 𝑁} ⊆ 𝑒)) | ||
Theorem | nbgrelOLD 26279* | Obsolete version of nbgrel 26278 as of 12-Feb-2022. (Contributed by Alexander van der Vekens and Mario Carneiro, 9-Oct-2017.) (Revised by AV, 26-Oct-2020.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ (𝐺 ∈ 𝑊 → (𝐾 ∈ (𝐺 NeighbVtx 𝑁) ↔ ((𝐾 ∈ 𝑉 ∧ 𝑁 ∈ 𝑉) ∧ 𝐾 ≠ 𝑁 ∧ ∃𝑒 ∈ 𝐸 {𝑁, 𝐾} ⊆ 𝑒))) | ||
Theorem | nbgrisvtx 26280 | Every neighbor 𝑁 of a vertex 𝐾 is a vertex. (Contributed by Alexander van der Vekens, 12-Oct-2017.) (Revised by AV, 26-Oct-2020.) (Revised by AV, 12-Feb-2022.) |
⊢ 𝑉 = (Vtx‘𝐺) ⇒ ⊢ (𝑁 ∈ (𝐺 NeighbVtx 𝐾) → 𝑁 ∈ 𝑉) | ||
Theorem | nbgrssvtx 26281 | The neighbors of a vertex 𝐾 in a graph form a subset of all vertices of the graph. (Contributed by Alexander van der Vekens, 12-Oct-2017.) (Revised by AV, 26-Oct-2020.) (Revised by AV, 12-Feb-2022.) |
⊢ 𝑉 = (Vtx‘𝐺) ⇒ ⊢ (𝐺 NeighbVtx 𝐾) ⊆ 𝑉 | ||
Theorem | nbgrisvtxOLD 26282 | Obsolete version of nbgrisvtx 26280 as of 12-Feb-2022. (Contributed by Alexander van der Vekens, 12-Oct-2017.) (Revised by AV, 26-Oct-2020.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ 𝑉 = (Vtx‘𝐺) ⇒ ⊢ ((𝐺 ∈ 𝑊 ∧ 𝑁 ∈ (𝐺 NeighbVtx 𝐾)) → 𝑁 ∈ 𝑉) | ||
Theorem | nbgrssvtxOLD 26283 | Obsolete version of nbgrssvtx 26281 as of 12-Feb-2022. (Contributed by Alexander van der Vekens, 12-Oct-2017.) (Revised by AV, 26-Oct-2020.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ 𝑉 = (Vtx‘𝐺) ⇒ ⊢ (𝐺 ∈ 𝑊 → (𝐺 NeighbVtx 𝑁) ⊆ 𝑉) | ||
Theorem | nbuhgr 26284* | The set of neighbors of a vertex in a hypergraph. This version of nbgrval 26274 (with 𝑁 being an arbitrary set instead of being a vertex) only holds for classes whose edges are subsets of the set of vertices (hypergraphs!). (Contributed by AV, 26-Oct-2020.) (Proof shortened by AV, 15-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ ((𝐺 ∈ UHGraph ∧ 𝑁 ∈ 𝑋) → (𝐺 NeighbVtx 𝑁) = {𝑛 ∈ (𝑉 ∖ {𝑁}) ∣ ∃𝑒 ∈ 𝐸 {𝑁, 𝑛} ⊆ 𝑒}) | ||
Theorem | nbupgr 26285* | The set of neighbors of a vertex in a pseudograph. (Contributed by AV, 5-Nov-2020.) (Proof shortened by AV, 30-Dec-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ ((𝐺 ∈ UPGraph ∧ 𝑁 ∈ 𝑉) → (𝐺 NeighbVtx 𝑁) = {𝑛 ∈ (𝑉 ∖ {𝑁}) ∣ {𝑁, 𝑛} ∈ 𝐸}) | ||
Theorem | nbupgrel 26286 | A neighbor of a vertex in a pseudograph. (Contributed by AV, 5-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ (((𝐺 ∈ UPGraph ∧ 𝐾 ∈ 𝑉) ∧ (𝑁 ∈ 𝑉 ∧ 𝑁 ≠ 𝐾)) → (𝑁 ∈ (𝐺 NeighbVtx 𝐾) ↔ {𝑁, 𝐾} ∈ 𝐸)) | ||
Theorem | nbumgrvtx 26287* | The set of neighbors of a vertex in a multigraph. (Contributed by AV, 27-Nov-2020.) (Proof shortened by AV, 30-Dec-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ ((𝐺 ∈ UMGraph ∧ 𝑁 ∈ 𝑉) → (𝐺 NeighbVtx 𝑁) = {𝑛 ∈ 𝑉 ∣ {𝑁, 𝑛} ∈ 𝐸}) | ||
Theorem | nbumgr 26288* | The set of neighbors of an arbitrary class in a multigraph. (Contributed by AV, 27-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ (𝐺 ∈ UMGraph → (𝐺 NeighbVtx 𝑁) = {𝑛 ∈ 𝑉 ∣ {𝑁, 𝑛} ∈ 𝐸}) | ||
Theorem | nbusgrvtx 26289* | The set of neighbors of a vertex in a simple graph. (Contributed by Alexander van der Vekens, 9-Oct-2017.) (Revised by AV, 26-Oct-2020.) (Proof shortened by AV, 27-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ ((𝐺 ∈ USGraph ∧ 𝑁 ∈ 𝑉) → (𝐺 NeighbVtx 𝑁) = {𝑛 ∈ 𝑉 ∣ {𝑁, 𝑛} ∈ 𝐸}) | ||
Theorem | nbusgr 26290* | The set of neighbors of an arbitrary class in a simple graph. (Contributed by Alexander van der Vekens, 9-Oct-2017.) (Revised by AV, 26-Oct-2020.) (Proof shortened by AV, 27-Nov-2020.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ (𝐺 ∈ USGraph → (𝐺 NeighbVtx 𝑁) = {𝑛 ∈ 𝑉 ∣ {𝑁, 𝑛} ∈ 𝐸}) | ||
Theorem | nbgr2vtx1edg 26291* | If a graph has two vertices, and there is an edge between the vertices, then each vertex is the neighbor of the other vertex. (Contributed by AV, 2-Nov-2020.) (Revised by AV, 25-Mar-2021.) (Proof shortened by AV, 13-Feb-2022.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ (((#‘𝑉) = 2 ∧ 𝑉 ∈ 𝐸) → ∀𝑣 ∈ 𝑉 ∀𝑛 ∈ (𝑉 ∖ {𝑣})𝑛 ∈ (𝐺 NeighbVtx 𝑣)) | ||
Theorem | nbuhgr2vtx1edgblem 26292* | Lemma for nbuhgr2vtx1edgb 26293. This reverse direction of nbgr2vtx1edg 26291 only holds for classes whose edges are subsets of the set of vertices, which is the property of hypergraphs. (Contributed by AV, 2-Nov-2020.) (Proof shortened by AV, 13-Feb-2022.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ ((𝐺 ∈ UHGraph ∧ 𝑉 = {𝑎, 𝑏} ∧ 𝑎 ∈ (𝐺 NeighbVtx 𝑏)) → {𝑎, 𝑏} ∈ 𝐸) | ||
Theorem | nbuhgr2vtx1edgb 26293* | If a hypergraph has two vertices, and there is an edge between the vertices, then each vertex is the neighbor of the other vertex. (Contributed by AV, 2-Nov-2020.) (Proof shortened by AV, 13-Feb-2022.) |
⊢ 𝑉 = (Vtx‘𝐺) & ⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ ((𝐺 ∈ UHGraph ∧ (#‘𝑉) = 2) → (𝑉 ∈ 𝐸 ↔ ∀𝑣 ∈ 𝑉 ∀𝑛 ∈ (𝑉 ∖ {𝑣})𝑛 ∈ (𝐺 NeighbVtx 𝑣))) | ||
Theorem | nbusgreledg 26294 | A class/vertex is a neighbor of another class/vertex in a simple graph iff the vertices are endpoints of an edge. (Contributed by Alexander van der Vekens, 11-Oct-2017.) (Revised by AV, 26-Oct-2020.) |
⊢ 𝐸 = (Edg‘𝐺) ⇒ ⊢ (𝐺 ∈ USGraph → (𝑁 ∈ (𝐺 NeighbVtx 𝐾) ↔ {𝑁, 𝐾} ∈ 𝐸)) | ||
Theorem | uhgrnbgr0nb 26295* | A vertex which is not endpoint of an edge has no neighbor in a hypergraph. (Contributed by Alexander van der Vekens, 12-Oct-2017.) (Revised by AV, 26-Oct-2020.) |
⊢ ((𝐺 ∈ UHGraph ∧ ∀𝑒 ∈ (Edg‘𝐺)𝑁 ∉ 𝑒) → (𝐺 NeighbVtx 𝑁) = ∅) | ||
Theorem | nbgr0vtxlem 26296* | Lemma for nbgr0vtx 26297 and nbgr0edg 26298. (Contributed by AV, 15-Nov-2020.) |
⊢ (𝜑 → ∀𝑛 ∈ ((Vtx‘𝐺) ∖ {𝐾}) ¬ ∃𝑒 ∈ (Edg‘𝐺){𝐾, 𝑛} ⊆ 𝑒) ⇒ ⊢ (𝜑 → (𝐺 NeighbVtx 𝐾) = ∅) | ||
Theorem | nbgr0vtx 26297 | In a null graph (with no vertices), all neighborhoods are empty. (Contributed by AV, 15-Nov-2020.) |
⊢ ((Vtx‘𝐺) = ∅ → (𝐺 NeighbVtx 𝐾) = ∅) | ||
Theorem | nbgr0edg 26298 | In an empty graph (with no edges), every vertex has no neighbor. (Contributed by Alexander van der Vekens, 12-Oct-2017.) (Revised by AV, 26-Oct-2020.) (Proof shortened by AV, 15-Nov-2020.) |
⊢ ((Edg‘𝐺) = ∅ → (𝐺 NeighbVtx 𝐾) = ∅) | ||
Theorem | nbgr1vtx 26299 | In a graph with one vertex, all neighborhoods are empty. (Contributed by AV, 15-Nov-2020.) |
⊢ ((#‘(Vtx‘𝐺)) = 1 → (𝐺 NeighbVtx 𝐾) = ∅) | ||
Theorem | nbgrnself 26300* | A vertex in a graph is not a neighbor of itself. (Contributed by Alexander van der Vekens, 12-Oct-2017.) (Revised by AV, 3-Nov-2020.) (Revised by AV, 21-Mar-2021.) |
⊢ 𝑉 = (Vtx‘𝐺) ⇒ ⊢ ∀𝑣 ∈ 𝑉 𝑣 ∉ (𝐺 NeighbVtx 𝑣) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |