Theory Plane1

Up to index of Isabelle/HOL/Tame

theory Plane1
imports Plane
begin

(*  Title:      Plane1.thy
    ID:         $Id: Plane1.thy,v 1.1 2006/01/04 13:43:59 nipkow Exp $
    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. *}


constdefs minimalFace:: "face list => face"
 "minimalFace ≡ minimal (length o vertices)"

constdefs minimalVertex :: "graph => face => vertex"
 "minimalVertex g f ≡ minimal (height g) (vertices f)" 

constdefs next_plane :: "nat => graph => graph list" ("next'_plane_")
 "next_planep g ≡
     let fs = nonFinals g in
     if fs = [] then [] 
     else let f = minimalFace fs; v = minimalVertex g f in
          \<Squnion>i∈[3..maxGon p] generatePolygon i v f g"

constdefs
 PlaneGraphsP  :: "nat => graph set" ("PlaneGraphs_")
"PlaneGraphsp ≡ {g. Seedp [next_planep]->* g ∧ final g}"

 PlaneGraphs :: "graph set"
"PlaneGraphs ≡ \<Union>p. PlaneGraphsp"

end