Up to index of Isabelle/HOL/SizeChange
theory Kleene_Algebras(* Title: HOL/Library/Kleene_Algebras.thy ID: $Id: Kleene_Algebras.thy,v 1.1 2007/11/06 16:44:54 krauss Exp $ Author: Alexander Krauss, TU Muenchen *) header "Kleene Algebras" theory Kleene_Algebras imports Main begin text {* A type class of kleene algebras *} class star = type + fixes star :: "'a => 'a" class idem_add = ab_semigroup_add + assumes add_idem [simp]: "x + x = x" lemma add_idem2[simp]: "(x::'a::idem_add) + (x + y) = x + y" unfolding add_assoc[symmetric] by simp class order_by_add = idem_add + ord + assumes order_def: "a ≤ b <-> a + b = b" assumes strict_order_def: "a < b <-> a ≤ b ∧ a ≠ b" lemma ord_simp1[simp]: "(x::'a::order_by_add) ≤ y ==> x + y = y" unfolding order_def . lemma ord_simp2[simp]: "(x::'a::order_by_add) ≤ y ==> y + x = y" unfolding order_def add_commute . lemma ord_intro: "(x::'a::order_by_add) + y = y ==> x ≤ y" unfolding order_def . instance order_by_add ⊆ order proof fix x y z :: 'a show "x ≤ x" unfolding order_def by simp show "[|x ≤ y; y ≤ z|] ==> x ≤ z" proof (rule ord_intro) assume "x ≤ y" "y ≤ z" have "x + z = x + y + z" by (simp add:`y ≤ z` add_assoc) also have "… = y + z" by (simp add:`x ≤ y`) also have "… = z" by (simp add:`y ≤ z`) finally show "x + z = z" . qed show "[|x ≤ y; y ≤ x|] ==> x = y" unfolding order_def by (simp add:add_commute) show "x < y <-> x ≤ y ∧ x ≠ y" by (fact strict_order_def) qed class pre_kleene = semiring_1 + order_by_add instance pre_kleene ⊆ pordered_semiring proof fix x y z :: 'a assume "x ≤ y" show "z + x ≤ z + y" proof (rule ord_intro) have "z + x + (z + y) = x + y + z" by (simp add:add_ac) also have "… = z + y" by (simp add:`x ≤ y` add_ac) finally show "z + x + (z + y) = z + y" . qed show "z * x ≤ z * y" proof (rule ord_intro) from `x ≤ y` have "z * (x + y) = z * y" by simp thus "z * x + z * y = z * y" by (simp add:right_distrib) qed show "x * z ≤ y * z" proof (rule ord_intro) from `x ≤ y` have "(x + y) * z = y * z" by simp thus "x * z + y * z = y * z" by (simp add:left_distrib) qed qed class kleene = pre_kleene + star + assumes star1: "1 + a * star a ≤ star a" and star2: "1 + star a * a ≤ star a" and star3: "a * x ≤ x ==> star a * x ≤ x" and star4: "x * a ≤ x ==> x * star a ≤ x" class kleene_by_complete_lattice = pre_kleene + complete_lattice + recpower + star + assumes star_cont: "a * star b * c = SUPR UNIV (λn. a * b ^ n * c)" lemma plus_leI: fixes x :: "'a :: order_by_add" shows "x ≤ z ==> y ≤ z ==> x + y ≤ z" unfolding order_def by (simp add:add_assoc) lemma le_SUPI': fixes l :: "'a :: complete_lattice" assumes "l ≤ M i" shows "l ≤ (SUP i. M i)" using assms by (rule order_trans) (rule le_SUPI [OF UNIV_I]) lemma zero_minimum[simp]: "(0::'a::pre_kleene) ≤ x" unfolding order_def by simp instance kleene_by_complete_lattice ⊆ kleene proof fix a x :: 'a have [simp]: "1 ≤ star a" unfolding star_cont[of 1 a 1, simplified] by (subst power_0[symmetric]) (rule le_SUPI [OF UNIV_I]) show "1 + a * star a ≤ star a" apply (rule plus_leI, simp) apply (simp add:star_cont[of a a 1, simplified]) apply (simp add:star_cont[of 1 a 1, simplified]) apply (subst power_Suc[symmetric]) by (intro SUP_leI le_SUPI UNIV_I) show "1 + star a * a ≤ star a" apply (rule plus_leI, simp) apply (simp add:star_cont[of 1 a a, simplified]) apply (simp add:star_cont[of 1 a 1, simplified]) by (auto intro: SUP_leI le_SUPI UNIV_I simp add: power_Suc[symmetric] power_commutes) show "a * x ≤ x ==> star a * x ≤ x" proof - assume a: "a * x ≤ x" { fix n have "a ^ (Suc n) * x ≤ a ^ n * x" proof (induct n) case 0 thus ?case by (simp add:a power_Suc) next case (Suc n) hence "a * (a ^ Suc n * x) ≤ a * (a ^ n * x)" by (auto intro: mult_mono) thus ?case by (simp add:power_Suc mult_assoc) qed } note a = this { fix n have "a ^ n * x ≤ x" proof (induct n) case 0 show ?case by simp next case (Suc n) with a[of n] show ?case by simp qed } note b = this show "star a * x ≤ x" unfolding star_cont[of 1 a x, simplified] by (rule SUP_leI) (rule b) qed show "x * a ≤ x ==> x * star a ≤ x" (* symmetric *) proof - assume a: "x * a ≤ x" { fix n have "x * a ^ (Suc n) ≤ x * a ^ n" proof (induct n) case 0 thus ?case by (simp add:a power_Suc) next case (Suc n) hence "(x * a ^ Suc n) * a ≤ (x * a ^ n) * a" by (auto intro: mult_mono) thus ?case by (simp add:power_Suc power_commutes mult_assoc) qed } note a = this { fix n have "x * a ^ n ≤ x" proof (induct n) case 0 show ?case by simp next case (Suc n) with a[of n] show ?case by simp qed } note b = this show "x * star a ≤ x" unfolding star_cont[of x a 1, simplified] by (rule SUP_leI) (rule b) qed qed lemma less_add[simp]: fixes a b :: "'a :: order_by_add" shows "a ≤ a + b" and "b ≤ a + b" unfolding order_def by (auto simp:add_ac) lemma add_est1: fixes a b c :: "'a :: order_by_add" assumes a: "a + b ≤ c" shows "a ≤ c" using less_add(1) a by (rule order_trans) lemma add_est2: fixes a b c :: "'a :: order_by_add" assumes a: "a + b ≤ c" shows "b ≤ c" using less_add(2) a by (rule order_trans) lemma star3': fixes a b x :: "'a :: kleene" assumes a: "b + a * x ≤ x" shows "star a * b ≤ x" proof (rule order_trans) from a have "b ≤ x" by (rule add_est1) show "star a * b ≤ star a * x" by (rule mult_mono) (auto simp:`b ≤ x`) from a have "a * x ≤ x" by (rule add_est2) with star3 show "star a * x ≤ x" . qed lemma star4': fixes a b x :: "'a :: kleene" assumes a: "b + x * a ≤ x" shows "b * star a ≤ x" proof (rule order_trans) from a have "b ≤ x" by (rule add_est1) show "b * star a ≤ x * star a" by (rule mult_mono) (auto simp:`b ≤ x`) from a have "x * a ≤ x" by (rule add_est2) with star4 show "x * star a ≤ x" . qed lemma star_idemp: fixes x :: "'a :: kleene" shows "star (star x) = star x" oops lemma star_unfold_left: fixes a :: "'a :: kleene" shows "1 + a * star a = star a" proof (rule order_antisym, rule star1) have "1 + a * (1 + a * star a) ≤ 1 + a * star a" apply (rule add_mono, rule) apply (rule mult_mono, auto) apply (rule star1) done with star3' have "star a * 1 ≤ 1 + a * star a" . thus "star a ≤ 1 + a * star a" by simp qed lemma star_unfold_right: fixes a :: "'a :: kleene" shows "1 + star a * a = star a" proof (rule order_antisym, rule star2) have "1 + (1 + star a * a) * a ≤ 1 + star a * a" apply (rule add_mono, rule) apply (rule mult_mono, auto) apply (rule star2) done with star4' have "1 * star a ≤ 1 + star a * a" . thus "star a ≤ 1 + star a * a" by simp qed lemma star_zero[simp]: shows "star (0::'a::kleene) = 1" by (rule star_unfold_left[of 0, simplified]) lemma star_commute: fixes a b x :: "'a :: kleene" assumes a: "a * x = x * b" shows "star a * x = x * star b" proof (rule order_antisym) show "star a * x ≤ x * star b" proof (rule star3', rule order_trans) from a have "a * x ≤ x * b" by simp hence "a * x * star b ≤ x * b * star b" by (rule mult_mono) auto thus "x + a * (x * star b) ≤ x + x * b * star b" using add_mono by (auto simp: mult_assoc) show "… ≤ x * star b" proof - have "x * (1 + b * star b) ≤ x * star b" by (rule mult_mono[OF _ star1]) auto thus ?thesis by (simp add:right_distrib mult_assoc) qed qed show "x * star b ≤ star a * x" proof (rule star4', rule order_trans) from a have b: "x * b ≤ a * x" by simp have "star a * x * b ≤ star a * a * x" unfolding mult_assoc by (rule mult_mono[OF _ b]) auto thus "x + star a * x * b ≤ x + star a * a * x" using add_mono by auto show "… ≤ star a * x" proof - have "(1 + star a * a) * x ≤ star a * x" by (rule mult_mono[OF star2]) auto thus ?thesis by (simp add:left_distrib mult_assoc) qed qed qed lemma star_assoc: fixes c d :: "'a :: kleene" shows "star (c * d) * c = c * star (d * c)" by (auto simp:mult_assoc star_commute) lemma star_dist: fixes a b :: "'a :: kleene" shows "star (a + b) = star a * star (b * star a)" oops lemma star_one: fixes a p p' :: "'a :: kleene" assumes "p * p' = 1" and "p' * p = 1" shows "p' * star a * p = star (p' * a * p)" proof - from assms have "p' * star a * p = p' * star (p * p' * a) * p" by simp also have "… = p' * p * star (p' * a * p)" by (simp add: mult_assoc star_assoc) also have "… = star (p' * a * p)" by (simp add: assms) finally show ?thesis . qed lemma star_mono: fixes x y :: "'a :: kleene" assumes "x ≤ y" shows "star x ≤ star y" oops (* Own lemmas *) lemma x_less_star[simp]: fixes x :: "'a :: kleene" shows "x ≤ x * star a" proof - have "x ≤ x * (1 + a * star a)" by (simp add:right_distrib) also have "… = x * star a" by (simp only: star_unfold_left) finally show ?thesis . qed subsection {* Transitive Closure *} definition "tcl (x::'a::kleene) = star x * x" lemma tcl_zero: "tcl (0::'a::kleene) = 0" unfolding tcl_def by simp lemma tcl_unfold_right: "tcl a = a + tcl a * a" proof - from star_unfold_right[of a] have "a * (1 + star a * a) = a * star a" by simp from this[simplified right_distrib, simplified] show ?thesis by (simp add:tcl_def star_commute mult_ac) qed lemma less_tcl: "a ≤ tcl a" proof - have "a ≤ a + tcl a * a" by simp also have "… = tcl a" by (rule tcl_unfold_right[symmetric]) finally show ?thesis . qed subsection {* Naive Algorithm to generate the transitive closure *} function (default "λx. 0", tailrec, domintros) mk_tcl :: "('a::{plus,times,ord,zero}) => 'a => 'a" where "mk_tcl A X = (if X * A ≤ X then X else mk_tcl A (X + X * A))" by pat_completeness simp declare mk_tcl.simps[simp del] (* loops *) lemma mk_tcl_code[code]: "mk_tcl A X = (let XA = X * A in if XA ≤ X then X else mk_tcl A (X + XA))" unfolding mk_tcl.simps[of A X] Let_def .. lemma mk_tcl_lemma1: fixes X :: "'a :: kleene" shows "(X + X * A) * star A = X * star A" proof - have "A * star A ≤ 1 + A * star A" by simp also have "… = star A" by (simp add:star_unfold_left) finally have "star A + A * star A = star A" by simp hence "X * (star A + A * star A) = X * star A" by simp thus ?thesis by (simp add:left_distrib right_distrib mult_ac) qed lemma mk_tcl_lemma2: fixes X :: "'a :: kleene" shows "X * A ≤ X ==> X * star A = X" by (rule order_antisym) (auto simp:star4) lemma mk_tcl_correctness: fixes A X :: "'a :: {kleene}" assumes "mk_tcl_dom (A, X)" shows "mk_tcl A X = X * star A" using assms by induct (auto simp:mk_tcl_lemma1 mk_tcl_lemma2) lemma graph_implies_dom: "mk_tcl_graph x y ==> mk_tcl_dom x" by (rule mk_tcl_graph.induct) (auto intro:accp.accI elim:mk_tcl_rel.cases) lemma mk_tcl_default: "¬ mk_tcl_dom (a,x) ==> mk_tcl a x = 0" unfolding mk_tcl_def by (rule fundef_default_value[OF mk_tcl_sum_def graph_implies_dom]) text {* We can replace the dom-Condition of the correctness theorem with something executable *} lemma mk_tcl_correctness2: fixes A X :: "'a :: {kleene}" assumes "mk_tcl A A ≠ 0" shows "mk_tcl A A = tcl A" using assms mk_tcl_default mk_tcl_correctness unfolding tcl_def by (auto simp:star_commute) end
lemma add_idem2:
x + (x + y) = x + y
lemma ord_simp1:
x ≤ y ==> x + y = y
lemma ord_simp2:
x ≤ y ==> y + x = y
lemma ord_intro:
x + y = y ==> x ≤ y
lemma plus_leI:
[| x ≤ z; y ≤ z |] ==> x + y ≤ z
lemma le_SUPI':
l ≤ M i ==> l ≤ (SUP i. M i)
lemma zero_minimum:
(0::'a) ≤ x
lemma less_add(1):
a ≤ a + b
and less_add(2):
b ≤ a + b
lemma add_est1:
a + b ≤ c ==> a ≤ c
lemma add_est2:
a + b ≤ c ==> b ≤ c
lemma star3':
b + a * x ≤ x ==> star a * b ≤ x
lemma star4':
b + x * a ≤ x ==> b * star a ≤ x
lemma star_unfold_left:
(1::'a) + a * star a = star a
lemma star_unfold_right:
(1::'a) + star a * a = star a
lemma star_zero:
star (0::'a) = (1::'a)
lemma star_commute:
a * x = x * b ==> star a * x = x * star b
lemma star_assoc:
star (c * d) * c = c * star (d * c)
lemma star_one:
[| p * p' = (1::'a); p' * p = (1::'a) |] ==> p' * star a * p = star (p' * a * p)
lemma x_less_star:
x ≤ x * star a
lemma tcl_zero:
tcl (0::'a) = (0::'a)
lemma tcl_unfold_right:
tcl a = a + tcl a * a
lemma less_tcl:
a ≤ tcl a
lemma mk_tcl_code:
mk_tcl A X = (let XA = X * A in if XA ≤ X then X else mk_tcl A (X + XA))
lemma mk_tcl_lemma1:
(X + X * A) * star A = X * star A
lemma mk_tcl_lemma2:
X * A ≤ X ==> X * star A = X
lemma mk_tcl_correctness:
mk_tcl_dom (A, X) ==> mk_tcl A X = X * star A
lemma graph_implies_dom:
mk_tcl_graph x y ==> mk_tcl_dom x
lemma mk_tcl_default:
¬ mk_tcl_dom (a, x) ==> mk_tcl a x = (0::'a)
lemma mk_tcl_correctness2:
mk_tcl A A ≠ (0::'a) ==> mk_tcl A A = tcl A