# HG changeset patch # User krauss # Date 1176234758 -7200 # Node ID 7e35b6c8ab5b10c0da18676fd94b58cb9d8ae12a # Parent a2967023d674d8aedf744013b016122f64115bd3 some restructuring diff -r a2967023d674 -r 7e35b6c8ab5b src/HOL/Library/Graphs.thy --- a/src/HOL/Library/Graphs.thy Tue Apr 10 21:51:08 2007 +0200 +++ b/src/HOL/Library/Graphs.thy Tue Apr 10 21:52:38 2007 +0200 @@ -60,21 +60,87 @@ by (cases G, cases H, auto simp:split_paired_all has_edge_def) -instance graph :: (type, times) times - graph_mult_def: "G * H == grcomp G H" .. +instance graph :: (type, type) "{comm_monoid_add}" + graph_zero_def: "0 == Graph {}" + graph_plus_def: "G + H == Graph (dest_graph G \ dest_graph H)" +proof + fix x y z :: "('a,'b) graph" + + show "x + y + z = x + (y + z)" + and "x + y = y + x" + and "0 + x = x" + unfolding graph_plus_def graph_zero_def + by auto +qed + + +instance graph :: (type, type) "{distrib_lattice, complete_lattice}" + graph_leq_def: "G \ H \ dest_graph G \ dest_graph H" + graph_less_def: "G < H \ dest_graph G \ dest_graph H" + "inf G H \ Graph (dest_graph G \ dest_graph H)" + "sup G H \ G + H" + Inf_graph_def: "Inf == \Gs. Graph (\(dest_graph ` Gs))" +proof + fix x y z :: "('a,'b) graph" + + show "(x < y) = (x \ y \ x \ y)" + unfolding graph_leq_def graph_less_def + by (cases x, cases y) auto + + show "x \ x" unfolding graph_leq_def .. + + show "\x \ y; y \ x\ \ x = y" unfolding graph_leq_def + by (cases x, cases y) simp + + show "inf x y \ x" "inf x y \ y" + "\x \ y; x \ z\ \ x \ inf y z" + unfolding inf_graph_def graph_leq_def by auto + + show "x \ sup x y" "y \ sup x y" + "\y \ x; z \ x\ \ sup y z \ x" + unfolding sup_graph_def graph_leq_def graph_plus_def by auto + + show "sup x (inf y z) = inf (sup x y) (sup x z)" + unfolding inf_graph_def sup_graph_def graph_leq_def graph_plus_def by auto + + fix A :: "('a, 'b) graph set" + show "x \ A \ Inf A \ x" + and "(\x. x \ A \ z \ x) \ z \ Inf A" + unfolding Inf_graph_def graph_leq_def by auto + + from order_trans + show "\x \ y; y \ z\ \ x \ z" + unfolding graph_leq_def . +qed auto + +lemma in_grplus: + "has_edge (G + H) p b q = (has_edge G p b q \ has_edge H p b q)" + by (cases G, cases H, auto simp:has_edge_def graph_plus_def) + +lemma in_grzero: + "has_edge 0 p b q = False" + by (simp add:graph_zero_def has_edge_def) + + + +subsection {* Multiplicative Structure *} + +instance graph :: (type, times) mult_zero + graph_mult_def: "G * H == grcomp G H" +proof + fix a :: "('a, 'b) graph" + + show "0 * a = 0" + unfolding graph_mult_def graph_zero_def + by (cases a) (simp add:grcomp.simps) + show "a * 0 = 0" + unfolding graph_mult_def graph_zero_def + by (cases a) (simp add:grcomp.simps) +qed instance graph :: (type, one) one graph_one_def: "1 == Graph { (x, 1, x) |x. True}" .. -instance graph :: (type, type) zero - graph_zero_def: "0 == Graph {}" .. - -instance graph :: (type, type) plus - graph_plus_def: "G + H == Graph (dest_graph G \ dest_graph H)" .. - - -subsection {* Simprules for the graph operations *} - lemma in_grcomp: "has_edge (G * H) p b q = (\k e e'. has_edge G p e k \ has_edge H k e' q \ b = e * e')" @@ -84,15 +150,6 @@ "has_edge 1 p b q = (p = q \ b = 1)" by (auto simp:graph_one_def has_edge_def) -lemma in_grplus: - "has_edge (G + H) p b q = (has_edge G p b q \ has_edge H p b q)" - by (cases G, cases H, auto simp:has_edge_def graph_plus_def) - -lemma in_grzero: - "has_edge 0 p b q = False" - by (simp add:graph_zero_def has_edge_def) - - instance graph :: (type, semigroup_mult) semigroup_mult proof fix G1 G2 G3 :: "('a,'b) graph" @@ -112,80 +169,22 @@ qed qed -instance graph :: (type, monoid_mult) monoid_mult -proof - fix G1 G2 G3 :: "('a,'b) graph" - - show "1 * G1 = G1" - by (rule graph_ext) (auto simp:in_grcomp in_grunit) - show "G1 * 1 = G1" - by (rule graph_ext) (auto simp:in_grcomp in_grunit) -qed - - -lemma grcomp_rdist: - fixes G :: "('a::type, 'b::semigroup_mult) graph" - shows "G * (H + I) = G * H + G * I" - by (rule graph_ext, simp add:in_grcomp in_grplus) blast - -lemma grcomp_ldist: - fixes G :: "('a::type, 'b::semigroup_mult) graph" - shows "(G + H) * I = G * I + H * I" - by (rule graph_ext, simp add:in_grcomp in_grplus) blast - fun grpow :: "nat \ ('a::type, 'b::monoid_mult) graph \ ('a, 'b) graph" where "grpow 0 A = 1" | "grpow (Suc n) A = A * (grpow n A)" - -instance graph :: (type, monoid_mult) recpower - graph_pow_def: "A ^ n == grpow n A" - by default (simp_all add:graph_pow_def) - -subsection {* Order on Graphs *} - -instance graph :: (type, type) order - graph_leq_def: "G \ H \ dest_graph G \ dest_graph H" - graph_less_def: "G < H \ dest_graph G \ dest_graph H" -proof - fix x y z :: "('a,'b) graph" - - show "x \ x" unfolding graph_leq_def .. - - from order_trans - show "\x \ y; y \ z\ \ x \ z" unfolding graph_leq_def . - - show "\x \ y; y \ x\ \ x = y" unfolding graph_leq_def - by (cases x, cases y) simp - - show "(x < y) = (x \ y \ x \ y)" - unfolding graph_leq_def graph_less_def - by (cases x, cases y) auto -qed - -instance graph :: (type, type) distrib_lattice - "inf G H \ Graph (dest_graph G \ dest_graph H)" - "sup G H \ G + H" - by default (auto simp add: split_graph_all graph_plus_def inf_graph_def sup_graph_def graph_leq_def graph_less_def) - -instance graph :: (type, type) complete_lattice - Inf_graph_def: "Inf == \Gs. Graph (\(dest_graph ` Gs))" - by default (auto simp: Inf_graph_def graph_leq_def le_fun_def le_bool_def in_grplus has_edge_def) - -instance graph :: (type, monoid_mult) semiring_1 +instance graph :: (type, monoid_mult) + "{semiring_1,idem_add,recpower,star}" + graph_pow_def: "A ^ n == grpow n A" + graph_star_def: "star G == (SUP n. G ^ n)" proof fix a b c :: "('a, 'b) graph" - - show "a + b + c = a + (b + c)" - and "a + b = b + a" unfolding graph_plus_def - by auto - - show "0 + a = a" unfolding graph_zero_def graph_plus_def - by simp - - show "0 * a = 0" "a * 0 = 0" unfolding graph_zero_def graph_mult_def - by (cases a, simp)+ + + show "1 * a = a" + by (rule graph_ext) (auto simp:in_grcomp in_grunit) + show "a * 1 = a" + by (rule graph_ext) (auto simp:in_grcomp in_grunit) show "(a + b) * c = a * c + b * c" by (rule graph_ext, simp add:in_grcomp in_grplus) blast @@ -195,22 +194,14 @@ show "(0::('a,'b) graph) \ 1" unfolding graph_zero_def graph_one_def by simp -qed -instance graph :: (type, monoid_mult) idem_add -proof - fix a :: "('a, 'b) graph" show "a + a = a" unfolding graph_plus_def by simp + + show "a ^ 0 = 1" "\n. a ^ (Suc n) = a * a ^ n" + unfolding graph_pow_def by simp_all qed -(* define star on graphs *) - - -instance graph :: (type, monoid_mult) star - graph_star_def: "star G == (SUP n. G ^ n)" .. - - lemma graph_leqI: assumes "\n e n'. has_edge G n e n' \ has_edge H n e n'" shows "G \ H" @@ -227,8 +218,6 @@ using prems by (auto simp: in_grplus) - - lemma assumes "x \ S k" shows "x \ (\k. S k)" @@ -242,15 +231,13 @@ lemma Sup_graph_eq: "(SUP n. Graph (G n)) = Graph (\n. G n)" - unfolding SUPR_def - apply (rule order_antisym) - apply (rule Sup_least) - apply auto - apply (simp add:graph_leq_def) - apply auto - apply (rule graph_union_least) - apply (rule Sup_upper) - by auto +proof (rule order_antisym) + show "(SUP n. Graph (G n)) \ Graph (\n. G n)" + by (rule SUP_leI) (auto simp add: graph_leq_def) + + show "Graph (\n. G n) \ (SUP n. Graph (G n))" + by (rule graph_union_least, rule le_SUPI', rule) +qed lemma has_edge_leq: "has_edge G p b q = (Graph {(p,b,q)} \ G)" unfolding has_edge_def graph_leq_def