Theory Plane1

(*  Title:      Plane1.thy
    Author:     Gertrud Bauer, Tobias Nipkow

Fixing a single face and vertex in each refinement step.
*)

theory Plane1
imports Plane
begin

text‹This is an optimized definition of plane graphs and the one we
adopt as our point of reference. In every step only one fixed nonfinal
face (the smallest one) and one edge in that face are picked.›


definition minimalFace :: "face list  face" where
 "minimalFace  minimal (length  vertices)"

definition minimalVertex :: "graph  face  vertex" where
 "minimalVertex g f  minimal (height g) (vertices f)" 

definition next_plane :: "nat  graph  graph list" ("next'_plane⇘_") where
 "next_plane⇘pg 
     let fs = nonFinals g in
     if fs = [] then [] 
     else let f = minimalFace fs; v = minimalVertex g f in
          ⨆⇘i[3..<Suc(maxGon p)]generatePolygon i v f g"

definition PlaneGraphsP :: "nat  graph set" ("PlaneGraphs⇘_") where
"PlaneGraphs⇘p {g. Seed⇘p[next_plane⇘p]→* g  final g}"

definition PlaneGraphs :: "graph set" where
"PlaneGraphs  p. PlaneGraphs⇘p⇙"

end