Skip to content

Commit

Permalink
Add a couple of simple automatic rewrites for generic graphs
Browse files Browse the repository at this point in the history
  • Loading branch information
mn200 committed Jan 8, 2023
1 parent 6b449d4 commit fcd4fa5
Showing 1 changed file with 13 additions and 0 deletions.
13 changes: 13 additions & 0 deletions examples/generic_finite_graphs/genericGraphScript.sml
Expand Up @@ -396,6 +396,13 @@ Proof
simp[emptyG0_def]
QED

Theorem nlabelfun_empty[simp]:
nlabelfun emptyG = (λn. ARB)
Proof
simp[nlabelfun_def, emptyG_def, #repabs_pseudo_id tydefrec] >>
simp[emptyG0_def, FUN_EQ_THM]
QED

Definition adjacent_def:
adjacent G n1 n2 ⇔ ∃l. (n1, n2, l) ∈ edges G
End
Expand Down Expand Up @@ -682,6 +689,12 @@ Proof
simp[oneEdge_max_def, itself2set_def]
QED

Theorem oneEdge_max_fdirgraph[simp]:
¬oneEdge_max (g : (α,β,γ)fdirgraph)
Proof
simp[oneEdge_max_def, itself2set_def]
QED

Theorem selfloops_ok_graph[simp]:
selfloops_ok (g : ('a,'d,'ec,'el,'nf,'nl,unit) graph)
Proof
Expand Down

0 comments on commit fcd4fa5

Please sign in to comment.