Up to index of Isabelle/HOL/SizeChange
theory Graphs(* Title: HOL/Library/Graphs.thy ID: $Id: Graphs.thy,v 1.1 2007/11/06 16:44:54 krauss Exp $ Author: Alexander Krauss, TU Muenchen *) header {* General Graphs as Sets *} theory Graphs imports Main Misc_Tools Kleene_Algebras begin subsection {* Basic types, Size Change Graphs *} datatype ('a, 'b) graph = Graph "('a × 'b × 'a) set" fun dest_graph :: "('a, 'b) graph => ('a × 'b × 'a) set" where "dest_graph (Graph G) = G" lemma graph_dest_graph[simp]: "Graph (dest_graph G) = G" by (cases G) simp lemma split_graph_all: "(!!gr. PROP P gr) ≡ (!!set. PROP P (Graph set))" proof fix set assume "!!gr. PROP P gr" then show "PROP P (Graph set)" . next fix gr assume "!!set. PROP P (Graph set)" then have "PROP P (Graph (dest_graph gr))" . then show "PROP P gr" by simp qed definition has_edge :: "('n,'e) graph => 'n => 'e => 'n => bool" ("_ \<turnstile> _ \<leadsto>_ _") where "has_edge G n e n' = ((n, e, n') ∈ dest_graph G)" subsection {* Graph composition *} fun grcomp :: "('n, 'e::times) graph => ('n, 'e) graph => ('n, 'e) graph" where "grcomp (Graph G) (Graph H) = Graph {(p,b,q) | p b q. (∃k e e'. (p,e,k)∈G ∧ (k,e',q)∈H ∧ b = e * e')}" declare grcomp.simps[code del] lemma graph_ext: assumes "!!n e n'. has_edge G n e n' = has_edge H n e n'" shows "G = H" using assms by (cases G, cases H) (auto simp:split_paired_all has_edge_def) 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 lemmas [code func del] = graph_plus_def 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 (\<Inter>(dest_graph ` Gs))" Sup_graph_def: "Sup ≡ λGs. Graph (\<Union>(dest_graph ` Gs))" proof fix x y z :: "('a,'b) graph" fix A :: "('a, 'b) graph set" 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 .. { assume "x ≤ y" "y ≤ z" with order_trans show "x ≤ z" unfolding graph_leq_def . } { assume "x ≤ y" "y ≤ x" thus "x = y" unfolding graph_leq_def by (cases x, cases y) simp } show "inf x y ≤ x" "inf x y ≤ y" unfolding inf_graph_def graph_leq_def by auto { assume "x ≤ y" "x ≤ z" thus "x ≤ inf y z" unfolding inf_graph_def graph_leq_def by auto } show "x ≤ sup x y" "y ≤ sup x y" unfolding sup_graph_def graph_leq_def graph_plus_def by auto { assume "y ≤ x" "z ≤ x" thus "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 { assume "x ∈ A" thus "Inf A ≤ x" unfolding Inf_graph_def graph_leq_def by auto } { assume "!!x. x ∈ A ==> z ≤ x" thus "z ≤ Inf A" unfolding Inf_graph_def graph_leq_def by auto } { assume "x ∈ A" thus "x ≤ Sup A" unfolding Sup_graph_def graph_leq_def by auto } { assume "!!x. x ∈ A ==> x ≤ z" thus "Sup A ≤ z" unfolding Sup_graph_def graph_leq_def by auto } qed lemmas [code func del] = graph_leq_def graph_less_def inf_graph_def sup_graph_def Inf_graph_def Sup_graph_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) subsubsection {* 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 lemmas [code func del] = graph_mult_def instance graph :: (type, one) one graph_one_def: "1 == Graph { (x, 1, x) |x. True}" .. 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')" by (cases G, cases H) (auto simp:graph_mult_def has_edge_def image_def) lemma in_grunit: "has_edge 1 p b q = (p = q ∧ b = 1)" by (auto simp:graph_one_def has_edge_def) instance graph :: (type, semigroup_mult) semigroup_mult proof fix G1 G2 G3 :: "('a,'b) graph" show "G1 * G2 * G3 = G1 * (G2 * G3)" proof (rule graph_ext, rule trans) fix p J q show "has_edge ((G1 * G2) * G3) p J q = (∃G i H j I. has_edge G1 p G i ∧ has_edge G2 i H j ∧ has_edge G3 j I q ∧ J = (G * H) * I)" by (simp only:in_grcomp) blast show "… = has_edge (G1 * (G2 * G3)) p J q" by (simp only:in_grcomp mult_assoc) blast qed qed 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) "{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 "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 show "a * (b + c) = a * b + a * c" by (rule graph_ext, simp add:in_grcomp in_grplus) blast show "(0::('a,'b) graph) ≠ 1" unfolding graph_zero_def graph_one_def by simp 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 lemma graph_leqI: assumes "!!n e n'. has_edge G n e n' ==> has_edge H n e n'" shows "G ≤ H" using assms unfolding graph_leq_def has_edge_def by auto lemma in_graph_plusE: assumes "has_edge (G + H) n e n'" assumes "has_edge G n e n' ==> P" assumes "has_edge H n e n' ==> P" shows P using assms by (auto simp: in_grplus) lemma in_graph_compE: assumes GH: "has_edge (G * H) n e n'" obtains e1 k e2 where "has_edge G n e1 k" "has_edge H k e2 n'" "e = e1 * e2" using GH by (auto simp: in_grcomp) lemma assumes "x ∈ S k" shows "x ∈ (\<Union>k. S k)" using assms by blast lemma graph_union_least: assumes "!!n. Graph (G n) ≤ C" shows "Graph (\<Union>n. G n) ≤ C" using assms unfolding graph_leq_def by auto lemma Sup_graph_eq: "(SUP n. Graph (G n)) = Graph (\<Union>n. G n)" proof (rule order_antisym) show "(SUP n. Graph (G n)) ≤ Graph (\<Union>n. G n)" by (rule SUP_leI) (auto simp add: graph_leq_def) show "Graph (\<Union>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 by (cases G) simp lemma Sup_graph_eq2: "(SUP n. G n) = Graph (\<Union>n. dest_graph (G n))" using Sup_graph_eq[of "λn. dest_graph (G n)", simplified] by simp lemma in_SUP: "has_edge (SUP x. Gs x) p b q = (∃x. has_edge (Gs x) p b q)" unfolding Sup_graph_eq2 has_edge_leq graph_leq_def by simp instance graph :: (type, monoid_mult) kleene_by_complete_lattice proof fix a b c :: "('a, 'b) graph" show "a ≤ b <-> a + b = b" unfolding graph_leq_def graph_plus_def by (cases a, cases b) auto from order_less_le show "a < b <-> a ≤ b ∧ a ≠ b" . show "a * star b * c = (SUP n. a * b ^ n * c)" unfolding graph_star_def by (rule graph_ext) (force simp:in_SUP in_grcomp) qed lemma in_star: "has_edge (star G) a x b = (∃n. has_edge (G ^ n) a x b)" by (auto simp:graph_star_def in_SUP) lemma tcl_is_SUP: "tcl (G::('a::type, 'b::monoid_mult) graph) = (SUP n. G ^ (Suc n))" unfolding tcl_def using star_cont[of 1 G G] by (simp add:power_Suc power_commutes) lemma in_tcl: "has_edge (tcl G) a x b = (∃n>0. has_edge (G ^ n) a x b)" apply (auto simp: tcl_is_SUP in_SUP) apply (rule_tac x = "n - 1" in exI, auto) done subsection {* Infinite Paths *} types ('n, 'e) ipath = "('n × 'e) sequence" definition has_ipath :: "('n, 'e) graph => ('n, 'e) ipath => bool" where "has_ipath G p = (∀i. has_edge G (fst (p i)) (snd (p i)) (fst (p (Suc i))))" subsection {* Finite Paths *} types ('n, 'e) fpath = "('n × ('e × 'n) list)" inductive has_fpath :: "('n, 'e) graph => ('n, 'e) fpath => bool" for G :: "('n, 'e) graph" where has_fpath_empty: "has_fpath G (n, [])" | has_fpath_join: "[|G \<turnstile> n \<leadsto>e n'; has_fpath G (n', es)|] ==> has_fpath G (n, (e, n')#es)" definition "end_node p = (if snd p = [] then fst p else snd (snd p ! (length (snd p) - 1)))" definition path_nth :: "('n, 'e) fpath => nat => ('n × 'e × 'n)" where "path_nth p k = (if k = 0 then fst p else snd (snd p ! (k - 1)), snd p ! k)" lemma endnode_nth: assumes "length (snd p) = Suc k" shows "end_node p = snd (snd (path_nth p k))" using assms unfolding end_node_def path_nth_def by auto lemma path_nth_graph: assumes "k < length (snd p)" assumes "has_fpath G p" shows "(λ(n,e,n'). has_edge G n e n') (path_nth p k)" using assms proof (induct k arbitrary: p) case 0 thus ?case unfolding path_nth_def by (auto elim:has_fpath.cases) next case (Suc k p) from `has_fpath G p` show ?case proof (rule has_fpath.cases) case goal1 with Suc show ?case by simp next fix n e n' es assume st: "p = (n, (e, n') # es)" "G \<turnstile> n \<leadsto>e n'" "has_fpath G (n', es)" with Suc have "(λ(n, b, a). G \<turnstile> n \<leadsto>b a) (path_nth (n', es) k)" by simp with st show ?thesis by (cases k, auto simp:path_nth_def) qed qed lemma path_nth_connected: assumes "Suc k < length (snd p)" shows "fst (path_nth p (Suc k)) = snd (snd (path_nth p k))" using assms unfolding path_nth_def by auto definition path_loop :: "('n, 'e) fpath => ('n, 'e) ipath" ("omega") where "omega p ≡ (λi. (λ(n,e,n'). (n,e)) (path_nth p (i mod (length (snd p)))))" lemma fst_p0: "fst (path_nth p 0) = fst p" unfolding path_nth_def by simp lemma path_loop_connect: assumes "fst p = end_node p" and "0 < length (snd p)" (is "0 < ?l") shows "fst (path_nth p (Suc i mod (length (snd p)))) = snd (snd (path_nth p (i mod length (snd p))))" (is "… = snd (snd (path_nth p ?k))") proof - from `0 < ?l` have "i mod ?l < ?l" (is "?k < ?l") by simp show ?thesis proof (cases "Suc ?k < ?l") case True hence "Suc ?k ≠ ?l" by simp with path_nth_connected[OF True] show ?thesis by (simp add:mod_Suc) next case False with `?k < ?l` have wrap: "Suc ?k = ?l" by simp hence "fst (path_nth p (Suc i mod ?l)) = fst (path_nth p 0)" by (simp add: mod_Suc) also from fst_p0 have "… = fst p" . also have "… = end_node p" by fact also have "… = snd (snd (path_nth p ?k))" by (auto simp: endnode_nth wrap) finally show ?thesis . qed qed lemma path_loop_graph: assumes "has_fpath G p" and loop: "fst p = end_node p" and nonempty: "0 < length (snd p)" (is "0 < ?l") shows "has_ipath G (omega p)" proof - { fix i from `0 < ?l` have "i mod ?l < ?l" (is "?k < ?l") by simp from this and `has_fpath G p` have pk_G: "(λ(n,e,n'). has_edge G n e n') (path_nth p ?k)" by (rule path_nth_graph) from path_loop_connect[OF loop nonempty] pk_G have "has_edge G (fst (omega p i)) (snd (omega p i)) (fst (omega p (Suc i)))" unfolding path_loop_def has_edge_def split_def by simp } then show ?thesis by (auto simp:has_ipath_def) qed definition prod :: "('n, 'e::monoid_mult) fpath => 'e" where "prod p = foldr (op *) (map fst (snd p)) 1" lemma prod_simps[simp]: "prod (n, []) = 1" "prod (n, (e,n')#es) = e * (prod (n',es))" unfolding prod_def by simp_all lemma power_induces_path: assumes a: "has_edge (A ^ k) n G m" obtains p where "has_fpath A p" and "n = fst p" "m = end_node p" and "G = prod p" and "k = length (snd p)" using a proof (induct k arbitrary:m n G thesis) case (0 m n G) let ?p = "(n, [])" from 0 have "has_fpath A ?p" "m = end_node ?p" "G = prod ?p" by (auto simp:in_grunit end_node_def intro:has_fpath.intros) thus ?case using 0 by (auto simp:end_node_def) next case (Suc k m n G) hence "has_edge (A * A ^ k) n G m" by (simp add:power_Suc power_commutes) then obtain G' H j where a_A: "has_edge A n G' j" and H_pow: "has_edge (A ^ k) j H m" and [simp]: "G = G' * H" by (auto simp:in_grcomp) from H_pow and Suc obtain p where p_path: "has_fpath A p" and [simp]: "j = fst p" "m = end_node p" "H = prod p" "k = length (snd p)" by blast let ?p' = "(n, (G', j)#snd p)" from a_A and p_path have "has_fpath A ?p'" "m = end_node ?p'" "G = prod ?p'" by (auto simp:end_node_def nth.simps intro:has_fpath.intros split:nat.split) thus ?case using Suc by auto qed subsection {* Sub-Paths *} definition sub_path :: "('n, 'e) ipath => nat => nat => ('n, 'e) fpath" ("(_〈_,_〉)") where "p〈i,j〉 = (fst (p i), map (λk. (snd (p k), fst (p (Suc k)))) [i ..< j])" lemma sub_path_is_path: assumes ipath: "has_ipath G p" assumes l: "i ≤ j" shows "has_fpath G (p〈i,j〉)" using l proof (induct i rule:inc_induct) case base show ?case by (auto simp:sub_path_def intro:has_fpath.intros) next case (step i) with ipath upt_rec[of i j] show ?case by (auto simp:sub_path_def has_ipath_def intro:has_fpath.intros) qed lemma sub_path_start[simp]: "fst (p〈i,j〉) = fst (p i)" by (simp add:sub_path_def) lemma nth_upto[simp]: "k < j - i ==> [i ..< j] ! k = i + k" by (induct k) auto lemma sub_path_end[simp]: "i < j ==> end_node (p〈i,j〉) = fst (p j)" by (auto simp:sub_path_def end_node_def) lemma foldr_map: "foldr f (map g xs) = foldr (f o g) xs" by (induct xs) auto lemma upto_append[simp]: assumes "i ≤ j" "j ≤ k" shows "[ i ..< j ] @ [j ..< k] = [i ..< k]" using assms and upt_add_eq_append[of i j "k - j"] by simp lemma foldr_monoid: "foldr (op *) xs 1 * foldr (op *) ys 1 = foldr (op *) (xs @ ys) (1::'a::monoid_mult)" by (induct xs) (auto simp:mult_assoc) lemma sub_path_prod: assumes "i < j" assumes "j < k" shows "prod (p〈i,k〉) = prod (p〈i,j〉) * prod (p〈j,k〉)" using assms unfolding prod_def sub_path_def by (simp add:map_compose[symmetric] comp_def) (simp only:foldr_monoid map_append[symmetric] upto_append) lemma path_acgpow_aux: assumes "length es = l" assumes "has_fpath G (n,es)" shows "has_edge (G ^ l) n (prod (n,es)) (end_node (n,es))" using assms proof (induct l arbitrary:n es) case 0 thus ?case by (simp add:in_grunit end_node_def) next case (Suc l n es) hence "es ≠ []" by auto let ?n' = "snd (hd es)" let ?es' = "tl es" let ?e = "fst (hd es)" from Suc have len: "length ?es' = l" by auto from Suc have [simp]: "end_node (n, es) = end_node (?n', ?es')" by (cases es) (auto simp:end_node_def nth.simps split:nat.split) from `has_fpath G (n,es)` have "has_fpath G (?n', ?es')" by (rule has_fpath.cases) (auto intro:has_fpath.intros) with Suc len have "has_edge (G ^ l) ?n' (prod (?n', ?es')) (end_node (?n', ?es'))" by auto moreover from `es ≠ []` have "prod (n, es) = ?e * (prod (?n', ?es'))" by (cases es) auto moreover from `has_fpath G (n,es)` have c:"has_edge G n ?e ?n'" by (rule has_fpath.cases) (insert `es ≠ []`, auto) ultimately show ?case unfolding power_Suc by (auto simp:in_grcomp) qed lemma path_acgpow: "has_fpath G p ==> has_edge (G ^ length (snd p)) (fst p) (prod p) (end_node p)" by (cases p) (rule path_acgpow_aux[of "snd p" "length (snd p)" _ "fst p", simplified]) lemma star_paths: "has_edge (star G) a x b = (∃p. has_fpath G p ∧ a = fst p ∧ b = end_node p ∧ x = prod p)" proof assume "has_edge (star G) a x b" then obtain n where pow: "has_edge (G ^ n) a x b" by (auto simp:in_star) then obtain p where "has_fpath G p" "a = fst p" "b = end_node p" "x = prod p" by (rule power_induces_path) thus "∃p. has_fpath G p ∧ a = fst p ∧ b = end_node p ∧ x = prod p" by blast next assume "∃p. has_fpath G p ∧ a = fst p ∧ b = end_node p ∧ x = prod p" then obtain p where "has_fpath G p" "a = fst p" "b = end_node p" "x = prod p" by blast hence "has_edge (G ^ length (snd p)) a x b" by (auto intro:path_acgpow) thus "has_edge (star G) a x b" by (auto simp:in_star) qed lemma plus_paths: "has_edge (tcl G) a x b = (∃p. has_fpath G p ∧ a = fst p ∧ b = end_node p ∧ x = prod p ∧ 0 < length (snd p))" proof assume "has_edge (tcl G) a x b" then obtain n where pow: "has_edge (G ^ n) a x b" and "0 < n" by (auto simp:in_tcl) from pow obtain p where "has_fpath G p" "a = fst p" "b = end_node p" "x = prod p" "n = length (snd p)" by (rule power_induces_path) with `0 < n` show "∃p. has_fpath G p ∧ a = fst p ∧ b = end_node p ∧ x = prod p ∧ 0 < length (snd p) " by blast next assume "∃p. has_fpath G p ∧ a = fst p ∧ b = end_node p ∧ x = prod p ∧ 0 < length (snd p)" then obtain p where "has_fpath G p" "a = fst p" "b = end_node p" "x = prod p" "0 < length (snd p)" by blast hence "has_edge (G ^ length (snd p)) a x b" by (auto intro:path_acgpow) with `0 < length (snd p)` show "has_edge (tcl G) a x b" by (auto simp:in_tcl) qed definition "contract s p = (λi. (fst (p (s i)), prod (p〈s i,s (Suc i)〉)))" lemma ipath_contract: assumes [simp]: "increasing s" assumes ipath: "has_ipath G p" shows "has_ipath (tcl G) (contract s p)" unfolding has_ipath_def proof fix i let ?p = "p〈s i,s (Suc i)〉" from increasing_strict have "fst (p (s (Suc i))) = end_node ?p" by simp moreover from increasing_strict[of s i "Suc i"] have "snd ?p ≠ []" by (simp add:sub_path_def) moreover from ipath increasing_weak[of s] have "has_fpath G ?p" by (rule sub_path_is_path) auto ultimately show "has_edge (tcl G) (fst (contract s p i)) (snd (contract s p i)) (fst (contract s p (Suc i)))" unfolding contract_def plus_paths by (intro exI) auto qed lemma prod_unfold: "i < j ==> prod (p〈i,j〉) = snd (p i) * prod (p〈Suc i, j〉)" unfolding prod_def by (simp add:sub_path_def upt_rec[of "i" j]) lemma sub_path_loop: assumes "0 < k" assumes k: "k = length (snd loop)" assumes loop: "fst loop = end_node loop" shows "(omega loop)〈k * i,k * Suc i〉 = loop" (is "?ω = loop") proof (rule prod_eqI) show "fst ?ω = fst loop" by (auto simp:path_loop_def path_nth_def split_def k) show "snd ?ω = snd loop" proof (rule nth_equalityI[rule_format]) show leneq: "length (snd ?ω) = length (snd loop)" unfolding sub_path_def k by simp fix j assume "j < length (snd (?ω))" with leneq and k have "j < k" by simp have a: "!!i. fst (path_nth loop (Suc i mod k)) = snd (snd (path_nth loop (i mod k)))" unfolding k apply (rule path_loop_connect[OF loop]) using `0 < k` and k apply auto done from `j < k` show "snd ?ω ! j = snd loop ! j" unfolding sub_path_def apply (simp add:path_loop_def split_def add_ac) apply (simp add:a k[symmetric]) apply (simp add:path_nth_def) done qed qed end
lemma graph_dest_graph:
Graph (dest_graph G) = G
lemma split_graph_all:
(!!gr. PROP P gr) == (!!set. PROP P (Graph set))
lemma graph_ext:
(!!n e n'. G \<turnstile> n \<leadsto>e n' = H \<turnstile> n \<leadsto>e n')
==> G = H
lemma
G + H == Graph (dest_graph G ∪ dest_graph H)
lemma
G ≤ H == dest_graph G ⊆ dest_graph H
G < H == dest_graph G ⊂ dest_graph H
inf G H == Graph (dest_graph G ∩ dest_graph H)
sup G H == G + H
Inf == λGs. Graph (Inter (dest_graph ` Gs))
Sup == λGs. Graph (Union (dest_graph ` Gs))
lemma in_grplus:
G + H \<turnstile> p \<leadsto>b q =
(G \<turnstile> p \<leadsto>b q ∨ H \<turnstile> p \<leadsto>b q)
lemma in_grzero:
0 \<turnstile> p \<leadsto>b q = False
lemma
G * H == grcomp G H
lemma in_grcomp:
G * H \<turnstile> p \<leadsto>b q =
(∃k e e'.
G \<turnstile> p \<leadsto>e k ∧
H \<turnstile> k \<leadsto>e' q ∧ b = e * e')
lemma in_grunit:
1 \<turnstile> p \<leadsto>b q = (p = q ∧ b = (1::'b))
lemma graph_leqI:
(!!n e n'. G \<turnstile> n \<leadsto>e n' ==> H \<turnstile> n \<leadsto>e n')
==> G ≤ H
lemma in_graph_plusE:
[| G + H \<turnstile> n \<leadsto>e n'; G \<turnstile> n \<leadsto>e n' ==> P;
H \<turnstile> n \<leadsto>e n' ==> P |]
==> P
lemma in_graph_compE:
[| G * H \<turnstile> n \<leadsto>e n';
!!e1 k e2.
[| G \<turnstile> n \<leadsto>e1 k; H \<turnstile> k \<leadsto>e2 n';
e = e1 * e2 |]
==> thesis |]
==> thesis
lemma
x ∈ S k ==> x ∈ (UN k. S k)
lemma graph_union_least:
(!!n. Graph (G n) ≤ C) ==> Graph (UN n. G n) ≤ C
lemma Sup_graph_eq:
(SUP n. Graph (G n)) = Graph (UN n. G n)
lemma has_edge_leq:
G \<turnstile> p \<leadsto>b q = (Graph {(p, b, q)} ≤ G)
lemma Sup_graph_eq2:
(SUP n. G n) = Graph (UN n. dest_graph (G n))
lemma in_SUP:
SUP x. Gs x \<turnstile> p \<leadsto>b q =
(∃x. Gs x \<turnstile> p \<leadsto>b q)
lemma in_star:
star G \<turnstile> a \<leadsto>x b = (∃n. G ^ n \<turnstile> a \<leadsto>x b)
lemma tcl_is_SUP:
tcl G = (SUP n. G ^ Suc n)
lemma in_tcl:
tcl G \<turnstile> a \<leadsto>x b = (∃n>0. G ^ n \<turnstile> a \<leadsto>x b)
lemma endnode_nth:
length (snd p) = Suc k ==> end_node p = snd (snd (path_nth p k))
lemma path_nth_graph:
[| k < length (snd p); has_fpath G p |]
==> (λ(n, e, n'). G \<turnstile> n \<leadsto>e n') (path_nth p k)
lemma path_nth_connected:
Suc k < length (snd p) ==> fst (path_nth p (Suc k)) = snd (snd (path_nth p k))
lemma fst_p0:
fst (path_nth p 0) = fst p
lemma path_loop_connect:
[| fst p = end_node p; 0 < length (snd p) |]
==> fst (path_nth p (Suc i mod length (snd p))) =
snd (snd (path_nth p (i mod length (snd p))))
lemma path_loop_graph:
[| has_fpath G p; fst p = end_node p; 0 < length (snd p) |]
==> has_ipath G (omega p)
lemma prod_simps:
prod (n, []) = (1::'a)
prod (n, (e, n') # es) = e * prod (n', es)
lemma power_induces_path:
[| A ^ k \<turnstile> n \<leadsto>G m;
!!p. [| has_fpath A p; n = fst p; m = end_node p; G = prod p;
k = length (snd p) |]
==> thesis |]
==> thesis
lemma sub_path_is_path:
[| has_ipath G p; i ≤ j |] ==> has_fpath G p〈i,j〉
lemma sub_path_start:
fst p〈i,j〉 = fst (p i)
lemma nth_upto:
k < j - i ==> [i..<j] ! k = i + k
lemma sub_path_end:
i < j ==> end_node p〈i,j〉 = fst (p j)
lemma foldr_map:
foldr f (map g xs) = foldr (f o g) xs
lemma upto_append:
[| i ≤ j; j ≤ k |] ==> [i..<j] @ [j..<k] = [i..<k]
lemma foldr_monoid:
foldr op * xs (1::'a) * foldr op * ys (1::'a) = foldr op * (xs @ ys) (1::'a)
lemma sub_path_prod:
[| i < j; j < k |] ==> prod p〈i,k〉 = prod p〈i,j〉 * prod p〈j,k〉
lemma path_acgpow_aux:
[| length es = l; has_fpath G (n, es) |]
==> G ^ l \<turnstile> n \<leadsto>prod (n, es) end_node (n, es)
lemma path_acgpow:
has_fpath G p
==> G ^ length (snd p) \<turnstile> fst p \<leadsto>prod p end_node p
lemma star_paths:
star G \<turnstile> a \<leadsto>x b =
(∃p. has_fpath G p ∧ a = fst p ∧ b = end_node p ∧ x = prod p)
lemma plus_paths:
tcl G \<turnstile> a \<leadsto>x b =
(∃p. has_fpath G p ∧
a = fst p ∧ b = end_node p ∧ x = prod p ∧ 0 < length (snd p))
lemma ipath_contract:
[| increasing s; has_ipath G p |] ==> has_ipath (tcl G) (contract s p)
lemma prod_unfold:
i < j ==> prod p〈i,j〉 = snd (p i) * prod p〈Suc i,j〉
lemma sub_path_loop:
[| 0 < k; k = length (snd loop); fst loop = end_node loop |]
==> omega loop〈k * i,k * Suc i〉 = loop