(* Walter Guttmann, 2015.06.16 *) theory AACP imports Main begin (* theory Base imports Main begin *) -- {* This theory and the subsequent theories have been developed with Isabelle2014. *} nitpick_params [ timeout = 600 ] declare [[ smt_timeout = 600 ]] class mult = times begin notation times (infixl "\" 70) and times (infixl ";" 70) end class neg = uminus begin no_notation uminus ("- _" [81] 80) notation uminus ("- _" [80] 80) end class while = fixes while :: "'a \ 'a \ 'a" (infixr "\" 59) class L = fixes L :: "'a" class n = fixes n :: "'a \ 'a" class d = fixes d :: "'a \ 'a" class diamond = fixes diamond :: "'a \ 'a \ 'a" ("| _ > _" [50,90] 95) class box = fixes box :: "'a \ 'a \ 'a" ("| _ ] _" [50,90] 95) context ord begin definition isotone :: "('a \ 'a) \ bool" where "isotone f \ (\x y . x \ y \ f(x) \ f(y))" definition lifted_less_eq :: "('a \ 'a) \ ('a \ 'a) \ bool" ("(_ \\ _)" [51, 51] 50) where "f \\ g \ (\x . f(x) \ g(x))" definition galois :: "('a \ 'a) \ ('a \ 'a) \ bool" where "galois l u \ (\x y . l(x) \ y \ x \ u(y))" definition ascending_chain :: "(nat \ 'a) \ bool" where "ascending_chain f \ (\n . f n \ f (Suc n))" definition descending_chain :: "(nat \ 'a) \ bool" where "descending_chain f \ (\n . f (Suc n) \ f n)" definition directed :: "'a set \ bool" where "directed X \ X \ {} \ (\x\X . \y\X . \z\X . x \ z \ y \ z)" definition codirected :: "'a set \ bool" where "codirected X \ X \ {} \ (\x\X . \y\X . \z\X . z \ x \ z \ y)" definition chain :: "'a set \ bool" where "chain X \ (\x\X . \y\X . x \ y \ y \ x)" end context order begin lemma lifted_reflexive: "f = g \ f \\ g" by (metis lifted_less_eq_def order_refl) lemma lifted_transitive: "f \\ g \ g \\ h \ f \\ h" by (smt lifted_less_eq_def order_trans) lemma lifted_antisymmetric: "f \\ g \ g \\ f \ f = g" by (metis (full_types) antisym ext lifted_less_eq_def) lemma galois_char: "galois l u \ (\x . x \ u(l(x))) \ (\x . l(u(x)) \ x) \ isotone l \ isotone u" apply (rule iffI) apply (metis (full_types) galois_def isotone_def order_refl order_trans) apply (metis galois_def isotone_def order_trans) done lemma galois_closure: "galois l u \ l(x) = l(u(l(x))) \ u(x) = u(l(u(x)))" by (smt antisym galois_char isotone_def) lemma ascending_chain_k: "ascending_chain f \ f m \ f (m + k)" apply (induct k) apply simp apply (metis add_Suc_right ascending_chain_def order_trans) done lemma ascending_chain_isotone: "ascending_chain f \ m \ k \ f m \ f k" by (metis ascending_chain_k le_iff_add) lemma ascending_chain_comparable: "ascending_chain f \ f k \ f m \ f m \ f k" by (metis nat_le_linear ascending_chain_isotone) lemma ascending_chain_chain: "ascending_chain f \ chain (range f)" by (smt ascending_chain_comparable chain_def image_iff) lemma chain_directed: "X \ {} \ chain X \ directed X" by (metis chain_def directed_def) lemma ascending_chain_directed: "ascending_chain f \ directed (range f)" by (metis UNIV_not_empty ascending_chain_chain chain_directed empty_is_image) lemma descending_chain_k: "descending_chain f \ f (m + k) \ f m" apply (induct k) apply simp apply (metis add_Suc_right descending_chain_def order_trans) done lemma descending_chain_antitone: "descending_chain f \ m \ k \ f k \ f m" by (metis descending_chain_k le_iff_add) lemma descending_chain_comparable: "descending_chain f \ f k \ f m \ f m \ f k" by (metis nat_le_linear descending_chain_antitone) lemma descending_chain_chain: "descending_chain f \ chain (range f)" unfolding chain_def apply simp apply (smt descending_chain_comparable) done lemma chain_codirected: "X \ {} \ chain X \ codirected X" by (metis chain_def codirected_def) lemma descending_chain_codirected: "descending_chain f \ codirected (range f)" by (metis UNIV_not_empty descending_chain_chain chain_codirected empty_is_image) end context complete_lattice begin lemma sup_Sup: assumes nonempty: "A \ {}" shows "sup x (Sup A) = Sup ((sup x) ` A)" apply (rule antisym) apply (metis Sup_mono Sup_upper2 assms ex_in_conv imageI le_supI sup_ge1 sup_ge2) apply (smt Sup_least Sup_upper image_iff le_iff_sup sup.commute sup_ge1 sup_left_commute) done lemma sup_SUP: "Y \ {} \ sup x (SUP y:Y . f y) = (SUP y:Y. sup x (f y))" unfolding SUP_def apply rule apply (subst sup_Sup) apply (smt empty_is_image) apply (metis image_image) done lemma inf_Inf: assumes nonempty: "A \ {}" shows "inf x (Inf A) = Inf ((inf x) ` A)" apply (rule antisym) apply (smt Inf_greatest Inf_lower image_iff le_iff_inf inf.commute inf_le1 inf_left_commute) apply (metis Inf_mono Inf_lower2 assms ex_in_conv imageI le_infI inf_le1 inf_le2) done lemma inf_INF: "Y \ {} \ inf x (INF y:Y . f y) = (INF y:Y. inf x (f y))" unfolding INF_def apply rule apply (subst inf_Inf) apply (smt empty_is_image) apply (metis image_image) done lemma SUP_image_id[simp]: "(SUP x:f`A . x) = (SUP x:A . f x)" by simp lemma INF_image_id[simp]: "(INF x:f`A . x) = (INF x:A . f x)" by simp end lemma image_Collect_2: "f ` { g x | x . P x } = { f (g x) | x . P x }" by auto -- {* The following instantiation and four lemmas are from Jose Divason Mallagaray. *} instantiation "fun" :: (type, type) power begin definition one_fun :: "'a \ 'a" where one_fun_def: "one_fun \ id" definition times_fun :: "('a \ 'a) \ ('a \ 'a) \ ('a \ 'a)" where times_fun_def: "times_fun \ comp" instance by intro_classes end lemma id_power: "id^m = id" apply (induct m) apply (metis one_fun_def power_0) apply (simp add: times_fun_def) done lemma power_zero_id: "f^0 = id" by (metis one_fun_def power_0) lemma power_succ_unfold: "f^Suc m = f \ f^m" by (metis power_Suc times_fun_def) lemma power_succ_unfold_ext: "(f^Suc m) x = f ((f^m) x)" by (metis o_apply power_succ_unfold) (* end theory Fixpoint imports Base begin *) context order begin definition is_fixpoint :: "('a \ 'a) \ 'a \ bool" where "is_fixpoint f x \ f(x) = x" definition is_prefixpoint :: "('a \ 'a) \ 'a \ bool" where "is_prefixpoint f x \ f(x) \ x" definition is_postfixpoint :: "('a \ 'a) \ 'a \ bool" where "is_postfixpoint f x \ f(x) \ x" definition is_least_fixpoint :: "('a \ 'a) \ 'a \ bool" where "is_least_fixpoint f x \ f(x) = x \ (\y . f(y) = y \ x \ y)" definition is_greatest_fixpoint :: "('a \ 'a) \ 'a \ bool" where "is_greatest_fixpoint f x \ f(x) = x \ (\y . f(y) = y \ x \ y)" definition is_least_prefixpoint :: "('a \ 'a) \ 'a \ bool" where "is_least_prefixpoint f x \ f(x) \ x \ (\y . f(y) \ y \ x \ y)" definition is_greatest_postfixpoint :: "('a \ 'a) \ 'a \ bool" where "is_greatest_postfixpoint f x \ f(x) \ x \ (\y . f(y) \ y \ x \ y)" definition has_fixpoint :: "('a \ 'a) \ bool" where "has_fixpoint f \ (\x . is_fixpoint f x)" definition has_prefixpoint :: "('a \ 'a) \ bool" where "has_prefixpoint f \ (\x . is_prefixpoint f x)" definition has_postfixpoint :: "('a \ 'a) \ bool" where "has_postfixpoint f \ (\x . is_postfixpoint f x)" definition has_least_fixpoint :: "('a \ 'a) \ bool" where "has_least_fixpoint f \ (\x . is_least_fixpoint f x)" definition has_greatest_fixpoint :: "('a \ 'a) \ bool" where "has_greatest_fixpoint f \ (\x . is_greatest_fixpoint f x)" definition has_least_prefixpoint :: "('a \ 'a) \ bool" where "has_least_prefixpoint f \ (\x . is_least_prefixpoint f x)" definition has_greatest_postfixpoint :: "('a \ 'a) \ bool" where "has_greatest_postfixpoint f \ (\x . is_greatest_postfixpoint f x)" definition the_least_fixpoint :: "('a \ 'a) \ 'a" ("\ _" [201] 200) where "\ f = (THE x . is_least_fixpoint f x)" definition the_greatest_fixpoint :: "('a \ 'a) \ 'a" ("\ _" [201] 200) where "\ f = (THE x . is_greatest_fixpoint f x)" definition the_least_prefixpoint :: "('a \ 'a) \ 'a" ("p\ _" [201] 200) where "p\ f = (THE x . is_least_prefixpoint f x)" definition the_greatest_postfixpoint :: "('a \ 'a) \ 'a" ("p\ _" [201] 200) where "p\ f = (THE x . is_greatest_postfixpoint f x)" lemma least_fixpoint_unique: "has_least_fixpoint f \ (\!x . is_least_fixpoint f x)" by (smt antisym has_least_fixpoint_def is_least_fixpoint_def) lemma greatest_fixpoint_unique: "has_greatest_fixpoint f \ (\!x . is_greatest_fixpoint f x)" by (smt antisym has_greatest_fixpoint_def is_greatest_fixpoint_def) lemma least_prefixpoint_unique: "has_least_prefixpoint f \ (\!x . is_least_prefixpoint f x)" by (smt antisym has_least_prefixpoint_def is_least_prefixpoint_def) lemma greatest_postfixpoint_unique: "has_greatest_postfixpoint f \ (\!x . is_greatest_postfixpoint f x)" by (smt antisym has_greatest_postfixpoint_def is_greatest_postfixpoint_def) lemma least_fixpoint: "has_least_fixpoint f \ is_least_fixpoint f (\ f)" proof assume "has_least_fixpoint f" hence "is_least_fixpoint f (THE x . is_least_fixpoint f x)" by (smt least_fixpoint_unique theI') thus "is_least_fixpoint f (\ f)" by (simp add: is_least_fixpoint_def the_least_fixpoint_def) qed lemma greatest_fixpoint: "has_greatest_fixpoint f \ is_greatest_fixpoint f (\ f)" proof assume "has_greatest_fixpoint f" hence "is_greatest_fixpoint f (THE x . is_greatest_fixpoint f x)" by (smt greatest_fixpoint_unique theI') thus "is_greatest_fixpoint f (\ f)" by (simp add: is_greatest_fixpoint_def the_greatest_fixpoint_def) qed lemma least_prefixpoint: "has_least_prefixpoint f \ is_least_prefixpoint f (p\ f)" proof assume "has_least_prefixpoint f" hence "is_least_prefixpoint f (THE x . is_least_prefixpoint f x)" by (smt least_prefixpoint_unique theI') thus "is_least_prefixpoint f (p\ f)" by (simp add: is_least_prefixpoint_def the_least_prefixpoint_def) qed lemma greatest_postfixpoint: "has_greatest_postfixpoint f \ is_greatest_postfixpoint f (p\ f)" proof assume "has_greatest_postfixpoint f" hence "is_greatest_postfixpoint f (THE x . is_greatest_postfixpoint f x)" by (smt greatest_postfixpoint_unique theI') thus "is_greatest_postfixpoint f (p\ f)" by (simp add: is_greatest_postfixpoint_def the_greatest_postfixpoint_def) qed lemma least_fixpoint_same: "is_least_fixpoint f x \ x = \ f" by (metis least_fixpoint least_fixpoint_unique has_least_fixpoint_def) lemma greatest_fixpoint_same: "is_greatest_fixpoint f x \ x = \ f" by (metis greatest_fixpoint greatest_fixpoint_unique has_greatest_fixpoint_def) lemma least_prefixpoint_same: "is_least_prefixpoint f x \ x = p\ f" by (metis least_prefixpoint least_prefixpoint_unique has_least_prefixpoint_def) lemma greatest_postfixpoint_same: "is_greatest_postfixpoint f x \ x = p\ f" by (metis greatest_postfixpoint greatest_postfixpoint_unique has_greatest_postfixpoint_def) lemma least_fixpoint_char: "is_least_fixpoint f x \ has_least_fixpoint f \ x = \ f" by (metis least_fixpoint_same has_least_fixpoint_def) lemma least_prefixpoint_char: "is_least_prefixpoint f x \ has_least_prefixpoint f \ x = p\ f" by (metis least_prefixpoint_same has_least_prefixpoint_def) lemma greatest_fixpoint_char: "is_greatest_fixpoint f x \ has_greatest_fixpoint f \ x = \ f" by (metis greatest_fixpoint_same has_greatest_fixpoint_def) lemma greatest_postfixpoint_char: "is_greatest_postfixpoint f x \ has_greatest_postfixpoint f \ x = p\ f" by (metis greatest_postfixpoint_same has_greatest_postfixpoint_def) lemma mu_unfold: "has_least_fixpoint f \ f (\ f) = \ f" by (metis is_least_fixpoint_def least_fixpoint) lemma pmu_unfold: "has_least_prefixpoint f \ f (p\ f) \ p\ f" by (metis is_least_prefixpoint_def least_prefixpoint) lemma nu_unfold: "has_greatest_fixpoint f \ \ f = f (\ f)" by (metis is_greatest_fixpoint_def greatest_fixpoint) lemma pnu_unfold: "has_greatest_postfixpoint f \ p\ f \ f (p\ f)" by (metis is_greatest_postfixpoint_def greatest_postfixpoint) lemma least_prefixpoint_fixpoint: "has_least_prefixpoint f \ isotone f \ is_least_fixpoint f (p\ f)" by (smt eq_iff is_least_fixpoint_def is_least_prefixpoint_def isotone_def least_prefixpoint) lemma pmu_mu: "has_least_prefixpoint f \ isotone f \ p\ f = \ f" by (smt has_least_fixpoint_def is_least_fixpoint_def least_fixpoint_unique least_prefixpoint_fixpoint least_fixpoint) lemma greatest_postfixpoint_fixpoint: "has_greatest_postfixpoint f \ isotone f \ is_greatest_fixpoint f (p\ f)" by (smt eq_iff is_greatest_fixpoint_def is_greatest_postfixpoint_def isotone_def greatest_postfixpoint) lemma pnu_nu: "has_greatest_postfixpoint f \ isotone f \ p\ f = \ f" by (smt has_greatest_fixpoint_def is_greatest_fixpoint_def greatest_fixpoint_unique greatest_postfixpoint_fixpoint greatest_fixpoint) lemma pmu_isotone: "has_least_prefixpoint f \ has_least_prefixpoint g \ f \\ g \ p\ f \ p\ g" by (smt is_least_prefixpoint_def least_prefixpoint lifted_less_eq_def order_trans) lemma mu_isotone: "has_least_prefixpoint f \ has_least_prefixpoint g \ isotone f \ isotone g \ f \\ g \ \ f \ \ g" by (metis pmu_isotone pmu_mu) lemma pnu_isotone: "has_greatest_postfixpoint f \ has_greatest_postfixpoint g \ f \\ g \ p\ f \ p\ g" by (smt is_greatest_postfixpoint_def lifted_less_eq_def order_trans greatest_postfixpoint) lemma nu_isotone: "has_greatest_postfixpoint f \ has_greatest_postfixpoint g \ isotone f \ isotone g \ f \\ g \ \ f \ \ g" by (metis pnu_isotone pnu_nu) lemma mu_square: "isotone f \ has_least_fixpoint f \ has_least_fixpoint (f \ f) \ \ f = \ (f \ f)" by (metis (no_types, hide_lams) antisym is_least_fixpoint_def ord.isotone_def least_fixpoint_char least_fixpoint_unique o_apply) lemma nu_square: "isotone f \ has_greatest_fixpoint f \ has_greatest_fixpoint (f \ f) \ \ f = \ (f \ f)" by (metis (no_types, hide_lams) antisym is_greatest_fixpoint_def ord.isotone_def greatest_fixpoint_char greatest_fixpoint_unique o_apply) lemma mu_roll: "isotone g \ has_least_fixpoint (f \ g) \ has_least_fixpoint (g \ f) \ \ (g \ f) = g(\ (f \ g))" apply (rule impI) apply (rule antisym) apply (smt is_least_fixpoint_def least_fixpoint o_apply) apply (smt is_least_fixpoint_def isotone_def least_fixpoint o_apply) done lemma nu_roll: "isotone g \ has_greatest_fixpoint (f \ g) \ has_greatest_fixpoint (g \ f) \ \ (g \ f) = g(\ (f \ g))" apply (rule impI) apply (rule antisym) apply (smt is_greatest_fixpoint_def greatest_fixpoint isotone_def o_apply) apply (smt is_greatest_fixpoint_def greatest_fixpoint o_apply) done lemma mu_below_nu: "has_least_fixpoint f \ has_greatest_fixpoint f \ \ f \ \ f" by (metis is_greatest_fixpoint_def is_least_fixpoint_def least_fixpoint greatest_fixpoint) lemma pmu_below_pnu_fix: "has_fixpoint f \ has_least_prefixpoint f \ has_greatest_postfixpoint f \ p\ f \ p\ f" by (smt has_fixpoint_def is_fixpoint_def is_greatest_postfixpoint_def is_least_prefixpoint_def le_less order_trans least_prefixpoint greatest_postfixpoint) lemma pmu_below_pnu_iso: "isotone f \ has_least_prefixpoint f \ has_greatest_postfixpoint f \ p\ f \ p\ f" by (metis has_fixpoint_def is_fixpoint_def is_least_fixpoint_def least_prefixpoint_fixpoint pmu_below_pnu_fix) lemma mu_fusion_1: "galois l u \ isotone h \ has_least_prefixpoint g \ has_least_fixpoint h \ l(g(u(\ h))) \ h(l(u(\ h))) \ l(p\ g) \ \ h" proof assume 1: "galois l u \ isotone h \ has_least_prefixpoint g \ has_least_fixpoint h \ l(g(u(\ h))) \ h(l(u(\ h)))" hence "l(g(u(\ h))) \ \ h" by (metis galois_char least_fixpoint_same least_fixpoint_unique is_least_fixpoint_def isotone_def order_trans) thus "l(p\ g) \ \ h" using 1 by (metis galois_def least_prefixpoint is_least_prefixpoint_def least_fixpoint_same least_fixpoint_unique) qed lemma mu_fusion_2: "galois l u \ isotone h \ has_least_prefixpoint g \ has_least_fixpoint h \ l \ g \\ h \ l \ l(p\ g) \ \ h" by (metis lifted_less_eq_def mu_fusion_1 o_apply) lemma mu_fusion_equal_1: "galois l u \ isotone g \ isotone h \ has_least_prefixpoint g \ has_least_fixpoint h \ l(g(u(\ h))) \ h(l(u(\ h))) \ l(g(p\ g)) = h(l(p\ g)) \ \ h = l(p\ g) \ \ h = l(\ g)" by (metis antisym least_fixpoint least_prefixpoint_fixpoint is_least_fixpoint_def mu_fusion_1 pmu_mu) lemma mu_fusion_equal_2: "galois l u \ isotone h \ has_least_prefixpoint g \ has_least_prefixpoint h \ l(g(u(\ h))) \ h(l(u(\ h))) \ l(g(p\ g)) = h(l(p\ g)) \ p\ h = l(p\ g) \ \ h = l(p\ g)" by (smt antisym galois_char least_fixpoint_char least_prefixpoint least_prefixpoint_fixpoint is_least_prefixpoint_def isotone_def mu_fusion_1) lemma mu_fusion_equal_3: "galois l u \ isotone g \ isotone h \ has_least_prefixpoint g \ has_least_fixpoint h \ l \ g = h \ l \ \ h = l(p\ g) \ \ h = l(\ g)" by (metis mu_fusion_equal_1 o_apply order_refl) lemma mu_fusion_equal_4: "galois l u \ isotone h \ has_least_prefixpoint g \ has_least_prefixpoint h \ l \ g = h \ l \ p\ h = l(p\ g) \ \ h = l(p\ g)" by (metis mu_fusion_equal_2 o_apply order_refl) lemma nu_fusion_1: "galois l u \ isotone h \ has_greatest_postfixpoint g \ has_greatest_fixpoint h \ h(u(l(\ h))) \ u(g(l(\ h))) \ \ h \ u(p\ g)" proof assume 1: "galois l u \ isotone h \ has_greatest_postfixpoint g \ has_greatest_fixpoint h \ h(u(l(\ h))) \ u(g(l(\ h)))" hence "\ h \ u(g(l(\ h)))" using 1 by (metis galois_char greatest_fixpoint_same greatest_fixpoint_unique is_greatest_fixpoint_def isotone_def order_trans) thus "\ h \ u(p\ g)" using 1 by (smt galois_def greatest_postfixpoint is_greatest_postfixpoint_def greatest_fixpoint_same greatest_fixpoint_unique) qed lemma nu_fusion_2: "galois l u \ isotone h \ has_greatest_postfixpoint g \ has_greatest_fixpoint h \ h \ u \\ u \ g \ \ h \ u(p\ g)" by (metis lifted_less_eq_def nu_fusion_1 o_apply) lemma nu_fusion_equal_1: "galois l u \ isotone g \ isotone h \ has_greatest_postfixpoint g \ has_greatest_fixpoint h \ h(u(l(\ h))) \ u(g(l(\ h))) \ h(u(p\ g)) = u(g(p\ g)) \ \ h = u(p\ g) \ \ h = u(\ g)" by (metis antisym greatest_fixpoint greatest_postfixpoint_fixpoint is_greatest_fixpoint_def nu_fusion_1 pnu_nu) lemma nu_fusion_equal_2: "galois l u \ isotone h \ has_greatest_postfixpoint g \ has_greatest_postfixpoint h \ h(u(l(\ h))) \ u(g(l(\ h))) \ h(u(p\ g)) = u(g(p\ g)) \ p\ h = u(p\ g) \ \ h = u(p\ g)" by (smt antisym galois_char greatest_fixpoint_char greatest_postfixpoint greatest_postfixpoint_fixpoint is_greatest_postfixpoint_def isotone_def nu_fusion_1) lemma nu_fusion_equal_3: "galois l u \ isotone g \ isotone h \ has_greatest_postfixpoint g \ has_greatest_fixpoint h \ h \ u = u \ g \ \ h = u(p\ g) \ \ h = u(\ g)" by (metis nu_fusion_equal_1 o_apply order_refl) lemma nu_fusion_equal_4: "galois l u \ isotone h \ has_greatest_postfixpoint g \ has_greatest_postfixpoint h \ h \ u = u \ g \ p\ h = u(p\ g) \ \ h = u(p\ g)" by (metis nu_fusion_equal_2 o_apply order_refl) lemma mu_exchange_1: "galois l u \ isotone g \ isotone h \ has_least_prefixpoint (l \ h) \ has_least_prefixpoint (h \ g) \ has_least_fixpoint (g \ h) \ l \ h \ g \\ g \ h \ l \ \(l \ h) \ \(g \ h)" proof assume 1: "galois l u \ isotone g \ isotone h \ has_least_prefixpoint (l \ h) \ has_least_prefixpoint (h \ g) \ has_least_fixpoint (g \ h) \ l \ h \ g \\ g \ h \ l" hence "l \ (h \ g) \\ (g \ h) \ l" by (metis o_assoc) thus "\(l \ h) \ \(g \ h)" using 1 by (smt galois_char is_least_prefixpoint_def isotone_def least_fixpoint_char least_prefixpoint least_prefixpoint_fixpoint mu_fusion_2 mu_roll o_apply) qed lemma mu_exchange_2: "galois l u \ isotone g \ isotone h \ has_least_prefixpoint (l \ h) \ has_least_prefixpoint (h \ l) \ has_least_prefixpoint (h \ g) \ has_least_fixpoint (g \ h) \ has_least_fixpoint (h \ g) \ l \ h \ g \\ g \ h \ l \ \(h \ l) \ \(h \ g)" by (smt galois_char isotone_def least_fixpoint_char least_prefixpoint_fixpoint mu_exchange_1 mu_roll o_apply) lemma mu_exchange_equal: "galois l u \ galois k t \ isotone h \ has_least_prefixpoint (l \ h) \ has_least_prefixpoint (h \ l) \ has_least_prefixpoint (k \ h) \ has_least_prefixpoint (h \ k) \ l \ h \ k = k \ h \ l \ \(l \ h) = \(k \ h) \ \(h \ l) = \(h \ k)" by (smt antisym galois_char isotone_def least_fixpoint_char least_prefixpoint_fixpoint lifted_reflexive mu_exchange_1 mu_exchange_2 o_apply) lemma nu_exchange_1: "galois l u \ isotone g \ isotone h \ has_greatest_postfixpoint (u \ h) \ has_greatest_postfixpoint (h \ g) \ has_greatest_fixpoint (g \ h) \ g \ h \ u \\ u \ h \ g \ \(g \ h) \ \(u \ h)" proof assume 1: "galois l u \ isotone g \ isotone h \ has_greatest_postfixpoint (u \ h) \ has_greatest_postfixpoint (h \ g) \ has_greatest_fixpoint (g \ h) \ g \ h \ u \\ u \ h \ g" hence "(g \ h) \ u \\ u \ (h \ g)" by (metis o_assoc) thus "\(g \ h) \ \(u \ h)" using 1 by (smt galois_char is_greatest_postfixpoint_def isotone_def greatest_fixpoint_char greatest_postfixpoint greatest_postfixpoint_fixpoint nu_fusion_2 nu_roll o_apply) qed lemma nu_exchange_2: "galois l u \ isotone g \ isotone h \ has_greatest_postfixpoint (u \ h) \ has_greatest_postfixpoint (h \ u) \ has_greatest_postfixpoint (h \ g) \ has_greatest_fixpoint (g \ h) \ has_greatest_fixpoint (h \ g) \ g \ h \ u \\ u \ h \ g \ \(h \ g) \ \(h \ u)" by (smt galois_char isotone_def greatest_fixpoint_char greatest_postfixpoint_fixpoint nu_exchange_1 nu_roll o_apply) lemma nu_exchange_equal: "galois l u \ galois k t \ isotone h \ has_greatest_postfixpoint (u \ h) \ has_greatest_postfixpoint (h \ u) \ has_greatest_postfixpoint (t \ h) \ has_greatest_postfixpoint (h \ t) \ u \ h \ t = t \ h \ u \ \(u \ h) = \(t \ h) \ \(h \ u) = \(h \ t)" by (smt antisym galois_char isotone_def greatest_fixpoint_char greatest_postfixpoint_fixpoint lifted_reflexive nu_exchange_1 nu_exchange_2 o_apply) lemma mu_commute_fixpoint_1: "isotone f \ has_least_fixpoint (f \ g) \ f \ g = g \ f \ is_fixpoint f (\(f \ g))" by (metis is_fixpoint_def mu_roll) lemma mu_commute_fixpoint_2: "isotone g \ has_least_fixpoint (f \ g) \ f \ g = g \ f \ is_fixpoint g (\(f \ g))" by (metis is_fixpoint_def mu_roll) lemma mu_commute_least_fixpoint: "isotone f \ isotone g \ has_least_fixpoint f \ has_least_fixpoint g \ has_least_fixpoint (f \ g) \ f \ g = g \ f \ (\(f \ g) = \ f \ \ g \ \ f)" by (metis is_least_fixpoint_def least_fixpoint_same least_fixpoint_unique mu_roll) (* converse claimed by [Davey & Priestley 2002, page 200, exercise 8.27(iii)] for continuous f, g on a CPO ; does it hold without additional assumptions? lemma mu_commute_least_fixpoint_converse: "isotone f \ isotone g \ has_least_prefixpoint f \ has_least_prefixpoint g \ has_least_prefixpoint (f \ g) \ f \ g = g \ f \ (\ g \ \ f \ \(f \ g) = \ f)" *) lemma nu_commute_fixpoint_1: "isotone f \ has_greatest_fixpoint (f \ g) \ f \ g = g \ f \ is_fixpoint f (\(f \ g))" by (metis is_fixpoint_def nu_roll) lemma nu_commute_fixpoint_2: "isotone g \ has_greatest_fixpoint (f \ g) \ f \ g = g \ f \ is_fixpoint g (\(f \ g))" by (metis is_fixpoint_def nu_roll) lemma nu_commute_greatest_fixpoint: "isotone f \ isotone g \ has_greatest_fixpoint f \ has_greatest_fixpoint g \ has_greatest_fixpoint (f \ g) \ f \ g = g \ f \ (\(f \ g) = \ f \ \ f \ \ g)" by (smt is_greatest_fixpoint_def greatest_fixpoint_same greatest_fixpoint_unique nu_roll) lemma mu_diagonal_1: "isotone (\x . f x x) \ (\x . isotone (\y . f x y)) \ isotone (\x . \(\y . f x y)) \ (\x . has_least_fixpoint (\y . f x y)) \ has_least_prefixpoint (\x . \(\y . f x y)) \ \(\x . f x x) = \(\x . \(\y . f x y))" by (smt is_least_fixpoint_def is_least_prefixpoint_def least_fixpoint_same least_fixpoint_unique least_prefixpoint least_prefixpoint_fixpoint) lemma mu_diagonal_2: "(\x . isotone (\y . f x y) \ isotone (\y . f y x) \ has_least_prefixpoint (\y . f x y)) \ has_least_prefixpoint (\x . \(\y . f x y)) \ \(\x . f x x) = \(\x . \(\y . f x y))" proof assume 1: "(\x . isotone (\y . f x y) \ isotone (\y . f y x) \ has_least_prefixpoint (\y . f x y)) \ has_least_prefixpoint (\x . \(\y . f x y))" hence "isotone (\x . \(\y . f x y))" by (smt isotone_def lifted_less_eq_def mu_isotone) thus "\(\x . f x x) = \(\x . \(\y . f x y))" using 1 by (smt is_least_fixpoint_def is_least_prefixpoint_def least_fixpoint_same least_prefixpoint least_prefixpoint_fixpoint) qed lemma nu_diagonal_1: "isotone (\x . f x x) \ (\x . isotone (\y . f x y)) \ isotone (\x . \(\y . f x y)) \ (\x . has_greatest_fixpoint (\y . f x y)) \ has_greatest_postfixpoint (\x . \(\y . f x y)) \ \(\x . f x x) = \(\x . \(\y . f x y))" by (smt is_greatest_fixpoint_def is_greatest_postfixpoint_def greatest_fixpoint_same greatest_fixpoint_unique greatest_postfixpoint greatest_postfixpoint_fixpoint) lemma nu_diagonal_2: "(\x . isotone (\y . f x y) \ isotone (\y . f y x) \ has_greatest_postfixpoint (\y . f x y)) \ has_greatest_postfixpoint (\x . \(\y . f x y)) \ \(\x . f x x) = \(\x . \(\y . f x y))" proof assume 1: "(\x . isotone (\y . f x y) \ isotone (\y . f y x) \ has_greatest_postfixpoint (\y . f x y)) \ has_greatest_postfixpoint (\x . \(\y . f x y))" hence "isotone (\x . \(\y . f x y))" by (smt isotone_def lifted_less_eq_def nu_isotone) thus "\(\x . f x x) = \(\x . \(\y . f x y))" using 1 by (smt greatest_fixpoint_same greatest_postfixpoint greatest_postfixpoint_fixpoint is_greatest_fixpoint_def is_greatest_postfixpoint_def) qed end (* end theory Lattice imports Base begin *) class join_semilattice = plus + ord + assumes add_associative: "(x + y) + z = x + (y + z)" assumes add_commutative: "x + y = y + x" assumes add_idempotent: "x + x = x" assumes less_eq_def: "x \ y \ x + y = y" assumes less_def: "x < y \ x \ y \ \ (y \ x)" begin subclass order apply unfold_locales apply (metis less_def) apply (metis add_idempotent less_eq_def) apply (metis add_associative less_eq_def) apply (metis add_commutative less_eq_def) done lemma add_left_isotone: "x \ y \ x + z \ y + z" by (smt add_associative add_commutative add_idempotent less_eq_def) lemma add_right_isotone: "x \ y \ z + x \ z + y" by (metis add_commutative add_left_isotone) lemma add_isotone: "w \ y \ x \ z \ w + x \ y + z" by (smt add_associative add_commutative less_eq_def) lemma add_left_upper_bound: "x \ x + y" by (metis add_associative add_idempotent less_eq_def) lemma add_right_upper_bound: "y \ x + y" by (metis add_commutative add_left_upper_bound) lemma add_least_upper_bound: "x \ z \ y \ z \ x + y \ z" by (smt add_associative add_commutative add_left_upper_bound less_eq_def) lemma add_left_divisibility: "x \ y \ (\z . x + z = y)" by (metis add_left_upper_bound less_eq_def) lemma add_right_divisibility: "x \ y \ (\z . z + x = y)" by (metis add_commutative add_left_divisibility) lemma add_same_context:"x \ y + z \ y \ x + z \ x + z = y + z" by (smt add_associative add_commutative less_eq_def) lemma add_relative_same_increasing: "x \ y \ x + z = x + w \ y + z = y + w" by (smt add_associative add_right_divisibility) lemma ascending_chain_left_add: "ascending_chain f \ ascending_chain (\n . x + f n)" by (metis ascending_chain_def add_right_isotone) lemma ascending_chain_right_add: "ascending_chain f \ ascending_chain (\n . f n + x)" by (metis ascending_chain_def add_left_isotone) lemma descending_chain_left_add: "descending_chain f \ descending_chain (\n . x + f n)" by (metis descending_chain_def add_right_isotone) lemma descending_chain_right_add: "descending_chain f \ descending_chain (\n . f n + x)" by (metis descending_chain_def add_left_isotone) primrec pSum0 :: "(nat \ 'a) \ nat \ 'a" where "pSum0 f 0 = f 0" | "pSum0 f (Suc m) = pSum0 f m + f m" lemma pSum0_below: "(\i . f i \ x) \ pSum0 f m \ x" apply (induct m) apply (metis pSum0.simps(1)) by (metis add_least_upper_bound pSum0.simps(2)) end class bounded_join_semilattice = join_semilattice + zero + assumes add_left_zero: "0 + x = x" begin lemma add_right_zero: "x + 0 = x" by (metis add_commutative add_left_zero) lemma zero_least: "0 \ x" by (metis add_left_upper_bound add_left_zero) end class meet = fixes meet :: "'a \ 'a \ 'a" (infixl "\" 65) class meet_semilattice = meet + ord + assumes meet_associative: "(x \ y) \ z = x \ (y \ z)" assumes meet_commutative: "x \ y = y \ x" assumes meet_idempotent: "x \ x = x" assumes meet_less_eq_def: "x \ y \ x \ y = x" assumes meet_less_def: "x < y \ x \ y \ \ (y \ x)" sublocale meet_semilattice < meet!: join_semilattice where plus = meet and less_eq = "(\x y . y \ x)" and less = "(\x y . y < x)" apply unfold_locales apply (rule meet_associative) apply (rule meet_commutative) apply (rule meet_idempotent) apply (metis meet_commutative meet_less_eq_def) apply (metis meet_less_def) done class T = fixes T :: "'a" ("\") class bounded_meet_semilattice = meet_semilattice + T + assumes meet_left_top: "T \ x = x" sublocale bounded_meet_semilattice < meet!: bounded_join_semilattice where plus = meet and less_eq = "(\x y . y \ x)" and less = "(\x y . y < x)" and zero = T apply unfold_locales apply (rule meet_left_top) done class bounded_distributive_lattice = bounded_join_semilattice + bounded_meet_semilattice + assumes meet_left_dist_add: "x \ (y + z) = (x \ y) + (x \ z)" assumes add_left_dist_meet: "x + (y \ z) = (x + y) \ (x + z)" assumes meet_absorb: "x \ (x + y) = x" assumes add_absorb: "x + (x \ y) = x" begin lemma meet_left_zero: "0 \ x = 0" by (metis add_absorb add_left_zero) lemma meet_right_zero: "x \ 0 = 0" by (metis meet_commutative meet_left_zero) lemma add_left_top_1: "T + x = T" by (metis add_absorb meet_left_top) lemma add_right_top_1: "x + T = T" by (metis add_commutative add_left_top_1) lemma meet_same_context:"x \ y \ z \ y \ x \ z \ x \ z = y \ z" by (metis eq_iff meet.add_least_upper_bound) lemma relative_equality: "x + z = y + z \ x \ z = y \ z \ x = y" by (metis add_absorb add_commutative add_left_dist_meet) end (* end theory Semiring imports Fixpoint Lattice begin *) class monoid = mult + one + assumes mult_associative: "(x ; y) ; z = x ; (y ; z)" assumes mult_left_one_1: "1 ; x = x" assumes mult_right_one: "x ; 1 = x" class non_associative_left_semiring = bounded_join_semilattice + mult + one + assumes mult_left_sub_dist_add: "x ; y + x ; z \ x ; (y + z)" assumes mult_right_dist_add: "(x + y) ; z = x ; z + y ; z" assumes mult_left_zero: "0 ; x = 0" assumes mult_left_one: "1 ; x = x" assumes mult_sub_right_one: "x \ x ; 1" begin lemma mult_left_isotone: "x \ y \ x ; z \ y ; z" by (metis less_eq_def mult_right_dist_add) lemma mult_right_isotone: "x \ y \ z ; x \ z ; y" by (metis add_least_upper_bound less_eq_def mult_left_sub_dist_add) lemma mult_isotone: "w \ y \ x \ z \ w ; x \ y ; z" by (smt mult_left_isotone mult_right_isotone order_trans) lemma affine_isotone: "isotone (\x . y ; x + z)" by (smt add_commutative add_right_isotone isotone_def mult_right_isotone) lemma mult_left_sub_dist_add_left: "x ; y \ x ; (y + z)" by (metis add_left_upper_bound mult_right_isotone) lemma mult_left_sub_dist_add_right: "x ; z \ x ; (y + z)" by (metis add_right_upper_bound mult_right_isotone) lemma mult_right_sub_dist_add_left: "x ; z \ (x + y) ; z" by (metis add_left_upper_bound mult_right_dist_add) lemma mult_right_sub_dist_add_right: "y ; z \ (x + y) ; z" by (metis add_right_upper_bound mult_right_dist_add) lemma case_split_left: "1 \ w + z \ w ; x \ y \ z ; x \ y \ x \ y" by (smt add_associative add_commutative less_eq_def mult_left_one mult_right_dist_add order_refl) lemma case_split_left_equal: "w + z = 1 \ w ; x = w ; y \ z ; x = z ; y \ x = y" by (metis mult_left_one mult_right_dist_add) lemma ascending_chain_left_mult: "ascending_chain f \ ascending_chain (\n . x ; f n)" by (metis ascending_chain_def mult_right_isotone) lemma ascending_chain_right_mult: "ascending_chain f \ ascending_chain (\n . f n ; x)" by (metis ascending_chain_def mult_left_isotone) lemma descending_chain_left_mult: "descending_chain f \ descending_chain (\n . x ; f n)" by (metis descending_chain_def mult_right_isotone) lemma descending_chain_right_mult: "descending_chain f \ descending_chain (\n . f n ; x)" by (metis descending_chain_def mult_left_isotone) -- {* Some results about transitive closures in this class and the next two classes are taken from a joint paper with Rudolf Berghammer. *} abbreviation Lf :: "'a \ ('a \ 'a)" where "Lf y \ (\ x . 1 + x ; y)" abbreviation Rf :: "'a \ ('a \ 'a)" where "Rf y \ (\ x . 1 + y ; x)" abbreviation Sf :: "'a \ ('a \ 'a)" where "Sf y \ (\ x . 1 + y + x ; x)" abbreviation lstar :: "'a \ 'a" where "lstar y \ p\ (Lf y)" abbreviation rstar :: "'a \ 'a" where "rstar y \ p\ (Rf y)" abbreviation sstar :: "'a \ 'a" where "sstar y \ p\ (Sf y)" lemma lstar_rec_isotone: "isotone (Lf y)" by (smt2 add_isotone add_right_divisibility isotone_def mult_right_sub_dist_add_right order.refl) lemma rstar_rec_isotone: "isotone (Rf y)" by (metis add_isotone isotone_def less_eq_def mult_left_sub_dist_add_left order_refl) lemma sstar_rec_isotone: "isotone (Sf y)" by (simp add: add_right_isotone isotone_def mult_isotone) lemma lstar_fixpoint: "has_least_prefixpoint (Lf y) \ lstar y = \ (Lf y)" by (metis pmu_mu lstar_rec_isotone) lemma rstar_fixpoint: "has_least_prefixpoint (Rf y) \ rstar y = \ (Rf y)" by (metis pmu_mu rstar_rec_isotone) lemma sstar_fixpoint: "has_least_prefixpoint (Sf y) \ sstar y = \ (Sf y)" by (metis pmu_mu sstar_rec_isotone) lemma sstar_increasing: "has_least_prefixpoint (Sf y) \ y \ sstar y" by (metis add_least_upper_bound add_left_upper_bound order.trans pmu_unfold) lemma rstar_below_sstar: "has_least_prefixpoint (Rf y) \ has_least_prefixpoint (Sf y) \ rstar y \ sstar y" proof assume 1: "has_least_prefixpoint (Rf y) \ has_least_prefixpoint (Sf y)" hence "Rf y (sstar y) \ Sf y (sstar y)" by (smt2 add_isotone add_left_upper_bound add_right_upper_bound dual_order.trans mult_left_isotone pmu_unfold) also have "... \ sstar y" using 1 by (metis (erased, lifting) pmu_unfold) finally show "rstar y \ sstar y" using 1 by (metis (erased, lifting) is_least_prefixpoint_def least_prefixpoint) qed end class pre_left_semiring = non_associative_left_semiring + assumes mult_semi_associative: "(x ; y) ; z \ x ; (y ; z)" begin lemma mult_one_associative: "x ; 1 ; y = x ; y" by (metis dual_order.antisym mult_left_isotone mult_left_one mult_semi_associative mult_sub_right_one) lemma mult_sup_associative_one: "(x ; (y ; 1)) ; z \ x ; (y ; z)" by (metis mult_semi_associative mult_one_associative) lemma rstar_increasing: "has_least_prefixpoint (Rf y) \ y \ rstar y" proof assume "has_least_prefixpoint (Rf y)" hence "Rf y (rstar y) \ rstar y" by (metis pmu_unfold) thus "y \ rstar y" by (metis add_least_upper_bound mult_right_isotone mult_sub_right_one order.trans) qed end class residuated_pre_left_semiring = pre_left_semiring + inverse + assumes lres_galois: "x ; y \ z \ x \ z / y" begin -- {* Theorem 32.2 *} lemma lres_left_isotone: "x \ y \ x / z \ y / z" by (metis lres_galois order.refl order.trans) -- {* Theorem 32.2 *} lemma lres_right_antitone: "x \ y \ z / y \ z / x" by (metis lres_galois mult_right_isotone order.refl order_trans) lemma lres_inverse: "(x / y) ; y \ x" by (metis lres_galois order_refl) lemma lres_one: "x / 1 \ x" by (metis dual_order.trans mult_sub_right_one lres_inverse) lemma lres_mult_sub_lres_lres: "x / (z ; y) \ (x / y) / z" by (metis lres_galois lres_inverse mult_semi_associative order.trans) -- {* Theorem 32.4 *} lemma mult_lres_sub_assoc: "x ; (y / z) \ (x ; y) / z" by (metis (full_types) lres_galois lres_inverse mult_right_isotone mult_semi_associative order_trans) lemma lstar_below_rstar: "has_least_prefixpoint (Lf y) \ has_least_prefixpoint (Rf y) \ lstar y \ rstar y" proof assume 1: "has_least_prefixpoint (Lf y) \ has_least_prefixpoint (Rf y)" have "y ; (rstar y / y) ; y \ y ; rstar y" by (metis mult_right_isotone mult_semi_associative order_trans lres_inverse) also have "... \ rstar y" using 1 by (metis add_least_upper_bound pmu_unfold) finally have "y ; (rstar y / y) \ rstar y / y" by (metis lres_galois) hence "Rf y (rstar y / y) \ rstar y / y" using 1 by (metis add_least_upper_bound lres_galois mult_left_one rstar_increasing) hence "rstar y \ rstar y / y" using 1 by (metis is_least_prefixpoint_def least_prefixpoint) hence "Lf y (rstar y) \ rstar y" using 1 by (metis add_least_upper_bound lres_galois pmu_unfold) thus "lstar y \ rstar y" using 1 by (metis (erased, lifting) is_least_prefixpoint_def least_prefixpoint) qed -- {* The next proof follows an argument of Rudolf Berghammer; see Satz 10.1.5 in his 2012 book. *} lemma rstar_sstar: "has_least_prefixpoint (Rf y) \ has_least_prefixpoint (Sf y) \ rstar y = sstar y" proof assume 1: "has_least_prefixpoint (Rf y) \ has_least_prefixpoint (Sf y)" have "Rf y (rstar y / rstar y) ; rstar y \ rstar y + y ; ((rstar y / rstar y) ; rstar y)" by (metis add_isotone mult_left_one mult_right_dist_add mult_semi_associative) also have "... \ rstar y + y ; rstar y" by (metis add_right_isotone mult_right_isotone lres_inverse) also have "... \ rstar y" using 1 by (metis (full_types) add_least_upper_bound order_refl pmu_unfold) finally have "Rf y (rstar y / rstar y) \ rstar y / rstar y" by (metis lres_galois) hence "rstar y ; rstar y \ rstar y" using 1 by (metis (erased, lifting) is_least_prefixpoint_def least_prefixpoint lres_galois) hence "y + rstar y ; rstar y \ rstar y" using 1 by (metis add_least_upper_bound rstar_increasing) hence "Sf y (rstar y) \ rstar y" using 1 by (metis (full_types) add_least_upper_bound pmu_unfold) hence "sstar y \ rstar y" using 1 by (metis (erased, lifting) is_least_prefixpoint_def least_prefixpoint) thus "rstar y = sstar y" using 1 by (metis antisym rstar_below_sstar) qed end class idempotent_left_semiring = non_associative_left_semiring + monoid begin subclass pre_left_semiring apply unfold_locales apply (metis mult_associative order_refl) done lemma zero_right_mult_decreasing: "x ; 0 \ x" by (metis add_right_zero mult_left_sub_dist_add_right mult_right_one) lemma test_preserves_equation: "p \ p ; p \ p \ 1 \ (p ; x \ x ; p \ p ; x = p ; x ; p)" apply rule apply rule apply (rule antisym) apply (smt antisym mult_associative mult_right_isotone mult_right_one) apply (metis mult_right_isotone mult_right_one) apply (metis mult_left_isotone mult_left_one) done end class idempotent_left_zero_semiring = idempotent_left_semiring + assumes mult_left_dist_add: "x ; (y + z) = x ; y + x ; z" begin lemma case_split_right: "1 \ w + z \ x ; w \ y \ x ; z \ y \ x \ y" by (smt add_associative add_commutative less_eq_def mult_left_dist_add mult_right_one) lemma case_split_right_equal: "w + z = 1 \ x ; w = y ; w \ x ; z = y ; z \ x = y" by (metis mult_left_dist_add mult_right_one) end class idempotent_semiring = idempotent_left_zero_semiring + assumes mult_right_zero: "x ; 0 = 0" class bounded_pre_left_semiring = pre_left_semiring + T + assumes add_right_top: "x + T = T" begin lemma add_left_top: "T + x = T" by (metis add_commutative add_right_top) lemma top_greatest: "x \ T" by (metis add_left_top add_right_upper_bound) lemma top_left_mult_increasing: "x \ T ; x" by (metis mult_left_isotone mult_left_one top_greatest) lemma top_right_mult_increasing: "x \ x ; T" by (metis order_trans mult_right_isotone mult_sub_right_one top_greatest) lemma top_mult_top: "T ; T = T" by (metis add_right_divisibility add_right_top top_right_mult_increasing) definition vector :: "'a \ bool" where "vector x \ x = x ; T" lemma vector_zero: "vector 0" by (metis mult_left_zero vector_def) lemma vector_top: "vector T" by (metis top_mult_top vector_def) lemma vector_add_closed: "vector x \ vector y \ vector (x + y)" by (metis mult_right_dist_add vector_def) lemma vector_left_mult_closed: "vector y \ vector (x ; y)" by (metis antisym mult_semi_associative top_right_mult_increasing vector_def) end class bounded_residuated_pre_left_semiring = residuated_pre_left_semiring + bounded_pre_left_semiring begin -- {* Theorem 32.8 *} lemma lres_top_decreasing: "x / T \ x" by (metis lres_one lres_right_antitone order_trans top_greatest) -- {* Theorem 32.9 *} lemma top_lres_absorb: "T / x = T" by (metis eq_iff lres_galois top_greatest) end class bounded_idempotent_left_semiring = bounded_pre_left_semiring + idempotent_left_semiring class bounded_idempotent_left_zero_semiring = bounded_idempotent_left_semiring + idempotent_left_zero_semiring class bounded_idempotent_semiring = bounded_idempotent_left_zero_semiring + idempotent_semiring (* end theory Itering imports Semiring begin *) class circ = fixes circ :: "'a \ 'a" ("_\<^sup>\" [100] 100) class left_conway_semiring = idempotent_left_semiring + circ + assumes circ_left_unfold: "1 + x ; x\<^sup>\ = x\<^sup>\" assumes circ_left_slide: "(x ; y)\<^sup>\ ; x \ x ; (y ; x)\<^sup>\" assumes circ_add_1: "(x + y)\<^sup>\ = x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" begin -- {* Theorem 14.19 *} lemma circ_mult_sub: "1 + x ; (y ; x)\<^sup>\ ; y \ (x ; y)\<^sup>\" by (metis add_right_isotone circ_left_slide circ_left_unfold mult_associative mult_right_isotone) -- {* Theorem 14.8 *} lemma circ_right_unfold_sub: "1 + x\<^sup>\ ; x \ x\<^sup>\" by (metis circ_mult_sub mult_left_one mult_right_one) -- {* Theorem 1.1 and Theorem 14.1 *} lemma circ_zero: "0\<^sup>\ = 1" by (metis add_right_zero circ_left_unfold mult_left_zero) -- {* Theorem 1.4 and Theorem 1.13 and Theorem 14.4 and Theorem 14.13 *} lemma circ_increasing: "x \ x\<^sup>\" by (metis add_right_upper_bound circ_left_unfold circ_right_unfold_sub mult_left_one mult_right_sub_dist_add_left order_trans) -- {* Theorem 1.3 and Theorem 14.3 *} lemma circ_reflexive: "1 \ x\<^sup>\" by (metis add_left_divisibility circ_left_unfold) -- {* Theorem 1.5 *} lemma circ_mult_increasing: "x \ x ; x\<^sup>\" by (metis circ_reflexive mult_right_isotone mult_right_one) -- {* Theorem 14.5 *} lemma circ_mult_increasing_2: "x \ x\<^sup>\ ; x" by (metis circ_reflexive mult_left_isotone mult_left_one_1) -- {* Theorem 1.11 and Theorem 14.11 *} lemma circ_transitive_equal: "x\<^sup>\ ; x\<^sup>\ = x\<^sup>\" by (metis add_idempotent circ_add_1 circ_left_unfold mult_associative) -- {* Theorem 1.17 and Theorem 14.17 *} lemma circ_circ_circ: "x\<^sup>\\<^sup>\\<^sup>\ = x\<^sup>\\<^sup>\" by (metis add_idempotent circ_add_1 circ_increasing circ_transitive_equal less_eq_def) -- {* Theorem 1.18 and Theorem 14.18 *} lemma circ_one: "1\<^sup>\ = 1\<^sup>\\<^sup>\" by (metis circ_circ_circ circ_zero) -- {* Theorem 14.21 *} lemma circ_add_sub: "(x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ \ (x + y)\<^sup>\" by (metis circ_add_1 circ_left_slide) lemma circ_plus_one: "x\<^sup>\ = 1 + x\<^sup>\" by (metis less_eq_def circ_reflexive) -- {* Theorem 1.12 and Theorem 14.12 *} lemma circ_rtc_2: "1 + x + x\<^sup>\ ; x\<^sup>\ = x\<^sup>\" by (metis add_associative circ_increasing circ_plus_one circ_transitive_equal less_eq_def) -- {* Theorem 1.2 and Theorem 14.2 *} lemma mult_zero_circ: "(x ; 0)\<^sup>\ = 1 + x ; 0" by (metis circ_left_unfold mult_associative mult_left_zero) lemma mult_zero_add_circ: "(x + y ; 0)\<^sup>\ = x\<^sup>\ ; (y ; 0)\<^sup>\" by (metis circ_add_1 mult_associative mult_left_zero) -- {* Theorem 14.6 *} lemma circ_plus_sub: "x\<^sup>\ ; x \ x ; x\<^sup>\" by (metis circ_left_slide mult_left_one mult_right_one) lemma circ_loop_fixpoint: "y ; (y\<^sup>\ ; z) + z = y\<^sup>\ ; z" by (metis add_commutative circ_left_unfold mult_associative mult_left_one mult_right_dist_add) -- {* Theorem 1.6 and Theorem 14.7 *} lemma left_plus_below_circ: "x ; x\<^sup>\ \ x\<^sup>\" by (metis add_right_upper_bound circ_left_unfold) lemma right_plus_below_circ: "x\<^sup>\ ; x \ x\<^sup>\" by (metis add_least_upper_bound circ_right_unfold_sub) lemma circ_add_upper_bound: "x \ z\<^sup>\ \ y \ z\<^sup>\ \ x + y \ z\<^sup>\" by (metis add_least_upper_bound) lemma circ_mult_upper_bound: "x \ z\<^sup>\ \ y \ z\<^sup>\ \ x ; y \ z\<^sup>\" by (metis mult_isotone circ_transitive_equal) lemma circ_sub_dist: "x\<^sup>\ \ (x + y)\<^sup>\" by (metis circ_add_sub circ_plus_one mult_left_one mult_right_sub_dist_add_left order_trans) lemma circ_sub_dist_1: "x \ (x + y)\<^sup>\" by (metis add_least_upper_bound circ_increasing) lemma circ_sub_dist_2: "x ; y \ (x + y)\<^sup>\" by (metis add_commutative circ_mult_upper_bound circ_sub_dist_1) -- {* Theorem 1.20 and Theorem 14.23 *} lemma circ_sub_dist_3: "x\<^sup>\ ; y\<^sup>\ \ (x + y)\<^sup>\" by (metis add_commutative circ_mult_upper_bound circ_sub_dist) -- {* Theorem 1 and Theorem 14 *} lemma circ_isotone: "x \ y \ x\<^sup>\ \ y\<^sup>\" by (metis circ_sub_dist less_eq_def) -- {* Theorem 1.21 and Theorem 14.24 *} lemma circ_add_2: "(x + y)\<^sup>\ \ (x\<^sup>\ ; y\<^sup>\)\<^sup>\" by (metis add_least_upper_bound circ_increasing circ_isotone circ_reflexive mult_isotone mult_left_one mult_right_one) lemma circ_sup_one_left_unfold: "1 \ x \ x ; x\<^sup>\ = x\<^sup>\" by (metis antisym less_eq_def mult_left_one mult_right_sub_dist_add_left left_plus_below_circ) lemma circ_sup_one_right_unfold: "1 \ x \ x\<^sup>\ ; x = x\<^sup>\" by (metis antisym less_eq_def mult_left_sub_dist_add_left mult_right_one right_plus_below_circ) -- {* Theorem 1.23 and Theorem 14.26 *} lemma circ_decompose_4: "(x\<^sup>\ ; y\<^sup>\)\<^sup>\ = x\<^sup>\ ; (y\<^sup>\ ; x\<^sup>\)\<^sup>\" by (metis add_associative add_commutative circ_add_1 circ_loop_fixpoint circ_plus_one circ_rtc_2 circ_transitive_equal mult_associative) -- {* Theorem 1.22 and Theorem 14.25 *} lemma circ_decompose_5: "(x\<^sup>\ ; y\<^sup>\)\<^sup>\ = (y\<^sup>\ ; x\<^sup>\)\<^sup>\" by (smt add_associative add_commutative add_left_zero circ_add_1 circ_decompose_4 mult_left_zero mult_right_one) lemma circ_decompose_6: "x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ = y\<^sup>\ ; (x ; y\<^sup>\)\<^sup>\" by (metis add_commutative circ_add_1) lemma circ_decompose_7: "(x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\ ; (x + y)\<^sup>\" by (metis circ_add_1 circ_decompose_6 circ_transitive_equal mult_associative) lemma circ_decompose_8: "(x + y)\<^sup>\ = (x + y)\<^sup>\ ; x\<^sup>\ ; y\<^sup>\" by (metis antisym eq_refl mult_associative mult_isotone mult_right_one circ_mult_upper_bound circ_reflexive circ_sub_dist_3) lemma circ_decompose_9: "(x\<^sup>\ ; y\<^sup>\)\<^sup>\ = x\<^sup>\ ; y\<^sup>\ ; (x\<^sup>\ ; y\<^sup>\)\<^sup>\" by (metis circ_decompose_4 mult_associative) lemma circ_decompose_10: "(x\<^sup>\ ; y\<^sup>\)\<^sup>\ = (x\<^sup>\ ; y\<^sup>\)\<^sup>\ ; x\<^sup>\ ; y\<^sup>\" by (metis add_right_upper_bound circ_loop_fixpoint circ_reflexive circ_sup_one_right_unfold mult_associative order_trans) lemma circ_back_loop_prefixpoint: "(z ; y\<^sup>\) ; y + z \ z ; y\<^sup>\" by (metis add_least_upper_bound circ_left_unfold mult_associative mult_left_sub_dist_add_left mult_right_isotone mult_right_one right_plus_below_circ) -- {* Theorem 1 and Theorem 14 *} lemma circ_loop_is_fixpoint: "is_fixpoint (\x . y ; x + z) (y\<^sup>\ ; z)" by (metis circ_loop_fixpoint is_fixpoint_def) -- {* Theorem 14 *} lemma circ_back_loop_is_prefixpoint: "is_prefixpoint (\x . x ; y + z) (z ; y\<^sup>\)" by (metis circ_back_loop_prefixpoint is_prefixpoint_def) -- {* Theorem 1.16 and Theorem 14.16 *} lemma circ_circ_add: "(1 + x)\<^sup>\ = x\<^sup>\\<^sup>\" by (metis add_commutative circ_add_1 circ_decompose_4 circ_zero mult_right_one) -- {* Theorem 14.14 *} lemma circ_circ_mult_sub: "x\<^sup>\ ; 1\<^sup>\ \ x\<^sup>\\<^sup>\" by (metis circ_increasing circ_isotone circ_mult_upper_bound circ_reflexive) -- {* Theorem 14.9 *} lemma left_plus_circ: "(x ; x\<^sup>\)\<^sup>\ = x\<^sup>\" by (smt add_idempotent circ_add_1 circ_loop_fixpoint circ_transitive_equal mult_right_dist_add) -- {* Theorem 1.10 and Theorem 14.10 *} lemma right_plus_circ: "(x\<^sup>\ ; x)\<^sup>\ = x\<^sup>\" by (metis add_commutative circ_isotone circ_loop_fixpoint circ_plus_sub circ_sub_dist eq_iff left_plus_circ) lemma circ_square: "(x ; x)\<^sup>\ \ x\<^sup>\" by (metis circ_increasing circ_isotone left_plus_circ mult_right_isotone) -- {* Theorem 1.19 *} lemma circ_mult_sub_add: "(x ; y)\<^sup>\ \ (x + y)\<^sup>\" by (metis add_left_upper_bound add_right_upper_bound circ_isotone circ_square mult_isotone order_trans) lemma circ_add_mult_zero: "x\<^sup>\ ; y = (x + y ; 0)\<^sup>\ ; y" proof - have "(x + y ; 0)\<^sup>\ ; y = x\<^sup>\ ; (1 + y ; 0) ; y" by (metis mult_zero_add_circ mult_zero_circ) also have "... = x\<^sup>\ ; (y + y ; 0)" by (metis mult_associative mult_left_one mult_left_zero mult_right_dist_add) also have "... = x\<^sup>\ ; y" by (metis add_commutative less_eq_def zero_right_mult_decreasing) finally show ?thesis by metis qed -- {* Theorem 14.22 *} lemma troeger_1: "(x + y)\<^sup>\ = x\<^sup>\ ; (1 + y ; (x + y)\<^sup>\)" by (metis circ_add_1 circ_left_unfold mult_associative) lemma troeger_2: "(x + y)\<^sup>\ ; z = x\<^sup>\ ; (y ; (x + y)\<^sup>\ ; z + z)" by (metis circ_add_1 circ_loop_fixpoint mult_associative) lemma troeger_3: "(x + y ; 0)\<^sup>\ = x\<^sup>\ ; (1 + y ; 0)" by (metis mult_zero_add_circ mult_zero_circ) lemma circ_add_sub_add_one_1: "x + y \ x\<^sup>\ ; (1 + y)" by (smt add_associative add_commutative add_idempotent circ_increasing circ_loop_fixpoint less_eq_def mult_left_sub_dist_add_left mult_right_one) lemma circ_add_sub_add_one_2: "x\<^sup>\ ; (x + y) \ x\<^sup>\ ; (1 + y)" by (metis circ_add_sub_add_one_1 circ_transitive_equal mult_associative mult_right_isotone) lemma circ_add_sub_add_one: "x ; x\<^sup>\ ; (x + y) \ x ; x\<^sup>\ ; (1 + y)" by (metis circ_add_sub_add_one_2 mult_associative mult_right_isotone) lemma circ_square_2: "(x ; x)\<^sup>\ ; (x + 1) \ x\<^sup>\" by (metis add_least_upper_bound circ_increasing circ_mult_upper_bound circ_reflexive circ_square) -- {* Theorem 1.25 and Theorem 14.28 *} lemma circ_extra_circ: "(y ; x\<^sup>\)\<^sup>\ = (y ; y\<^sup>\ ; x\<^sup>\)\<^sup>\" by (smt add_commutative add_idempotent circ_add_1 circ_left_unfold mult_associative) -- {* Theorem 14.15 *} lemma circ_circ_sub_mult: "1\<^sup>\ ; x\<^sup>\ \ x\<^sup>\\<^sup>\" by (metis circ_increasing circ_isotone circ_mult_upper_bound circ_reflexive) -- {* Theorem 14.27 *} lemma circ_decompose_11: "(x\<^sup>\ ; y\<^sup>\)\<^sup>\ = (x\<^sup>\ ; y\<^sup>\)\<^sup>\ ; x\<^sup>\" by (smt circ_add_1 circ_circ_add circ_decompose_4 circ_decompose_8 circ_rtc_2 circ_transitive_equal mult_associative) -- {* Theorem 14.20 *} lemma circ_mult_below_circ_circ: "(x ; y)\<^sup>\ \ (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" by (metis circ_increasing circ_isotone circ_reflexive dual_order.trans mult_left_isotone mult_right_isotone mult_right_one) -- {* Theorem 14 counterexamples *} lemma circ_right_unfold: "1 + x\<^sup>\ ; x = x\<^sup>\" nitpick [expect=genuine] oops lemma circ_mult: "1 + x ; (y ; x)\<^sup>\ ; y = (x ; y)\<^sup>\" nitpick [expect=genuine] oops lemma circ_slide: "(x ; y)\<^sup>\ ; x = x ; (y ; x)\<^sup>\" nitpick [expect=genuine] oops lemma circ_plus_same: "x\<^sup>\ ; x = x ; x\<^sup>\" nitpick [expect=genuine] oops lemma "1\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; 1\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_circ_mult_1: "x\<^sup>\ ; 1\<^sup>\ = x\<^sup>\\<^sup>\" nitpick [expect=genuine,card=7] oops lemma "x\<^sup>\ ; 1\<^sup>\ \ 1\<^sup>\ ; x\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_circ_mult: "1\<^sup>\ ; x\<^sup>\ = x\<^sup>\\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_add: "(x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ = (x + y)\<^sup>\" nitpick [expect=genuine,card=8] oops lemma circ_unfold_sum: "(x + y)\<^sup>\ = x\<^sup>\ + x\<^sup>\ ; y ; (x + y)\<^sup>\" nitpick [expect=genuine,card=7] oops lemma mult_zero_add_circ_2: "(x + y ; 0)\<^sup>\ = x\<^sup>\ + x\<^sup>\ ; y ; 0" nitpick [expect=genuine,card=7] oops lemma sub_mult_one_circ: "x ; 1\<^sup>\ \ 1\<^sup>\ ; x" nitpick [expect=genuine] oops lemma circ_back_loop_fixpoint: "(z ; y\<^sup>\) ; y + z = z ; y\<^sup>\" nitpick [expect=genuine] oops lemma circ_back_loop_is_fixpoint: "is_fixpoint (\x . x ; y + z) (z ; y\<^sup>\)" nitpick [expect=genuine] oops lemma "x\<^sup>\ ; y\<^sup>\ \ (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" nitpick [expect=genuine,card=7] oops end class bounded_left_conway_semiring = bounded_idempotent_left_semiring + left_conway_semiring begin lemma circ_top_1: "T\<^sup>\ = T" by (metis add_right_top antisym circ_left_unfold mult_left_sub_dist_add_left mult_right_one top_greatest) lemma circ_right_top_1: "x\<^sup>\ ; T = T" by (metis add_right_top circ_loop_fixpoint) lemma circ_left_top_1: "T ; x\<^sup>\ = T" by (metis antisym circ_plus_one mult_left_sub_dist_add_left mult_right_one top_greatest) lemma mult_top_circ_1: "(x ; T)\<^sup>\ = 1 + x ; T" by (metis circ_left_top_1 circ_left_unfold mult_associative) end class left_zero_conway_semiring = idempotent_left_zero_semiring + left_conway_semiring begin lemma mult_zero_add_circ_2: "(x + y ; 0)\<^sup>\ = x\<^sup>\ + x\<^sup>\ ; y ; 0" by (metis mult_associative mult_left_dist_add mult_right_one troeger_3) lemma circ_unfold_sum: "(x + y)\<^sup>\ = x\<^sup>\ + x\<^sup>\ ; y ; (x + y)\<^sup>\" by (metis mult_associative mult_left_dist_add mult_right_one troeger_1) end class left_conway_semiring_1 = left_conway_semiring + assumes circ_right_slide: "x ; (y ; x)\<^sup>\ \ (x ; y)\<^sup>\ ; x" begin -- {* Theorem 1.26 *} lemma circ_slide_1: "x ; (y ; x)\<^sup>\ = (x ; y)\<^sup>\ ; x" by (metis antisym circ_left_slide circ_right_slide) -- {* Theorem 1.9 *} lemma circ_right_unfold_1: "1 + x\<^sup>\ ; x = x\<^sup>\" by (metis circ_left_unfold circ_slide_1 mult_left_one mult_right_one) lemma circ_mult_1: "(x ; y)\<^sup>\ = 1 + x ; (y ; x)\<^sup>\ ; y" by (metis circ_left_unfold circ_slide_1 mult_associative) lemma circ_add_9: "(x + y)\<^sup>\ = (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" by (metis circ_add_1 circ_slide_1) -- {* Theorem 1.7 *} lemma circ_plus_same: "x\<^sup>\ ; x = x ; x\<^sup>\" by (metis circ_slide_1 mult_left_one mult_right_one) lemma circ_decompose_12: "x\<^sup>\ ; y\<^sup>\ \ (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" by (metis circ_add_9 circ_sub_dist_3) end class left_zero_conway_semiring_1 = left_zero_conway_semiring + left_conway_semiring_1 begin lemma circ_back_loop_fixpoint: "(z ; y\<^sup>\) ; y + z = z ; y\<^sup>\" by (metis add_commutative circ_left_unfold circ_plus_same mult_associative mult_left_dist_add mult_right_one) -- {* Theorem 1 *} lemma circ_back_loop_is_fixpoint: "is_fixpoint (\x . x ; y + z) (z ; y\<^sup>\)" by (metis circ_back_loop_fixpoint is_fixpoint_def) lemma circ_elimination: "x ; y = 0 \ x ; y\<^sup>\ \ x" by (metis add_left_zero circ_back_loop_fixpoint circ_plus_same mult_associative mult_left_zero order_refl) end class itering_1 = left_conway_semiring_1 + assumes circ_simulate: "z ; x \ y ; z \ z ; x\<^sup>\ \ y\<^sup>\ ; z" begin -- {* Theorem 1.15 *} lemma circ_circ_mult: "1\<^sup>\ ; x\<^sup>\ = x\<^sup>\\<^sup>\" by (metis antisym circ_circ_add circ_reflexive circ_simulate circ_sub_dist_3 circ_sup_one_left_unfold circ_transitive_equal mult_left_one order_refl) lemma sub_mult_one_circ: "x ; 1\<^sup>\ \ 1\<^sup>\ ; x" by (metis circ_simulate mult_left_one mult_right_one order_refl) lemma circ_import: "p \ p ; p \ p \ 1 \ p ; x \ x ; p \ p ; x\<^sup>\ = p ; (p ; x)\<^sup>\" apply rule apply (rule antisym) apply (smt antisym circ_simulate circ_slide_1 mult_associative mult_right_isotone mult_right_one order_refl test_preserves_equation) apply (metis circ_isotone mult_left_isotone mult_left_one mult_right_isotone) done end class itering_2 = left_conway_semiring_1 + assumes circ_simulate_right: "z ; x \ y ; z + w \ z ; x\<^sup>\ \ y\<^sup>\ ; (z + w ; x\<^sup>\)" assumes circ_simulate_left: "x ; z \ z ; y + w \ x\<^sup>\ ; z \ (z + x\<^sup>\ ; w) ; y\<^sup>\" begin subclass itering_1 apply unfold_locales apply (metis add_right_zero circ_simulate_right mult_left_zero) done lemma circ_simulate_left_1: "x ; z \ z ; y \ x\<^sup>\ ; z \ z ; y\<^sup>\ + x\<^sup>\ ; 0" by (metis add_right_zero circ_simulate_left mult_associative mult_left_zero mult_right_dist_add) lemma circ_separate_1: "y ; x \ x ; y \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" proof - have "y ; x \ x ; y \ y\<^sup>\ ; x ; y\<^sup>\ \ x ; y\<^sup>\ + y\<^sup>\ ; 0" by (smt circ_simulate_left_1 circ_transitive_equal mult_associative mult_left_isotone mult_left_zero mult_right_dist_add) thus ?thesis by (smt add_commutative circ_add_1 circ_simulate_right circ_sub_dist_3 less_eq_def mult_associative mult_left_zero zero_right_mult_decreasing) qed -- {* Theorem 1.14 *} lemma circ_circ_mult_1: "x\<^sup>\ ; 1\<^sup>\ = x\<^sup>\\<^sup>\" by (metis add_commutative circ_circ_add circ_separate_1 mult_left_one mult_right_one order_refl) end class itering_3 = itering_2 + left_zero_conway_semiring_1 begin lemma circ_simulate_1: "y ; x \ x ; y \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" by (smt add_associative add_right_zero circ_loop_fixpoint circ_simulate circ_simulate_left_1 mult_associative mult_left_zero mult_zero_add_circ_2) -- {* Theorem 4 *} lemma atomicity_refinement: "s = s ; q \ x = q ; x \ q ; b = 0 \ r ; b \ b ; r \ r ; l \ l ; r \ x ; l \ l ; x \ b ; l \ l ; b \ q ; l \ l ; q \ r\<^sup>\ ; q \ q ; r\<^sup>\ \ q \ 1 \ s ; (x + b + r + l)\<^sup>\ ; q \ s ; (x ; b\<^sup>\ ; q + r + l)\<^sup>\" proof assume 1: "s = s ; q \ x = q ; x \ q ; b = 0 \ r ; b \ b ; r \ r ; l \ l ; r \ x ; l \ l ; x \ b ; l \ l ; b \ q ; l \ l ; q \ r\<^sup>\ ; q \ q ; r\<^sup>\ \ q \ 1" hence "s ; (x + b + r + l)\<^sup>\ ; q = s ; l\<^sup>\ ; (x + b + r)\<^sup>\ ; q" by (smt add_commutative add_least_upper_bound circ_separate_1 mult_associative mult_left_sub_dist_add_right mult_right_dist_add order_trans) also have "... = s ; l\<^sup>\ ; b\<^sup>\ ; r\<^sup>\ ; q ; (x ; b\<^sup>\ ; r\<^sup>\ ; q)\<^sup>\" using 1 by (smt add_associative add_commutative circ_add_1 circ_separate_1 circ_slide_1 mult_associative) also have "... \ s ; l\<^sup>\ ; b\<^sup>\ ; r\<^sup>\ ; q ; (x ; b\<^sup>\ ; q ; r\<^sup>\)\<^sup>\" using 1 by (metis circ_isotone mult_associative mult_right_isotone) also have "... \ s ; q ; l\<^sup>\ ; b\<^sup>\ ; r\<^sup>\ ; (x ; b\<^sup>\ ; q ; r\<^sup>\)\<^sup>\" using 1 by (metis mult_left_isotone mult_right_isotone mult_right_one) also have "... \ s ; l\<^sup>\ ; q ; b\<^sup>\ ; r\<^sup>\ ; (x ; b\<^sup>\ ; q ; r\<^sup>\)\<^sup>\" using 1 by (metis circ_simulate mult_associative mult_left_isotone mult_right_isotone) also have "... \ s ; l\<^sup>\ ; r\<^sup>\ ; (x ; b\<^sup>\ ; q ; r\<^sup>\)\<^sup>\" using 1 by (metis add_left_zero circ_back_loop_fixpoint circ_plus_same mult_associative mult_left_zero mult_left_isotone mult_right_isotone mult_right_one) also have "... \ s ; (x ; b\<^sup>\ ; q + r + l)\<^sup>\" using 1 by (metis add_commutative circ_add_1 circ_sub_dist_3 mult_associative mult_right_isotone) finally show "s ; (x + b + r + l)\<^sup>\ ; q \ s ; (x ; b\<^sup>\ ; q + r + l)\<^sup>\" . qed end class itering = idempotent_left_zero_semiring + circ + assumes circ_add: "(x + y)\<^sup>\ = (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" assumes circ_mult: "(x ; y)\<^sup>\ = 1 + x ; (y ; x)\<^sup>\ ; y" assumes circ_simulate_right_plus: "z ; x \ y ; y\<^sup>\ ; z + w \ z ; x\<^sup>\ \ y\<^sup>\ ; (z + w ; x\<^sup>\)" assumes circ_simulate_left_plus: "x ; z \ z ; y\<^sup>\ + w \ x\<^sup>\ ; z \ (z + x\<^sup>\ ; w) ; y\<^sup>\" begin lemma circ_right_unfold: "1 + x\<^sup>\ ; x = x\<^sup>\" by (metis circ_mult mult_left_one mult_right_one) lemma circ_slide: "x ; (y ; x)\<^sup>\ = (x ; y)\<^sup>\ ; x" by (smt2 circ_mult mult_associative mult_left_dist_add mult_left_one mult_right_dist_add mult_right_one order_refl) -- {* Theorem 50.6 *} subclass itering_3 apply unfold_locales apply (metis circ_mult mult_left_one mult_right_one) -- {* Theorem 1.8 *} apply (metis circ_slide order_refl) apply (metis circ_add circ_slide) apply (metis circ_slide order_refl) apply (metis add_left_isotone circ_right_unfold mult_left_isotone mult_left_sub_dist_add_left mult_right_one order_trans circ_simulate_right_plus) apply (metis add_commutative add_left_upper_bound add_right_isotone circ_mult mult_right_isotone mult_right_one order_trans circ_simulate_left_plus) done lemma circ_simulate_right_plus_1: "z ; x \ y ; y\<^sup>\ ; z \ z ; x\<^sup>\ \ y\<^sup>\ ; z" by (metis add_right_zero circ_simulate_right_plus mult_left_zero) lemma circ_simulate_left_plus_1: "x ; z \ z ; y\<^sup>\ \ x\<^sup>\ ; z \ z ; y\<^sup>\ + x\<^sup>\ ; 0" by (smt add_right_zero circ_simulate_left_plus mult_associative mult_left_zero mult_right_dist_add) lemma circ_simulate_2: "y ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\ \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" apply (rule iffI) apply (smt add_associative add_right_zero circ_loop_fixpoint circ_simulate_left_plus_1 mult_associative mult_left_zero mult_zero_add_circ_2) apply (metis circ_increasing mult_left_isotone order_trans) done lemma circ_simulate_absorb: "y ; x \ x \ y\<^sup>\ ; x \ x + y\<^sup>\ ; 0" by (metis circ_simulate_left_plus_1 circ_zero mult_right_one) lemma circ_simulate_3: "y ; x\<^sup>\ \ x\<^sup>\ \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" by (metis add_least_upper_bound circ_reflexive circ_simulate_2 less_eq_def mult_right_isotone mult_right_one) lemma circ_separate_mult_1: "y ; x \ x ; y \ (x ; y)\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" by (metis circ_mult_sub_add circ_separate_1) -- {* Theorem 1.24 *} lemma circ_separate_unfold: "(y ; x\<^sup>\)\<^sup>\ = y\<^sup>\ + y\<^sup>\ ; y ; x ; x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" by (smt add_commutative circ_add circ_left_unfold circ_loop_fixpoint mult_associative mult_left_dist_add mult_right_one) -- {* Theorem 3 *} lemma separation: "y ; x \ x ; y\<^sup>\ \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" proof - have "y ; x \ x ; y\<^sup>\ \ y\<^sup>\ ; x ; y\<^sup>\ \ x ; y\<^sup>\ + y\<^sup>\ ; 0" by (smt circ_simulate_left_plus_1 circ_transitive_equal mult_associative mult_left_isotone mult_left_zero mult_right_dist_add) thus ?thesis by (smt add_commutative circ_add_1 circ_simulate_right circ_sub_dist_3 less_eq_def mult_associative mult_left_zero zero_right_mult_decreasing) qed -- {* Theorem 3 *} lemma simulation: "y ; x \ x ; y\<^sup>\ \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" by (metis add_right_upper_bound circ_isotone circ_mult_upper_bound circ_sub_dist separation) -- {* Theorem 3 *} lemma circ_simulate_4: "y ; x \ x ; x\<^sup>\ ; (1 + y) \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" proof assume "y ; x \ x ; x\<^sup>\ ; (1 + y)" hence "(1 + y) ; x \ x ; x\<^sup>\ ; (1 + y)" by (smt add_associative add_commutative add_left_upper_bound circ_back_loop_fixpoint less_eq_def mult_left_dist_add mult_left_one mult_right_dist_add mult_right_one) hence "y ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" by (metis circ_add_upper_bound circ_increasing circ_reflexive circ_simulate_right_plus_1 mult_right_isotone mult_right_sub_dist_add_right order_trans) thus "y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" by (metis circ_simulate_2) qed lemma circ_simulate_5: "y ; x \ x ; x\<^sup>\ ; (x + y) \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" by (metis circ_add_sub_add_one circ_simulate_4 order_trans) lemma circ_simulate_6: "y ; x \ x ; (x + y) \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" by (metis add_commutative circ_back_loop_fixpoint circ_simulate_5 mult_right_sub_dist_add_left order_trans) -- {* Theorem 3 *} lemma circ_separate_4: "y ; x \ x ; x\<^sup>\ ; (1 + y) \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" proof assume 1: "y ; x \ x ; x\<^sup>\ ; (1 + y)" hence "y ; x ; x\<^sup>\ \ x ; x\<^sup>\ + x ; x\<^sup>\ ; y ; x\<^sup>\" by (smt circ_transitive_equal less_eq_def mult_associative mult_left_dist_add mult_right_dist_add mult_right_one) also have "... \ x ; x\<^sup>\ + x ; x\<^sup>\ ; x\<^sup>\ ; y\<^sup>\" using 1 by (metis add_right_isotone circ_simulate_2 circ_simulate_4 mult_associative mult_right_isotone) finally have "y ; x ; x\<^sup>\ \ x ; x\<^sup>\ ; y\<^sup>\" by (metis circ_reflexive circ_transitive_equal less_eq_def mult_associative mult_right_isotone mult_right_one) hence "y\<^sup>\ ; (y\<^sup>\ ; x)\<^sup>\ \ x\<^sup>\ ; (y\<^sup>\ + y\<^sup>\ ; 0 ; (y\<^sup>\ ; x)\<^sup>\)" by (smt add_right_upper_bound circ_back_loop_fixpoint circ_simulate_left_plus_1 circ_simulate_right_plus circ_transitive_equal mult_associative order_trans) thus "(x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" by (smt add_commutative antisym circ_add_1 circ_slide circ_sub_dist_3 circ_transitive_equal less_eq_def mult_associative mult_left_zero mult_right_sub_dist_add_right zero_right_mult_decreasing) qed lemma circ_separate_5: "y ; x \ x ; x\<^sup>\ ; (x + y) \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" by (metis circ_add_sub_add_one circ_separate_4 order_trans) lemma circ_separate_6: "y ; x \ x ; (x + y) \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" by (metis add_commutative circ_back_loop_fixpoint circ_separate_5 mult_right_sub_dist_add_left order_trans) end class bounded_itering = bounded_idempotent_left_zero_semiring + itering begin -- {* Theorem 1 *} lemma circ_top: "T\<^sup>\ = T" by (metis add_right_top antisym circ_left_unfold mult_left_sub_dist_add_left mult_right_one top_greatest) lemma circ_right_top: "x\<^sup>\ ; T = T" by (metis add_right_top circ_loop_fixpoint) lemma circ_left_top: "T ; x\<^sup>\ = T" by (metis add_right_top circ_add circ_right_top circ_top) lemma mult_top_circ: "(x ; T)\<^sup>\ = 1 + x ; T" by (metis circ_left_top circ_mult mult_associative) -- {* Theorem 1 counterexamples *} lemma "1 = x\<^sup>\" nitpick [expect=genuine] oops lemma "x = x\<^sup>\" nitpick [expect=genuine] oops lemma "x = x ; x\<^sup>\" nitpick [expect=genuine] oops lemma "x ; x\<^sup>\ = x\<^sup>\" nitpick [expect=genuine] oops lemma "x\<^sup>\ = x\<^sup>\\<^sup>\" nitpick [expect=genuine] oops lemma "(x ; y)\<^sup>\ = (x + y)\<^sup>\" nitpick [expect=genuine] oops lemma "x\<^sup>\ ; y\<^sup>\ = (x + y)\<^sup>\" nitpick [expect=genuine,card=6] oops lemma "(x + y)\<^sup>\ = (x\<^sup>\ ; y\<^sup>\)\<^sup>\" nitpick [expect=genuine] oops lemma "1 = 1\<^sup>\" nitpick [expect=genuine] oops lemma "1 = (x ; 0)\<^sup>\" nitpick [expect=genuine] oops lemma "1 + x ; 0 = x\<^sup>\" nitpick [expect=genuine] oops lemma "x\<^sup>\ = x\<^sup>\ ; 1\<^sup>\" nitpick [expect=genuine] oops lemma "z + y ; x = x \ y\<^sup>\ ; z \ x" nitpick [expect=genuine] oops lemma "y ; x = x \ y\<^sup>\ ; x \ x" nitpick [expect=genuine] oops lemma "z + x ; y = x \ z ; y\<^sup>\ \ x" nitpick [expect=genuine] oops lemma "x ; y = x \ x ; y\<^sup>\ \ x" nitpick [expect=genuine] oops lemma "x = z + y ; x \ x \ y\<^sup>\ ; z" nitpick [expect=genuine] oops lemma "x = y ; x \ x \ y\<^sup>\" nitpick [expect=genuine] oops lemma "x ; z = z ; y \ x\<^sup>\ ; z \ z ; y\<^sup>\" nitpick [expect=genuine] oops lemma "x\<^sup>\ = (x ; x)\<^sup>\ ; (x + 1)" oops lemma "y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\ \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" oops lemma "y ; x \ (1 + x) ; y\<^sup>\ \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" oops lemma "y ; x \ x \ y\<^sup>\ ; x \ 1\<^sup>\ ; x" oops end class left_conway_semiring_L = left_conway_semiring + L + assumes one_circ_mult_split: "1\<^sup>\ ; x = L + x" assumes L_split_add: "x ; (y + L) \ x ; y + L" begin lemma L_def: "L = 1\<^sup>\ ; 0" by (metis add_right_zero one_circ_mult_split) lemma one_circ_split: "1\<^sup>\ = L + 1" by (metis mult_right_one one_circ_mult_split) lemma one_circ_circ_split: "1\<^sup>\\<^sup>\ = L + 1" by (metis circ_one one_circ_split) lemma sub_mult_one_circ: "x ; 1\<^sup>\ \ 1\<^sup>\ ; x" by (metis L_split_add add_commutative mult_right_one one_circ_mult_split) lemma one_circ_mult_split_2: "1\<^sup>\ ; x = x ; 1\<^sup>\ + L" by (smt add_associative add_commutative add_least_upper_bound circ_back_loop_prefixpoint less_eq_def one_circ_mult_split sub_mult_one_circ) lemma sub_mult_one_circ_split: "x ; 1\<^sup>\ \ x + L" by (metis add_commutative one_circ_mult_split sub_mult_one_circ) lemma sub_mult_one_circ_split_2: "x ; 1\<^sup>\ \ x + 1\<^sup>\" by (metis L_def add_right_isotone order_trans sub_mult_one_circ_split zero_right_mult_decreasing) lemma L_split: "x ; L \ x ; 0 + L" by (metis L_split_add add_left_zero) lemma L_left_zero: "L ; x = L" by (metis L_def mult_associative mult_left_zero) lemma one_circ_L: "1\<^sup>\ ; L = L" by (metis L_def circ_transitive_equal mult_associative) lemma mult_L_circ: "(x ; L)\<^sup>\ = 1 + x ; L" by (metis L_left_zero circ_left_unfold mult_associative) lemma mult_L_circ_mult: "(x ; L)\<^sup>\ ; y = y + x ; L" by (metis L_left_zero mult_L_circ mult_associative mult_left_one mult_right_dist_add) lemma circ_L: "L\<^sup>\ = L + 1" by (metis L_left_zero add_commutative circ_left_unfold) lemma L_below_one_circ: "L \ 1\<^sup>\" by (metis L_def zero_right_mult_decreasing) lemma circ_circ_mult_1: "x\<^sup>\ ; 1\<^sup>\ = x\<^sup>\\<^sup>\" by (metis L_left_zero add_commutative circ_add_1 circ_circ_add mult_zero_circ one_circ_split) lemma circ_circ_mult: "1\<^sup>\ ; x\<^sup>\ = x\<^sup>\\<^sup>\" by (metis antisym circ_circ_mult_1 circ_circ_sub_mult sub_mult_one_circ) lemma circ_circ_split: "x\<^sup>\\<^sup>\ = L + x\<^sup>\" by (metis circ_circ_mult one_circ_mult_split) lemma circ_add_6: "L + (x + y)\<^sup>\ = (x\<^sup>\ ; y\<^sup>\)\<^sup>\" by (metis add_associative add_commutative circ_add_1 circ_circ_add circ_circ_split circ_decompose_4) end class itering_L = itering + L + assumes L_def: "L = 1\<^sup>\ ; 0" begin lemma one_circ_split: "1\<^sup>\ = L + 1" by (metis L_def add_commutative antisym circ_add_upper_bound circ_reflexive circ_simulate_absorb mult_right_one order_refl zero_right_mult_decreasing) lemma one_circ_mult_split: "1\<^sup>\ ; x = L + x" by (metis L_def add_commutative circ_loop_fixpoint mult_associative mult_left_zero mult_zero_circ one_circ_split) lemma sub_mult_one_circ_split: "x ; 1\<^sup>\ \ x + L" by (metis add_commutative one_circ_mult_split sub_mult_one_circ) lemma sub_mult_one_circ_split_2: "x ; 1\<^sup>\ \ x + 1\<^sup>\" by (metis L_def add_right_isotone order_trans sub_mult_one_circ_split zero_right_mult_decreasing) lemma L_split: "x ; L \ x ; 0 + L" by (smt L_def mult_associative mult_left_isotone mult_right_dist_add sub_mult_one_circ_split_2) subclass left_conway_semiring_L apply unfold_locales apply (metis L_def add_commutative circ_loop_fixpoint mult_associative mult_left_zero mult_zero_circ one_circ_split) apply (metis add_commutative mult_associative mult_left_isotone one_circ_mult_split sub_mult_one_circ) done lemma circ_left_induct_mult_L: "L \ x \ x ; y \ x \ x ; y\<^sup>\ \ x" by (metis circ_one circ_simulate less_eq_def one_circ_mult_split) lemma circ_left_induct_mult_iff_L: "L \ x \ x ; y \ x \ x ; y\<^sup>\ \ x" by (smt add_least_upper_bound circ_back_loop_fixpoint circ_left_induct_mult_L less_eq_def) lemma circ_left_induct_L: "L \ x \ x ; y + z \ x \ z ; y\<^sup>\ \ x" by (metis add_least_upper_bound circ_left_induct_mult_L less_eq_def mult_right_dist_add) end (* end theory KleeneAlgebra imports Itering begin *) class star = fixes star :: "'a \ 'a" ("_\<^sup>\" [100] 100) class left_kleene_algebra = idempotent_left_semiring + star + assumes star_left_unfold : "1 + y ; y\<^sup>\ \ y\<^sup>\" assumes star_left_induct : "z + y ; x \ x \ y\<^sup>\ ; z \ x" begin lemma star_left_unfold_equal: "1 + x ; x\<^sup>\ = x\<^sup>\" by (smt add_right_isotone antisym mult_right_isotone mult_right_one star_left_induct star_left_unfold) lemma star_left_slide: "(x ; y)\<^sup>\ ; x \ x ; (y ; x)\<^sup>\" by (metis mult_associative mult_left_sub_dist_add mult_right_one star_left_induct star_left_unfold_equal) lemma star_isotone: "x \ y \ x\<^sup>\ \ y\<^sup>\" by (metis add_right_isotone mult_left_isotone order_trans star_left_unfold mult_right_one star_left_induct) lemma star_add_1_sub: "x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ \ (x + y)\<^sup>\" proof - have "x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ \ x\<^sup>\ ; (y ; (x + y)\<^sup>\)\<^sup>\" by (smt add_left_upper_bound mult_right_isotone star_isotone) also have "... \ x\<^sup>\ ; ((x + y) ; (x + y)\<^sup>\)\<^sup>\" by (smt add_right_upper_bound mult_left_isotone mult_right_isotone star_isotone) also have "... \ x\<^sup>\ ; (x + y)\<^sup>\\<^sup>\" by (smt add_least_upper_bound mult_right_isotone star_isotone star_left_unfold) also have "... \ (x + y)\<^sup>\ ; (x + y)\<^sup>\\<^sup>\" by (smt add_left_upper_bound mult_left_isotone star_isotone) also have "... \ (x + y)\<^sup>\" by (smt add_least_upper_bound mult_right_one star_left_induct star_left_unfold) finally show "x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ \ (x + y)\<^sup>\" by smt qed lemma star_add_1: "(x + y)\<^sup>\ = x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" apply (rule antisym) apply (smt add_least_upper_bound add_left_upper_bound add_right_upper_bound mult_associative mult_left_one mult_right_dist_add mult_right_one star_left_induct star_left_unfold_equal) apply (smt star_add_1_sub) done end -- {* Theorem 50.1 *} sublocale left_kleene_algebra < star!: left_conway_semiring where circ = star apply unfold_locales apply (metis star_left_unfold_equal) apply (metis star_left_slide) apply (metis star_add_1) done context left_kleene_algebra begin -- {* Many lemmas in this class are taken from Georg Struth's Isabelle theories. *} lemma star_sub_one: "x \ 1 \ x\<^sup>\ = 1" by (metis add_right_isotone eq_iff less_eq_def mult_right_one star.circ_plus_one star_left_induct) lemma star_one: "1\<^sup>\ = 1" by (metis eq_iff star_sub_one) lemma star_left_induct_mult: "x ; y \ y \ x\<^sup>\ ; y \ y" by (metis add_commutative less_eq_def order_refl star_left_induct) lemma star_left_induct_mult_iff: "x ; y \ y \ x\<^sup>\ ; y \ y" by (metis mult_associative mult_left_isotone mult_left_one mult_right_isotone order_trans star_left_induct_mult star.circ_reflexive star.left_plus_below_circ) lemma star_involutive: "x\<^sup>\ = x\<^sup>\\<^sup>\" by (smt antisym less_eq_def mult_left_sub_dist_add_left mult_right_one star_left_induct star.circ_plus_one star.left_plus_below_circ star.circ_transitive_equal) lemma star_sup_one: "(1 + x)\<^sup>\ = x\<^sup>\" by (metis star.circ_circ_add star_involutive) lemma star_left_induct_equal: "z + x ; y = y \ x\<^sup>\ ; z \ y" by (metis order_refl star_left_induct) lemma star_left_induct_mult_equal: "x ; y = y \ x\<^sup>\ ; y \ y" by (metis order_refl star_left_induct_mult) lemma star_star_upper_bound: "x\<^sup>\ \ z\<^sup>\ \ x\<^sup>\\<^sup>\ \ z\<^sup>\" by (metis star_involutive) lemma star_simulation_left: "x ; z \ z ; y \ x\<^sup>\ ; z \ z ; y\<^sup>\" by (smt add_commutative add_least_upper_bound mult_right_dist_add less_eq_def mult_associative mult_right_one star.left_plus_below_circ star.circ_increasing star_left_induct star_involutive star.circ_isotone star.circ_reflexive mult_left_sub_dist_add_left) lemma quasicomm_1: "y ; x \ x ; (x + y)\<^sup>\ \ y\<^sup>\ ; x \ x ; (x + y)\<^sup>\" by (smt mult_isotone order_refl order_trans star.circ_increasing star_involutive star_simulation_left) lemma star_rtc_3: "1 + x + y ; y = y \ x\<^sup>\ \ y" by (metis add_least_upper_bound less_eq_def mult_left_sub_dist_add_left mult_right_one star_left_induct_mult_iff star.circ_sub_dist) lemma star_decompose_1: "(x + y)\<^sup>\ = (x\<^sup>\ ; y\<^sup>\)\<^sup>\" apply (rule antisym) apply (smt add_least_upper_bound mult_isotone mult_left_one mult_right_one star.circ_increasing star.circ_isotone star.circ_reflexive) apply (smt star.circ_isotone star.circ_sub_dist_3 star_involutive) done lemma star_sum: "(x + y)\<^sup>\ = (x\<^sup>\ + y\<^sup>\)\<^sup>\" by (metis star_decompose_1 star_involutive) lemma star_decompose_3: "(x\<^sup>\ ; y\<^sup>\)\<^sup>\ = x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" by (metis star_decompose_1 star.circ_add_1) lemma star_loop_least_fixpoint: "y ; x + z = x \ y\<^sup>\ ; z \ x" by (metis add_commutative star_left_induct_equal) lemma star_loop_is_least_fixpoint: "is_least_fixpoint (\x . y ; x + z) (y\<^sup>\ ; z)" by (smt is_least_fixpoint_def star.circ_loop_fixpoint star_loop_least_fixpoint) lemma star_loop_mu: "\ (\x . y ; x + z) = y\<^sup>\ ; z" by (metis least_fixpoint_same star_loop_is_least_fixpoint) lemma affine_has_least_fixpoint: "has_least_fixpoint (\x . y ; x + z)" by (metis has_least_fixpoint_def star_loop_is_least_fixpoint) lemma circ_add: "(x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ = (x + y)\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_mult: "1 + x ; (y ; x)\<^sup>\ ; y = (x ; y)\<^sup>\" nitpick [expect=genuine] oops lemma circ_plus_same: "x\<^sup>\ ; x = x ; x\<^sup>\" nitpick [expect=genuine] oops lemma circ_unfold_sum: "(x + y)\<^sup>\ = x\<^sup>\ + x\<^sup>\ ; y ; (x + y)\<^sup>\" nitpick [expect=genuine,card=8] oops lemma mult_zero_add_circ_2: "(x + y ; 0)\<^sup>\ = x\<^sup>\ + x\<^sup>\ ; y ; 0" nitpick [expect=genuine,card=7] oops lemma circ_simulate_left: "x ; z \ z ; y + w \ x\<^sup>\ ; z \ (z + x\<^sup>\ ; w) ; y\<^sup>\" nitpick [expect=genuine] oops lemma circ_simulate_1: "y ; x \ x ; y \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_separate_1: "y ; x \ x ; y \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma atomicity_refinement: "s = s ; q \ x = q ; x \ q ; b = 0 \ r ; b \ b ; r \ r ; l \ l ; r \ x ; l \ l ; x \ b ; l \ l ; b \ q ; l \ l ; q \ r\<^sup>\ ; q \ q ; r\<^sup>\ \ q \ 1 \ s ; (x + b + r + l)\<^sup>\ ; q \ s ; (x ; b\<^sup>\ ; q + r + l)\<^sup>\" nitpick [expect=genuine] oops lemma circ_simulate_left_plus: "x ; z \ z ; y\<^sup>\ + w \ x\<^sup>\ ; z \ (z + x\<^sup>\ ; w) ; y\<^sup>\" nitpick [expect=genuine] oops lemma circ_separate_unfold: "(y ; x\<^sup>\)\<^sup>\ = y\<^sup>\ + y\<^sup>\ ; y ; x ; x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" nitpick [expect=genuine] oops lemma separation: "y ; x \ x ; y\<^sup>\ \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_simulate_4: "y ; x \ x ; x\<^sup>\ ; (1 + y) \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_simulate_5: "y ; x \ x ; x\<^sup>\ ; (x + y) \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_simulate_6: "y ; x \ x ; (x + y) \ y\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_separate_4: "y ; x \ x ; x\<^sup>\ ; (1 + y) \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_separate_5: "y ; x \ x ; x\<^sup>\ ; (x + y) \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_separate_6: "y ; x \ x ; (x + y) \ (x + y)\<^sup>\ = x\<^sup>\ ; y\<^sup>\" nitpick [expect=genuine,card=7] oops end class strong_left_kleene_algebra = left_kleene_algebra + assumes star_right_induct: "z + x ; y \ x \ z ; y\<^sup>\ \ x" begin lemma star_plus: "y\<^sup>\ ; y = y ; y\<^sup>\" by (smt add_least_upper_bound antisym less_eq_def mult_left_one mult_right_dist_add star.circ_plus_sub star_left_unfold star_right_induct) lemma star_slide: "(x ; y)\<^sup>\ ; x = x ; (y ; x)\<^sup>\" by (smt add_least_upper_bound antisym mult_associative mult_left_isotone mult_left_one mult_right_one order_trans star_left_slide star_left_unfold star_right_induct) lemma star_simulation_right: "z ; x \ y ; z \ z ; x\<^sup>\ \ y\<^sup>\ ; z" by (smt add_commutative add_least_upper_bound add_left_upper_bound mult_associative order_trans star.circ_loop_fixpoint star_left_induct star_plus star_right_induct) end sublocale strong_left_kleene_algebra < star!: itering_1 where circ = star apply unfold_locales apply (metis star_slide order_refl) apply (metis star_simulation_right) done context strong_left_kleene_algebra begin lemma star_right_induct_mult: "y ; x \ y \ y ; x\<^sup>\ \ y" by (metis add_least_upper_bound eq_refl star_right_induct) lemma star_right_induct_mult_iff: "y ; x \ y \ y ; x\<^sup>\ \ y" by (metis mult_right_isotone order_trans star.circ_increasing star_right_induct_mult) lemma star_simulation_right_equal: "z ; x = y ; z \ z ; x\<^sup>\ = y\<^sup>\ ; z" by (metis eq_iff star_simulation_left star_simulation_right) lemma star_simulation_star: "x ; y \ y ; x \ x\<^sup>\ ; y\<^sup>\ \ y\<^sup>\ ; x\<^sup>\" by (metis star_simulation_left star_simulation_right) lemma star_right_induct_equal: "z + y ; x = y \ z ; x\<^sup>\ \ y" by (metis order_refl star_right_induct) lemma star_right_induct_mult_equal: "y ; x = y \ y ; x\<^sup>\ \ y" by (metis order_refl star_right_induct_mult) lemma star_back_loop_least_fixpoint: "x ; y + z = x \ z ; y\<^sup>\ \ x" by (metis add_commutative star_right_induct_equal) lemma star_back_loop_is_least_fixpoint: "is_least_fixpoint (\x . x ; y + z) (z ; y\<^sup>\)" by (smt add_commutative add_right_isotone antisym is_least_fixpoint_def mult_left_isotone star.circ_back_loop_prefixpoint star_back_loop_least_fixpoint star_right_induct) lemma star_back_loop_mu: "\ (\x . x ; y + z) = z ; y\<^sup>\" by (metis least_fixpoint_same star_back_loop_is_least_fixpoint) lemma star_square: "x\<^sup>\ = (1 + x) ; (x ; x)\<^sup>\" proof - let ?f = "\y . y ; x + 1" have 1: "isotone ?f" by (smt add_left_isotone isotone_def mult_left_isotone) have 2: "?f \ ?f = (\y . y ; (x ; x) + (1 + x))" by (simp add: add_associative add_commutative mult_associative mult_left_one mult_right_dist_add o_def) thus ?thesis using 1 by (metis mu_square mult_left_one star_back_loop_mu has_least_fixpoint_def star_back_loop_is_least_fixpoint) qed lemma star_square_2: "x\<^sup>\ = (x ; x)\<^sup>\ ; (x + 1)" by (smt add_commutative antisym mult_left_one mult_left_sub_dist_add mult_right_dist_add mult_right_one star.circ_square_2 star_slide star_square) lemma star_circ_simulate_right_plus: "z ; x \ y ; y\<^sup>\ ; z + w \ z ; x\<^sup>\ \ y\<^sup>\ ; (z + w ; x\<^sup>\)" proof assume 1: "z ; x \ y ; y\<^sup>\ ; z + w" have "(z + w ; x\<^sup>\) ; x \ z ; x + w ; x\<^sup>\" by (metis add_right_isotone mult_associative mult_right_dist_add mult_right_isotone star.circ_increasing star.circ_transitive_equal) also have "... \ y ; y\<^sup>\ ; z + w + w ; x\<^sup>\" using 1 by (metis add_left_isotone) also have "... \ y ; y\<^sup>\ ; z + w ; x\<^sup>\" by (metis add_least_upper_bound add_right_isotone add_right_upper_bound star.circ_back_loop_prefixpoint) also have "... \ y\<^sup>\ ; (z + w ; x\<^sup>\)" by (metis add_least_upper_bound mult_isotone mult_left_isotone mult_left_one mult_left_sub_dist_add_left star.circ_reflexive star.left_plus_below_circ) finally have "y\<^sup>\ ; (z + w ; x\<^sup>\) ; x \ y\<^sup>\ ; (z + w ; x\<^sup>\)" by (metis mult_associative mult_right_isotone star.circ_transitive_equal) thus "z ; x\<^sup>\ \ y\<^sup>\ ; (z + w ; x\<^sup>\)" by (metis add_least_upper_bound star_right_induct mult_left_sub_dist_add_left star.circ_loop_fixpoint) qed lemma star_circ_simulate_left_plus: "x ; z \ z ; y\<^sup>\ + w \ x\<^sup>\ ; z \ (z + x\<^sup>\ ; w) ; y\<^sup>\" nitpick [expect=genuine,card=7] oops end class left_zero_kleene_algebra = idempotent_left_zero_semiring + strong_left_kleene_algebra begin lemma star_star_absorb: "y\<^sup>\ ; (y\<^sup>\ ; x)\<^sup>\ ; y\<^sup>\ = (y\<^sup>\ ; x)\<^sup>\ ; y\<^sup>\" by (metis add_commutative mult_associative star.circ_decompose_4 star.circ_slide_1 star_decompose_1 star_decompose_3) lemma star_circ_simulate_left_plus: "x ; z \ z ; y\<^sup>\ + w \ x\<^sup>\ ; z \ (z + x\<^sup>\ ; w) ; y\<^sup>\" proof assume 1: "x ; z \ z ; y\<^sup>\ + w" have "x ; ((z + x\<^sup>\ ; w) ; y\<^sup>\) \ x ; z ; y\<^sup>\ + x\<^sup>\ ; w ; y\<^sup>\" by (smt add_right_isotone mult_associative mult_left_dist_add mult_right_dist_add mult_right_sub_dist_add_left star.circ_loop_fixpoint) also have "... \ (z + w + x\<^sup>\ ; w) ; y\<^sup>\" using 1 by (smt add_left_divisibility add_left_isotone mult_associative mult_right_dist_add star.circ_transitive_equal) also have "... = (z + x\<^sup>\ ; w) ; y\<^sup>\" by (metis add_associative add_right_upper_bound less_eq_def star.circ_loop_fixpoint) finally show "x\<^sup>\ ; z \ (z + x\<^sup>\ ; w) ; y\<^sup>\" by (metis add_least_upper_bound mult_left_sub_dist_add_left mult_right_one star.circ_right_unfold_1 star_left_induct) qed end -- {* Theorem 2.1 *} sublocale left_zero_kleene_algebra < star!: itering where circ = star apply unfold_locales apply (metis star.circ_add_9) apply (metis star.circ_mult_1) apply (rule star_circ_simulate_right_plus) apply (rule star_circ_simulate_left_plus) done class kleene_algebra = left_zero_kleene_algebra + idempotent_semiring class left_kleene_conway_semiring = left_kleene_algebra + left_conway_semiring begin lemma star_below_circ: "x\<^sup>\ \ x\<^sup>\" by (metis circ_left_unfold mult_right_one order_refl star_left_induct) lemma star_zero_below_circ_mult: "x\<^sup>\ ; 0 \ x\<^sup>\ ; y" by (metis mult_isotone star_below_circ zero_least) lemma star_mult_circ: "x\<^sup>\ ; x\<^sup>\ = x\<^sup>\" by (metis add_right_divisibility antisym circ_left_unfold star_left_induct_mult star.circ_loop_fixpoint) lemma circ_mult_star: "x\<^sup>\ ; x\<^sup>\ = x\<^sup>\" by (metis add_associative add_least_upper_bound circ_left_unfold circ_rtc_2 eq_iff left_plus_circ star.circ_add_sub star.circ_back_loop_prefixpoint star.circ_increasing star_below_circ star_mult_circ star_sup_one) lemma circ_star: "x\<^sup>\\<^sup>\ = x\<^sup>\" by (metis circ_left_unfold left_plus_circ less_def less_le star.circ_increasing star_below_circ star_sup_one) lemma star_circ: "x\<^sup>\\<^sup>\ = x\<^sup>\\<^sup>\" by (metis antisym circ_circ_add circ_sub_dist less_eq_def star.circ_rtc_2 star_below_circ) lemma circ_add_3: "(x\<^sup>\ ; y\<^sup>\)\<^sup>\ \ (x + y)\<^sup>\" by (metis circ_add_1 circ_isotone circ_left_unfold circ_star mult_left_sub_dist_add_left mult_right_isotone mult_right_one star.circ_isotone) end class left_zero_kleene_conway_semiring = left_zero_kleene_algebra + itering begin subclass left_kleene_conway_semiring .. lemma circ_isolate: "x\<^sup>\ = x\<^sup>\ ; 0 + x\<^sup>\" by (metis add_commutative antisym circ_add_upper_bound circ_mult_star circ_simulate_absorb star.left_plus_below_circ star_below_circ zero_right_mult_decreasing) lemma circ_isolate_mult: "x\<^sup>\ ; y = x\<^sup>\ ; 0 + x\<^sup>\ ; y" by (metis circ_isolate mult_associative mult_left_zero mult_right_dist_add) lemma circ_isolate_mult_sub: "x\<^sup>\ ; y \ x\<^sup>\ + x\<^sup>\ ; y" by (metis add_left_isotone circ_isolate_mult zero_right_mult_decreasing) lemma circ_sub_decompose: "(x\<^sup>\ ; y)\<^sup>\ \ (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" by (smt add_commutative add_least_upper_bound add_right_upper_bound circ_back_loop_fixpoint circ_isolate_mult mult_zero_add_circ_2 zero_right_mult_decreasing) lemma circ_add_4: "(x + y)\<^sup>\ = (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" apply (rule antisym) apply (smt circ_add circ_sub_decompose circ_transitive_equal mult_associative mult_left_isotone) apply (smt circ_add circ_isotone mult_left_isotone star_below_circ) done lemma circ_add_5: "(x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ = (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" by (metis circ_add circ_add_4) lemma plus_circ: "(x\<^sup>\ ; x)\<^sup>\ = x\<^sup>\" by (smt add_idempotent circ_add_4 circ_decompose_7 circ_star star.circ_decompose_5 star.right_plus_circ) lemma "(x\<^sup>\ ; y ; x\<^sup>\)\<^sup>\ = (x\<^sup>\ ; y)\<^sup>\" nitpick [expect=genuine] oops end class bounded_left_kleene_algebra = bounded_idempotent_left_semiring + left_kleene_algebra sublocale bounded_left_kleene_algebra < star!: bounded_left_conway_semiring where circ = star .. class bounded_left_zero_kleene_algebra = bounded_idempotent_left_semiring + left_zero_kleene_algebra sublocale bounded_left_zero_kleene_algebra < star!: bounded_itering where circ = star .. class bounded_kleene_algebra = bounded_idempotent_semiring + kleene_algebra sublocale bounded_kleene_algebra < star!: bounded_itering where circ = star .. (* end theory OmegaAlgebra imports KleeneAlgebra begin *) class omega = fixes omega :: "'a \ 'a" ("_\<^sup>\" [100] 100) class left_omega_algebra = left_kleene_algebra + omega + assumes omega_unfold: "y\<^sup>\ = y ; y\<^sup>\" assumes omega_induct: "x \ z + y ; x \ x \ y\<^sup>\ + y\<^sup>\ ; z" begin -- {* Many lemmas in this class are taken from Georg Struth's Isabelle theories. *} lemma star_zero_below_omega: "x\<^sup>\ ; 0 \ x\<^sup>\" by (metis add_left_zero omega_unfold star_left_induct_equal) lemma star_zero_below_omega_zero: "x\<^sup>\ ; 0 \ x\<^sup>\ ; 0" by (metis add_left_zero mult_associative omega_unfold star_left_induct_equal) lemma omega_induct_mult: "y \ x ; y \ y \ x\<^sup>\" by (metis add_commutative add_left_zero less_eq_def omega_induct star_zero_below_omega) lemma omega_sub_dist: "x\<^sup>\ \ (x+y)\<^sup>\" by (metis mult_right_sub_dist_add_left omega_induct_mult omega_unfold) lemma omega_isotone: "x \ y \ x\<^sup>\ \ y\<^sup>\" by (metis less_eq_def omega_sub_dist) lemma omega_induct_equal: "y = z + x ; y \ y \ x\<^sup>\ + x\<^sup>\ ; z" by (metis omega_induct order_refl) lemma omega_zero: "0\<^sup>\ = 0" by (metis mult_left_zero omega_unfold) lemma omega_one_greatest: "x \ 1\<^sup>\" by (metis mult_left_one omega_induct_mult order_refl) lemma star_mult_omega: "x\<^sup>\ ; x\<^sup>\ = x\<^sup>\" by (metis antisym_conv mult_isotone omega_unfold star.circ_increasing star_left_induct_mult_equal star_left_induct_mult_iff) lemma omega_sub_vector: "x\<^sup>\ ; y \ x\<^sup>\" by (metis mult_associative omega_induct_mult omega_unfold order_refl) lemma omega_simulation: "z ; x \ y ; z \ z ; x\<^sup>\ \ y\<^sup>\ " by (smt less_eq_def mult_associative mult_right_sub_dist_add_left omega_induct_mult omega_unfold) lemma omega_omega: "x\<^sup>\\<^sup>\ \ x\<^sup>\" by (metis omega_sub_vector omega_unfold) lemma left_plus_omega: "(x ; x\<^sup>\)\<^sup>\ = x\<^sup>\" by (metis antisym mult_associative omega_induct_mult omega_unfold order_refl star.left_plus_circ star_mult_omega) lemma omega_slide: "x ; (y ; x)\<^sup>\ = (x ; y)\<^sup>\" by (metis antisym mult_associative mult_right_isotone omega_simulation omega_unfold order_refl) lemma omega_simulation_2: "y ; x \ x ; y \ (x ; y)\<^sup>\ \ x\<^sup>\" by (metis less_eq_def mult_right_isotone omega_induct_mult omega_slide omega_sub_dist) lemma wagner: "(x + y)\<^sup>\ = x ; (x + y)\<^sup>\ + z \ (x + y)\<^sup>\ = x\<^sup>\ + x\<^sup>\ ; z" by (metis add_commutative add_least_upper_bound eq_iff omega_induct omega_sub_dist star_left_induct) lemma right_plus_omega: "(x\<^sup>\ ; x)\<^sup>\ = x\<^sup>\" by (metis left_plus_omega omega_slide star_mult_omega) lemma omega_sub_dist_1: "(x ; y\<^sup>\)\<^sup>\ \ (x + y)\<^sup>\" by (metis add_least_upper_bound left_plus_omega mult_isotone mult_left_one mult_right_dist_add omega_isotone order_refl star_decompose_1 star.circ_increasing star.circ_plus_one) lemma omega_sub_dist_2: "(x\<^sup>\ ; y)\<^sup>\ \ (x + y)\<^sup>\" by (metis add_commutative mult_isotone omega_slide omega_sub_dist_1 star_mult_omega star.circ_sub_dist) lemma omega_star: "(x\<^sup>\)\<^sup>\ = 1 + x\<^sup>\" by (metis eq_iff mult_left_sub_dist_add_left mult_right_one omega_sub_vector star.circ_left_unfold) lemma omega_mult_omega_star: "x\<^sup>\ ; x\<^sup>\\<^sup>\ = x\<^sup>\" by (metis add_least_upper_bound antisym omega_sub_vector star.circ_back_loop_prefixpoint) lemma omega_sum_unfold_1: "(x + y)\<^sup>\ = x\<^sup>\ + x\<^sup>\ ; y ; (x + y)\<^sup>\" by (metis mult_associative mult_right_dist_add omega_unfold wagner) lemma omega_sum_unfold_2: "(x + y)\<^sup>\ \ (x\<^sup>\ ; y)\<^sup>\ + (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" by (metis omega_induct_equal omega_sum_unfold_1) lemma omega_sum_unfold_3: "(x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ \ (x + y)\<^sup>\" by (metis omega_sum_unfold_1 star_left_induct_equal) lemma omega_decompose: "(x + y)\<^sup>\ = (x\<^sup>\ ; y)\<^sup>\ + (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" by (metis add_least_upper_bound antisym omega_sub_dist_2 omega_sum_unfold_2 omega_sum_unfold_3) lemma omega_loop_fixpoint: "y ; (y\<^sup>\ + y\<^sup>\ ; z) + z = y\<^sup>\ + y\<^sup>\ ; z" apply (rule antisym) apply (smt add_commutative add_least_upper_bound add_right_isotone add_right_upper_bound mult_left_sub_dist_add_left mult_right_isotone omega_induct omega_unfold order_trans star.circ_loop_fixpoint) apply (smt add_associative add_left_isotone mult_left_sub_dist_add omega_unfold star.circ_loop_fixpoint) done lemma omega_loop_greatest_fixpoint: "y ; x + z = x \ x \ y\<^sup>\ + y\<^sup>\ ; z" by (metis add_commutative omega_induct_equal) lemma omega_square: "x\<^sup>\ = (x ; x)\<^sup>\" by (metis antisym mult_associative omega_induct_mult omega_mult_omega_star omega_slide omega_sub_vector omega_unfold) lemma mult_zero_omega: "(x ; 0)\<^sup>\ = x ; 0" by (metis mult_left_zero omega_slide) lemma mult_zero_add_omega: "(x + y ; 0)\<^sup>\ = x\<^sup>\ + x\<^sup>\ ; y ; 0" by (smt add_associative add_commutative add_idempotent mult_associative mult_left_one mult_left_zero mult_right_dist_add mult_zero_omega star.mult_zero_circ omega_decompose) lemma omega_mult_star: "x\<^sup>\ ; x\<^sup>\ = x\<^sup>\" by (metis antisym mult_left_sub_dist_add_left mult_right_one omega_sub_vector star.circ_plus_one) lemma omega_loop_is_greatest_fixpoint: "is_greatest_fixpoint (\x . y ; x + z) (y\<^sup>\ + y\<^sup>\ ; z)" by (smt is_greatest_fixpoint_def omega_loop_fixpoint omega_loop_greatest_fixpoint) lemma omega_loop_nu: "\ (\x . y ; x + z) = y\<^sup>\ + y\<^sup>\ ; z" by (metis greatest_fixpoint_same omega_loop_is_greatest_fixpoint) lemma omega_loop_zero_is_greatest_fixpoint: "is_greatest_fixpoint (\x . y ; x) (y\<^sup>\)" by (metis is_greatest_fixpoint_def omega_induct_mult omega_unfold order_refl) lemma omega_loop_zero_nu: "\ (\x . y ; x) = y\<^sup>\" by (metis greatest_fixpoint_same omega_loop_zero_is_greatest_fixpoint) lemma affine_has_greatest_fixpoint: "has_greatest_fixpoint (\x . y ; x + z)" by (metis has_greatest_fixpoint_def omega_loop_is_greatest_fixpoint) lemma omega_separate_unfold: "(x\<^sup>\ ; y)\<^sup>\ = y\<^sup>\ + y\<^sup>\ ; x ; (x\<^sup>\ ; y)\<^sup>\" by (metis add_commutative mult_associative omega_slide omega_sum_unfold_1 star.circ_loop_fixpoint) lemma omega_zero_left_slide: "(x ; y)\<^sup>\ ; ((x ; y)\<^sup>\ ; 0 + 1) ; x \ x ; (y ; x)\<^sup>\ ; ((y ; x)\<^sup>\ ; 0 + 1)" proof - have "x + x ; (y ; x) ; (y ; x)\<^sup>\ ; ((y ; x)\<^sup>\ ; 0 + 1) \ x ; (y ; x)\<^sup>\ ; ((y ; x)\<^sup>\ ; 0 + 1)" by (smt add_commutative add_least_upper_bound mult_associative mult_left_isotone mult_right_isotone star.circ_back_loop_prefixpoint star.left_plus_below_circ star.mult_zero_add_circ star.mult_zero_circ) hence "((x ; y)\<^sup>\ ; 0 + 1) ; x + x ; y ; (x ; (y ; x)\<^sup>\ ; ((y ; x)\<^sup>\ ; 0 + 1)) \ x ; (y ; x)\<^sup>\ ; ((y ; x)\<^sup>\ ; 0 + 1)" by (smt add_associative less_eq_def mult_associative mult_left_one mult_left_sub_dist_add_left mult_left_zero mult_right_dist_add omega_slide star_mult_omega) thus ?thesis by (metis mult_associative star_left_induct) qed lemma omega_zero_add_1: "(x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1) = x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1)" proof (rule antisym) have 1: "(x + y) ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1) \ x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1)" by (smt add_associative add_commutative less_eq_def mult_associative mult_left_isotone mult_right_dist_add star.circ_add_1 star.left_plus_below_circ star.mult_zero_add_circ star.mult_zero_circ star_decompose_1) have 2: "1 \ x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1)" by (smt add_commutative mult_associative star.circ_add_1 star.circ_reflexive star.mult_zero_add_circ star.mult_zero_circ) have "(y ; x\<^sup>\)\<^sup>\ ; 0 \ (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0" by (smt mult_left_isotone mult_left_sub_dist_add_right mult_right_one omega_isotone) also have 3: "... \ (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1)" by (smt add_commutative mult_associative mult_left_one mult_right_sub_dist_add_left order_trans star.circ_sub_dist_1 star.mult_zero_add_circ star.mult_zero_circ) finally have 4: "(x\<^sup>\ ; y)\<^sup>\ ; 0 \ x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1)" by (smt mult_associative mult_right_isotone omega_slide) have "y ; (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ ; 0 \ y ; (x\<^sup>\ ; (x\<^sup>\ ; 0 + y))\<^sup>\ ; x\<^sup>\ ; x\<^sup>\ ; 0 ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0" by (metis mult_left_isotone mult_left_sub_dist_add_right mult_right_isotone star.circ_isotone mult_associative mult_left_zero star_mult_omega) also have "... \ y ; (x\<^sup>\ ; (x\<^sup>\ ; 0 + y))\<^sup>\ ; (x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; y)\<^sup>\ ; 0" by (smt mult_associative mult_left_isotone mult_left_sub_dist_add_left omega_slide) also have "... = y ; (x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; y)\<^sup>\ ; 0" by (smt mult_associative mult_left_one mult_left_zero mult_right_dist_add star_mult_omega) finally have "x\<^sup>\ ; y ; (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ ; 0 \ x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1)" using 3 by (smt mult_associative mult_right_isotone omega_slide order_trans) hence "(x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ ; 0 \ x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1)" by (smt add_associative add_commutative less_eq_def mult_associative mult_isotone mult_left_one mult_right_one mult_right_sub_dist_add_left order_trans star.circ_loop_fixpoint star.circ_reflexive star.mult_zero_circ) hence "(x + y)\<^sup>\ ; 0 \ x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1)" using 4 by (metis add_least_upper_bound mult_right_dist_add omega_decompose) thus "(x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1) \ x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1)" using 1 2 by (smt add_least_upper_bound mult_associative star_left_induct) next have 5: "x\<^sup>\ ; 0 \ (x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1)" by (metis add_commutative add_left_zero mult_associative mult_left_isotone mult_left_one mult_right_dist_add omega_sub_dist order_trans star_mult_omega zero_right_mult_decreasing) have 6: "(y ; x\<^sup>\)\<^sup>\ ; 0 \ (x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1)" by (metis add_commutative mult_left_isotone omega_sub_dist_1 mult_associative mult_left_sub_dist_add_left order_trans star_mult_omega) have 7: "(y ; x\<^sup>\)\<^sup>\ \ (x + y)\<^sup>\" by (metis mult_left_one mult_right_sub_dist_add_left star.circ_add_1 star.circ_plus_one) hence "(y ; x\<^sup>\)\<^sup>\ ; x\<^sup>\ ; 0 \ (x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1)" by (smt add_associative less_eq_def mult_associative mult_isotone mult_right_dist_add omega_sub_dist) hence "(x\<^sup>\ ; 0 + y ; x\<^sup>\)\<^sup>\ ; 0 \ (x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1)" using 6 by (smt add_commutative add_least_upper_bound mult_associative mult_right_dist_add mult_zero_add_omega omega_unfold omega_zero) hence "(y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 \ y ; x\<^sup>\ ; (x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1)" by (smt mult_associative mult_left_one mult_left_zero mult_right_dist_add mult_right_isotone omega_slide) also have "... \ (x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1)" using 7 by (metis mult_left_isotone order_refl star.circ_mult_upper_bound star_left_induct_mult_iff) finally have "(y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1) \ (x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1)" using 5 by (smt add_commutative add_least_upper_bound mult_associative order_refl star.circ_mult_upper_bound star.circ_reflexive star.circ_sub_dist_1 star.mult_zero_add_circ star.mult_zero_circ star_left_induct) hence "(x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1) \ (x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1)" using 5 by (metis add_commutative mult_associative star.circ_isotone star.circ_mult_upper_bound star.mult_zero_add_circ star.mult_zero_circ star_involutive) thus "x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) ; (y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; ((y ; x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))\<^sup>\ ; 0 + 1) \ (x + y)\<^sup>\ ; ((x + y)\<^sup>\ ; 0 + 1)" by (smt add_associative add_commutative mult_associative star.circ_mult_upper_bound star.circ_sub_dist star.mult_zero_add_circ star.mult_zero_circ) qed lemma star_omega_greatest: "x\<^sup>\\<^sup>\ = 1\<^sup>\" by (metis add_commutative less_eq_def omega_one_greatest omega_sub_dist star.circ_plus_one) lemma omega_vector_greatest: "x\<^sup>\ ; 1\<^sup>\ = x\<^sup>\" by (metis antisym mult_isotone omega_mult_omega_star omega_one_greatest omega_sub_vector) lemma mult_greatest_omega: "(x ; 1\<^sup>\)\<^sup>\ \ x ; 1\<^sup>\" by (metis mult_right_isotone omega_slide omega_sub_vector) lemma omega_mult_star_2: "x\<^sup>\ ; y\<^sup>\ = x\<^sup>\" by (metis mult_associative omega_mult_star omega_vector_greatest star_involutive star_omega_greatest) lemma omega_import: "p \ p ; p \ p ; x \ x ; p \ p ; x\<^sup>\ = p ; (p ; x)\<^sup>\" proof assume 1: "p \ p ; p \ p ; x \ x ; p" hence "p ; x\<^sup>\ \ p ; (p ; x) ; x\<^sup>\" by (metis mult_associative mult_left_isotone omega_unfold) also have "... \ p ; x ; p ; x\<^sup>\" using 1 by (metis mult_associative mult_left_isotone mult_right_isotone) finally have "p ; x\<^sup>\ \ (p ; x)\<^sup>\" by (metis mult_associative omega_induct_mult) hence "p ; x\<^sup>\ \ p ; (p ; x)\<^sup>\" using 1 by (metis mult_associative mult_left_isotone mult_right_isotone order_trans) thus "p ; x\<^sup>\ = p ; (p ; x)\<^sup>\" using 1 by (metis add_left_divisibility antisym mult_right_isotone omega_induct_mult omega_slide omega_sub_dist) qed lemma omega_circ_simulate_right_plus: "z ; x \ y ; (y\<^sup>\ ; 0 + y\<^sup>\) ; z + w \ z ; (x\<^sup>\ ; 0 + x\<^sup>\) \ (y\<^sup>\ ; 0 + y\<^sup>\) ; (z + w ; (x\<^sup>\ ; 0 + x\<^sup>\))" nitpick [expect=genuine] oops lemma omega_circ_simulate_left_plus: "x ; z \ z ; (y\<^sup>\ ; 0 + y\<^sup>\) + w \ (x\<^sup>\ ; 0 + x\<^sup>\) ; z \ (z + (x\<^sup>\ ; 0 + x\<^sup>\) ; w) ; (y\<^sup>\ ; 0 + y\<^sup>\)" nitpick [expect=genuine] oops end -- {* Theorem 50.2 *} sublocale left_omega_algebra < comb0!: left_conway_semiring where circ = "(\x . x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))" apply unfold_locales apply (smt add_associative add_commutative less_eq_def mult_associative mult_left_sub_dist_add_left omega_unfold star.circ_loop_fixpoint star_mult_omega) apply (smt mult_associative omega_zero_left_slide) apply (smt mult_associative omega_zero_add_1) done class left_zero_omega_algebra = left_zero_kleene_algebra + left_omega_algebra begin lemma star_omega_absorb: "y\<^sup>\ ; (y\<^sup>\ ; x)\<^sup>\ ; y\<^sup>\ = (y\<^sup>\ ; x)\<^sup>\ ; y\<^sup>\" proof - have "y\<^sup>\ ; (y\<^sup>\ ; x)\<^sup>\ ; y\<^sup>\ = y\<^sup>\ ; y\<^sup>\ ; x ; (y\<^sup>\ ; x)\<^sup>\ ; y\<^sup>\ + y\<^sup>\ ; y\<^sup>\" by (metis add_commutative mult_associative mult_right_dist_add star.circ_back_loop_fixpoint star.circ_plus_same) thus ?thesis by (metis mult_associative star.circ_loop_fixpoint star.circ_transitive_equal star_mult_omega) qed lemma omega_circ_simulate_right_plus: "z ; x \ y ; (y\<^sup>\ ; 0 + y\<^sup>\) ; z + w \ z ; (x\<^sup>\ ; 0 + x\<^sup>\) \ (y\<^sup>\ ; 0 + y\<^sup>\) ; (z + w ; (x\<^sup>\ ; 0 + x\<^sup>\))" proof assume "z ; x \ y ; (y\<^sup>\ ; 0 + y\<^sup>\) ; z + w" hence 1: "z ; x \ y\<^sup>\ ; 0 + y ; y\<^sup>\ ; z + w" by (metis mult_associative mult_left_dist_add mult_left_zero mult_right_dist_add omega_unfold) hence "(y\<^sup>\ ; 0 + y\<^sup>\ ; z + y\<^sup>\ ; w ; x\<^sup>\ ; 0 + y\<^sup>\ ; w ; x\<^sup>\) ; x \ y\<^sup>\ ; 0 + y\<^sup>\ ; (y\<^sup>\ ; 0 + y ; y\<^sup>\ ; z + w) + y\<^sup>\ ; w ; x\<^sup>\ ; 0 + y\<^sup>\ ; w ; x\<^sup>\" by (smt add_associative add_left_upper_bound add_right_upper_bound less_eq_def mult_associative mult_left_dist_add mult_left_zero mult_right_dist_add star.circ_back_loop_fixpoint) also have "... = y\<^sup>\ ; 0 + y\<^sup>\ ; y ; y\<^sup>\ ; z + y\<^sup>\ ; w ; x\<^sup>\ ; 0 + y\<^sup>\ ; w ; x\<^sup>\" by (smt add_associative add_right_upper_bound less_eq_def mult_associative mult_left_dist_add star.circ_back_loop_fixpoint star_mult_omega) also have "... \ y\<^sup>\ ; 0 + y\<^sup>\ ; z + y\<^sup>\ ; w ; x\<^sup>\ ; 0 + y\<^sup>\ ; w ; x\<^sup>\" by (smt add_commutative add_left_isotone mult_left_isotone star.circ_increasing star.circ_plus_same star.circ_transitive_equal) finally have "z + (y\<^sup>\ ; 0 + y\<^sup>\ ; z + y\<^sup>\ ; w ; x\<^sup>\ ; 0 + y\<^sup>\ ; w ; x\<^sup>\) ; x \ y\<^sup>\ ; 0 + y\<^sup>\ ; z + y\<^sup>\ ; w ; x\<^sup>\ ; 0 + y\<^sup>\ ; w ; x\<^sup>\" by (smt add_least_upper_bound add_left_upper_bound star.circ_loop_fixpoint) hence 2: "z ; x\<^sup>\ \ y\<^sup>\ ; 0 + y\<^sup>\ ; z + y\<^sup>\ ; w ; x\<^sup>\ ; 0 + y\<^sup>\ ; w ; x\<^sup>\" by (metis star_right_induct) have "z ; x\<^sup>\ ; 0 \ (y\<^sup>\ ; 0 + y ; y\<^sup>\ ; z + w) ; x\<^sup>\ ; 0" using 1 by (smt add_left_divisibility mult_associative mult_right_sub_dist_add_left omega_unfold) hence "z ; x\<^sup>\ ; 0 \ y\<^sup>\ + y\<^sup>\ ; (y\<^sup>\ ; 0 + w ; x\<^sup>\ ; 0)" by (smt add_associative add_commutative left_plus_omega mult_associative mult_left_zero mult_right_dist_add omega_induct star.left_plus_circ) thus "z ; (x\<^sup>\ ; 0 + x\<^sup>\) \ (y\<^sup>\ ; 0 + y\<^sup>\) ; (z + w ; (x\<^sup>\ ; 0 + x\<^sup>\))" using 2 by (smt add_associative add_commutative less_eq_def mult_associative mult_left_dist_add mult_left_zero mult_right_dist_add omega_unfold omega_zero star_mult_omega zero_right_mult_decreasing) qed lemma omega_circ_simulate_left_plus: "x ; z \ z ; (y\<^sup>\ ; 0 + y\<^sup>\) + w \ (x\<^sup>\ ; 0 + x\<^sup>\) ; z \ (z + (x\<^sup>\ ; 0 + x\<^sup>\) ; w) ; (y\<^sup>\ ; 0 + y\<^sup>\)" proof assume 1: "x ; z \ z ; (y\<^sup>\ ; 0 + y\<^sup>\) + w" have "x ; (z ; y\<^sup>\ ; 0 + z ; y\<^sup>\ + x\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\) = x ; z ; y\<^sup>\ ; 0 + x ; z ; y\<^sup>\ + x\<^sup>\ ; 0 + x ; x\<^sup>\ ; w ; y\<^sup>\ ; 0 + x ; x\<^sup>\ ; w ; y\<^sup>\" by (smt mult_associative mult_left_dist_add omega_unfold) also have "... \ x ; z ; y\<^sup>\ ; 0 + x ; z ; y\<^sup>\ + x\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\" by (metis add_isotone add_right_isotone mult_left_isotone star.left_plus_below_circ) also have "... \ (z ; y\<^sup>\ ; 0 + z ; y\<^sup>\ + w) ; y\<^sup>\ ; 0 + (z ; y\<^sup>\ ; 0 + z ; y\<^sup>\ + w) ; y\<^sup>\ + x\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\" using 1 by (metis add_left_isotone mult_associative mult_left_dist_add mult_left_isotone) also have "... = z ; y\<^sup>\ ; 0 + z ; y\<^sup>\ ; y\<^sup>\ ; 0 + w ; y\<^sup>\ ; 0 + z ; y\<^sup>\ ; 0 + z ; y\<^sup>\ ; y\<^sup>\ + w ; y\<^sup>\ + x\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\" by (smt add_associative mult_associative mult_left_zero mult_right_dist_add) also have "... = z ; y\<^sup>\ ; 0 + z ; y\<^sup>\ + x\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\" by (smt add_associative add_commutative add_idempotent mult_associative mult_right_dist_add star.circ_loop_fixpoint star.circ_transitive_equal star_mult_omega) finally have "(x\<^sup>\ ; 0 + x\<^sup>\) ; z \ z ; y\<^sup>\ ; 0 + z ; y\<^sup>\ + x\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\ ; 0 + x\<^sup>\ ; w ; y\<^sup>\" by (smt add_least_upper_bound add_left_upper_bound mult_associative mult_left_zero mult_right_dist_add star.circ_back_loop_fixpoint star_left_induct) thus "(x\<^sup>\ ; 0 + x\<^sup>\) ; z \ (z + (x\<^sup>\ ; 0 + x\<^sup>\) ; w) ; (y\<^sup>\ ; 0 + y\<^sup>\)" by (smt add_associative mult_associative mult_left_dist_add mult_left_zero mult_right_dist_add) qed lemma omega_translate: "x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) = x\<^sup>\ ; 0 + x\<^sup>\" by (metis mult_associative mult_left_dist_add mult_right_one star_mult_omega) lemma omega_circ_simulate_right: "z ; x \ y ; z + w \ z ; (x\<^sup>\ ; 0 + x\<^sup>\) \ (y\<^sup>\ ; 0 + y\<^sup>\) ; (z + w ; (x\<^sup>\ ; 0 + x\<^sup>\))" proof assume "z ; x \ y ; z + w" also have "... \ y ; (y\<^sup>\ ; 0 + y\<^sup>\) ; z + w" by (metis add_left_isotone comb0.circ_reflexive mult_left_isotone mult_right_isotone mult_right_one omega_translate) finally show "z ; (x\<^sup>\ ; 0 + x\<^sup>\) \ (y\<^sup>\ ; 0 + y\<^sup>\) ; (z + w ; (x\<^sup>\ ; 0 + x\<^sup>\))" by (metis omega_circ_simulate_right_plus) qed end sublocale left_zero_omega_algebra < comb1!: left_conway_semiring_1 where circ = "(\x . x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))" apply unfold_locales apply (smt eq_iff mult_associative mult_left_dist_add mult_left_zero mult_right_dist_add mult_right_one omega_slide star_slide) done sublocale left_zero_omega_algebra < comb0!: itering where circ = "(\x . x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))" apply unfold_locales apply (metis comb1.circ_add_9) apply (metis comb1.circ_mult_1) apply (metis omega_circ_simulate_right_plus omega_translate) apply (metis omega_circ_simulate_left_plus omega_translate) done -- {* Theorem 2.2 *} sublocale left_zero_omega_algebra < comb2!: itering where circ = "(\x . x\<^sup>\ ; 0 + x\<^sup>\)" apply unfold_locales apply (metis comb1.circ_add_9 omega_translate) apply (metis comb1.circ_mult_1 omega_translate) apply (metis omega_circ_simulate_right_plus) apply (metis omega_circ_simulate_left_plus) done class omega_algebra = kleene_algebra + left_zero_omega_algebra class left_omega_conway_semiring = left_omega_algebra + left_conway_semiring begin subclass left_kleene_conway_semiring .. lemma circ_below_omega_star: "x\<^sup>\ \ x\<^sup>\ + x\<^sup>\" by (metis circ_left_unfold mult_right_one omega_induct order_refl) lemma omega_mult_circ: "x\<^sup>\ ; x\<^sup>\ = x\<^sup>\" by (metis circ_star mult_associative omega_mult_star omega_vector_greatest star_omega_greatest) lemma circ_mult_omega: "x\<^sup>\ ; x\<^sup>\ = x\<^sup>\" by (metis antisym add_right_divisibility circ_loop_fixpoint circ_plus_sub omega_simulation) lemma circ_omega_greatest: "x\<^sup>\\<^sup>\ = 1\<^sup>\" by (metis circ_star star_omega_greatest) lemma omega_circ: "x\<^sup>\\<^sup>\ = 1 + x\<^sup>\" by (metis antisym circ_left_unfold mult_left_sub_dist_add_left mult_right_one omega_sub_vector) end class bounded_left_omega_algebra = bounded_left_kleene_algebra + left_omega_algebra begin lemma omega_one: "1\<^sup>\ = T" by (smt add_left_top less_eq_def omega_one_greatest) lemma star_omega_top: "x\<^sup>\\<^sup>\ = T" by (metis add_left_top less_eq_def omega_one omega_sub_dist star.circ_plus_one) lemma omega_vector: "x\<^sup>\ ; T = x\<^sup>\" by (metis add_commutative less_eq_def omega_sub_vector top_right_mult_increasing) lemma mult_top_omega: "(x ; T)\<^sup>\ \ x ; T" by (metis mult_right_isotone omega_slide top_greatest) end sublocale bounded_left_omega_algebra < comb0!: bounded_left_conway_semiring where circ = "(\x . x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))" .. class bounded_left_zero_omega_algebra = bounded_left_zero_kleene_algebra + left_zero_omega_algebra begin subclass bounded_left_omega_algebra .. end sublocale bounded_left_zero_omega_algebra < comb0!: bounded_itering where circ = "(\x . x\<^sup>\ ; (x\<^sup>\ ; 0 + 1))" .. class bounded_omega_algebra = bounded_kleene_algebra + omega_algebra begin subclass bounded_left_zero_omega_algebra .. end class bounded_left_omega_conway_semiring = bounded_left_omega_algebra + left_omega_conway_semiring begin subclass left_kleene_conway_semiring .. subclass bounded_left_conway_semiring .. lemma circ_omega: "x\<^sup>\\<^sup>\ = T" by (metis circ_star star_omega_top) end class top_left_omega_algebra = bounded_left_omega_algebra + assumes top_left_zero: "T ; x = T" begin lemma omega_translate_3: "x\<^sup>\ ; (x\<^sup>\ ; 0 + 1) = x\<^sup>\ ; (x\<^sup>\ + 1)" by (metis mult_associative omega_mult_star_2 star.circ_top_1 top_left_zero) end -- {* Theorem 50.2 *} sublocale top_left_omega_algebra < comb4!: left_conway_semiring where circ = "(\x . x\<^sup>\ ; (x\<^sup>\ + 1))" apply unfold_locales apply (metis comb0.circ_left_unfold omega_translate_3) apply (metis comb0.circ_left_slide omega_translate_3) apply (metis comb0.circ_add_1 omega_translate_3) done class top_left_zero_omega_algebra = bounded_left_zero_omega_algebra + assumes top_left_zero: "T ; x = T" begin lemma omega_translate_2: "x\<^sup>\ ; 0 + x\<^sup>\ = x\<^sup>\ + x\<^sup>\" by (metis mult_associative omega_mult_star_2 star.circ_top top_left_zero) end -- {* Theorem 2.3 *} sublocale top_left_zero_omega_algebra < comb3!: itering where circ = "(\x . x\<^sup>\ + x\<^sup>\)" apply unfold_locales apply (metis comb2.circ_add_9 omega_translate_2) apply (metis comb2.circ_mult_1 omega_translate_2) apply (metis omega_circ_simulate_right_plus omega_translate_2) apply (metis omega_circ_simulate_left_plus omega_translate_2) done class Omega = fixes Omega :: "'a \ 'a" ("_\<^sup>\" [100] 100) (* end theory BinaryItering imports Semiring begin *) class binary_itering = idempotent_left_zero_semiring + while + assumes while_productstar: "(x ; y) \ z = z + x ; ((y ; x) \ (y ; z))" assumes while_sumstar: "(x + y) \ z = (x \ y) \ (x \ z)" assumes while_left_dist_add: "x \ (y + z) = (x \ y) + (x \ z)" assumes while_sub_associative: "(x \ y) ; z \ x \ (y ; z)" assumes while_simulate_left_plus: "x ; z \ z ; (y \ 1) + w \ x \ (z ; v) \ z ; (y \ v) + (x \ (w ; (y \ v)))" assumes while_simulate_right_plus: "z ; x \ y ; (y \ z) + w \ z ; (x \ v) \ y \ (z ; v + w ; (x \ v))" begin -- {* Theorem 9.1 *} lemma while_zero: "0 \ x = x" by (metis add_right_zero mult_left_zero while_productstar) -- {* Theorem 9.4 *} lemma while_mult_increasing: "x ; y \ x \ y" by (metis add_least_upper_bound mult_left_one order_refl while_productstar) -- {* Theorem 9.2 *} lemma while_one_increasing: "x \ x \ 1" by (metis mult_right_one while_mult_increasing) -- {* Theorem 9.3 *} lemma while_increasing: "y \ x \ y" by (metis add_left_divisibility mult_left_one while_productstar) -- {* Theorem 9.42 *} lemma while_right_isotone: "y \ z \ x \ y \ x \ z" by (metis less_eq_def while_left_dist_add) -- {* Theorem 9.41 *} lemma while_left_isotone: "x \ y \ x \ z \ y \ z" by (metis less_eq_def while_increasing while_sumstar) lemma while_isotone: "w \ x \ y \ z \ w \ y \ x \ z" by (smt order_trans while_left_isotone while_right_isotone) -- {* Theorem 9.17 *} lemma while_left_unfold: "x \ y = y + x ; (x \ y)" by (metis mult_left_one mult_right_one while_productstar) lemma while_simulate_left_plus_1: "x ; z \ z ; (y \ 1) \ x \ (z ; w) \ z ; (y \ w) + (x \ 0)" by (metis add_right_zero mult_left_zero while_simulate_left_plus) -- {* Theorem 11.1 *} lemma while_simulate_absorb: "y ; x \ x \ y \ x \ x + (y \ 0)" by (metis while_simulate_left_plus_1 while_zero mult_right_one) -- {* Theorem 9.10 *} lemma while_transitive: "x \ (x \ y) = x \ y" by (metis add_right_upper_bound add_right_zero antisym while_increasing while_left_dist_add while_left_unfold while_simulate_absorb) -- {* Theorem 9.25 *} lemma while_slide: "(x ; y) \ (x ; z) = x ; ((y ; x) \ z)" by (metis mult_associative mult_left_dist_add while_left_unfold while_productstar) -- {* Theorem 9.21 *} lemma while_zero_2: "(x ; 0) \ y = x ; 0 + y" by (metis add_commutative mult_associative mult_left_zero while_left_unfold) -- {* Theorem 9.5 *} lemma while_mult_star_exchange: "x ; (x \ y) = x \ (x ; y)" by (metis mult_left_one while_slide) -- {* Theorem 9.18 *} lemma while_right_unfold: "x \ y = y + (x \ (x ; y))" by (metis while_left_unfold while_mult_star_exchange) -- {* Theorem 9.7 *} lemma while_one_mult_below: "(x \ 1) ; y \ x \ y" by (metis mult_left_one while_sub_associative) lemma while_plus_one: "x \ y = y + (x \ y)" by (metis less_eq_def while_increasing) -- {* Theorem 9.19 *} lemma while_rtc_2: "y + x ; y + (x \ (x \ y)) = x \ y" by (metis add_associative less_eq_def while_mult_increasing while_plus_one while_transitive) -- {* Theorem 9.6 *} lemma while_left_plus_below: "x ; (x \ y) \ x \ y" by (metis add_right_divisibility while_left_unfold) lemma while_right_plus_below: "x \ (x ; y) \ x \ y" by (metis while_left_plus_below while_mult_star_exchange) lemma while_right_plus_below_2: "(x \ x) ; y \ x \ y" by (smt order_trans while_right_plus_below while_sub_associative) -- {* Theorem 9.47 *} lemma while_mult_transitive: "x \ z \ y \ y \ z \ w \ x \ z \ w" by (smt order_trans while_right_isotone while_transitive) -- {* Theorem 9.48 *} lemma while_mult_upper_bound: "x \ z \ 1 \ y \ z \ w \ x ; y \ z \ w" by (metis less_eq_def mult_right_sub_dist_add_left order_trans while_mult_transitive while_one_mult_below) lemma while_one_mult_while_below: "(y \ 1) ; (y \ v) \ y \ v" by (metis order_refl while_mult_upper_bound) -- {* Theorem 9.34 *} lemma while_sub_dist: "x \ z \ (x + y) \ z" by (metis add_left_upper_bound while_left_isotone) lemma while_sub_dist_1: "x ; z \ (x + y) \ z" by (metis order_trans while_mult_increasing while_sub_dist) lemma while_sub_dist_2: "x ; y ; z \ (x + y) \ z" by (metis add_commutative mult_associative while_mult_transitive while_sub_dist_1) -- {* Theorem 9.36 *} lemma while_sub_dist_3: "x \ (y \ z) \ (x + y) \ z" by (metis add_right_upper_bound while_left_isotone while_mult_transitive while_sub_dist) -- {* Theorem 9.44 *} lemma while_absorb_2: "x \ y \ y \ (x \ z) = y \ z" by (metis add_commutative less_eq_def while_left_dist_add while_plus_one while_sub_dist_3) lemma while_simulate_right_plus_1: "z ; x \ y ; (y \ z) \ z ; (x \ w) \ y \ (z ; w)" by (metis add_right_zero mult_left_zero while_simulate_right_plus) -- {* Theorem 9.39 *} lemma while_sumstar_1_below: "x \ ((y ; (x \ 1)) \ z) \ ((x \ 1) ; y) \ (x \ z)" proof - have 1: "x ; (((x \ 1) ; y) \ (x \ z)) \ ((x \ 1) ; y) \ (x \ z)" by (smt add_isotone add_right_upper_bound mult_associative mult_left_dist_add mult_right_sub_dist_add_right while_left_unfold) have "x \ ((y ; (x \ 1)) \ z) \ (x \ z) + (x \ (y ; (((x \ 1) ; y) \ ((x \ 1) ; z))))" by (metis eq_refl while_left_dist_add while_productstar) also have "... \ (x \ z) + (x \ ((x \ 1) ; y ; (((x \ 1) ; y) \ ((x \ 1) ; z))))" by (metis add_right_isotone mult_associative mult_left_one mult_right_sub_dist_add_left while_left_unfold while_right_isotone) also have "... \ (x \ z) + (x \ (((x \ 1) ; y) \ ((x \ 1) ; z)))" by (metis add_right_isotone add_right_upper_bound while_left_unfold while_right_isotone) also have "... \ x \ (((x \ 1) ; y) \ (x \ z))" by (smt add_associative add_left_upper_bound less_eq_def mult_left_one while_left_dist_add while_left_unfold while_sub_associative) also have "... \ (((x \ 1) ; y) \ (x \ z)) + (x \ 0)" using 1 by (metis while_simulate_absorb) also have "... = ((x \ 1) ; y) \ (x \ z)" by (smt add_associative add_commutative add_left_zero while_left_dist_add while_left_unfold) finally show ?thesis . qed lemma while_sumstar_2_below: "((x \ 1) ; y) \ (x \ z) \ (x \ y) \ (x \ z)" by (metis mult_left_one while_left_isotone while_sub_associative) -- {* Theorem 9.38 *} lemma while_add_1_below: "x \ ((y ; (x \ 1)) \ z) \ (x + y) \ z" proof - have "((x \ 1) ; y) \ ((x \ 1) ; z) \ (x + y) \ z" by (metis while_isotone while_one_mult_below while_sumstar) hence "(y ; (x \ 1)) \ z \ z + y ; ((x + y) \ z)" by (metis add_right_isotone mult_right_isotone while_productstar) also have "... \ (x + y) \ z" by (metis add_right_isotone add_right_upper_bound mult_left_isotone while_left_unfold) finally show ?thesis by (metis add_commutative add_right_upper_bound while_isotone while_transitive) qed -- {* Theorem 9.16 *} lemma while_while_while: "((x \ 1) \ 1) \ y = (x \ 1) \ y" by (smt add_commutative less_eq_def mult_right_one while_left_plus_below while_mult_star_exchange while_plus_one while_sumstar while_transitive) lemma while_one: "(1 \ 1) \ y = 1 \ y" by (metis while_while_while while_zero) -- {* Theorem 9.22 *} lemma while_add_below: "x + y \ x \ (y \ 1)" by (smt add_commutative add_isotone case_split_right order_trans while_increasing while_left_plus_below while_mult_increasing while_plus_one) -- {* Theorem 9.32 *} lemma while_add_2: "(x + y) \ z \ (x \ (y \ 1)) \ z" by (metis while_add_below while_left_isotone) -- {* Theorem 9.45 *} lemma while_sup_one_left_unfold: "1 \ x \ x ; (x \ y) = x \ y" by (metis less_eq_def mult_left_one mult_right_dist_add while_mult_star_exchange while_right_unfold while_transitive) lemma while_sup_one_right_unfold: "1 \ x \ x \ (x ; y) = x \ y" by (metis while_mult_star_exchange while_sup_one_left_unfold) -- {* Theorem 9.30 *} lemma while_decompose_7: "(x + y) \ z = x \ (y \ ((x + y) \ z))" by (metis eq_iff order_trans while_increasing while_sub_dist_3 while_transitive) -- {* Theorem 9.31 *} lemma while_decompose_8: "(x + y) \ z = (x + y) \ (x \ (y \ z))" by (metis add_commutative while_sumstar while_transitive) -- {* Theorem 9.27 *} lemma while_decompose_9: "(x \ (y \ 1)) \ z = x \ (y \ ((x \ (y \ 1)) \ z))" by (smt add_commutative less_eq_def order_trans while_add_below while_increasing while_sub_dist_3) lemma while_decompose_10: "(x \ (y \ 1)) \ z = (x \ (y \ 1)) \ (x \ (y \ z))" proof - have 1: "(x \ (y \ 1)) \ z \ (x \ (y \ 1)) \ (x \ (y \ z))" by (metis add_associative less_eq_def while_left_dist_add while_plus_one) have "x + (y \ 1) \ x \ (y \ 1)" by (metis add_least_upper_bound while_add_below while_increasing) hence "(x \ (y \ 1)) \ (x \ (y \ z)) \ (x \ (y \ 1)) \ z" by (smt add_least_upper_bound eq_refl order_trans while_absorb_2 while_one_increasing) thus ?thesis using 1 by (metis antisym) qed lemma while_back_loop_fixpoint: "z ; (y \ (y ; x)) + z ; x = z ; (y \ x)" by (metis add_commutative mult_left_dist_add while_right_unfold) lemma while_back_loop_prefixpoint: "z ; (y \ 1) ; y + z \ z ; (y \ 1)" by (metis add_least_upper_bound mult_associative mult_right_isotone mult_right_one order_refl while_increasing while_mult_upper_bound while_one_increasing) -- {* Theorem 9 *} lemma while_loop_is_fixpoint: "is_fixpoint (\x . y ; x + z) (y \ z)" by (smt add_commutative is_fixpoint_def while_left_unfold) -- {* Theorem 9 *} lemma while_back_loop_is_prefixpoint: "is_prefixpoint (\x . x ; y + z) (z ; (y \ 1))" by (metis is_prefixpoint_def while_back_loop_prefixpoint) -- {* Theorem 9.20 *} lemma while_while_add: "(1 + x) \ y = (x \ 1) \ y" by (metis add_commutative while_decompose_10 while_sumstar while_zero) lemma while_while_mult_sub: "x \ (1 \ y) \ (x \ 1) \ y" by (metis add_commutative while_sub_dist_3 while_while_add) -- {* Theorem 9.11 *} lemma while_right_plus: "(x \ x) \ y = x \ y" by (metis add_idempotent while_plus_one while_sumstar while_transitive) -- {* Theorem 9.12 *} lemma while_left_plus: "(x ; (x \ 1)) \ y = x \ y" by (metis mult_right_one while_mult_star_exchange while_right_plus) -- {* Theorem 9.9 *} lemma while_below_while_one: "x \ x \ x \ 1" by (metis while_one_increasing while_right_plus) lemma while_below_while_one_mult: "x ; (x \ x) \ x ; (x \ 1)" by (metis mult_right_isotone while_below_while_one) -- {* Theorem 9.23 *} lemma while_add_sub_add_one: "x \ (x + y) \ x \ (1 + y)" by (metis add_left_isotone while_below_while_one while_left_dist_add) lemma while_add_sub_add_one_mult: "x ; (x \ (x + y)) \ x ; (x \ (1 + y))" by (metis mult_right_isotone while_add_sub_add_one) lemma while_elimination: "x ; y = 0 \ x ; (y \ z) = x ; z" by (metis add_right_zero mult_associative mult_left_dist_add mult_left_zero while_left_unfold) -- {* Theorem 9.8 *} lemma while_square: "(x ; x) \ y \ x \ y" by (metis while_left_isotone while_mult_increasing while_right_plus) -- {* Theorem 9.35 *} lemma while_mult_sub_add: "(x ; y) \ z \ (x + y) \ z" by (metis while_increasing while_isotone while_mult_increasing while_sumstar) -- {* Theorem 9.43 *} lemma while_absorb_1: "x \ y \ x \ (y \ z) = y \ z" by (metis antisym less_eq_def while_increasing while_sub_dist_3) lemma while_absorb_3: "x \ y \ x \ (y \ z) = y \ (x \ z)" by (metis while_absorb_1 while_absorb_2) -- {* Theorem 9.24 *} lemma while_square_2: "(x ; x) \ ((x + 1) ; y) \ x \ y" by (metis add_least_upper_bound while_increasing while_mult_transitive while_mult_upper_bound while_one_increasing while_square) lemma while_separate_unfold_below: "(y ; (x \ 1)) \ z \ (y \ z) + (y \ (y ; x ; (x \ ((y ; (x \ 1)) \ z))))" proof - have "(y ; (x \ 1)) \ z = (y \ (y ; x ; (x \ 1))) \ (y \ z)" by (metis mult_associative mult_left_dist_add mult_right_one while_left_unfold while_sumstar) hence "(y ; (x \ 1)) \ z = (y \ z) + (y \ (y ; x ; (x \ 1))) ; ((y ; (x \ 1)) \ z)" by (metis while_left_unfold) also have "... \ (y \ z) + (y \ (y ; x ; (x \ 1)) ; ((y ; (x \ 1)) \ z))" by (metis add_right_isotone while_sub_associative) also have "... \ (y \ z) + (y \ (y ; x ; (x \ ((y ; (x \ 1)) \ z))))" by (smt add_right_isotone mult_associative mult_right_isotone while_one_mult_below while_right_isotone) finally show ?thesis . qed -- {* Theorem 9.33 *} lemma while_mult_zero_add: "(x + y ; 0) \ z = x \ ((y ; 0) \ z)" proof - have "(x + y ; 0) \ z = (x \ (y ; 0)) \ (x \ z)" by (metis while_sumstar) also have "... = (x \ z) + (x \ (y ; 0)) ; ((x \ (y ; 0)) \ (x \ z))" by (metis while_left_unfold) also have "... \ (x \ z) + (x \ (y ; 0))" by (metis add_right_isotone mult_associative mult_left_zero while_sub_associative) also have "... = x \ ((y ; 0) \ z)" by (metis add_commutative while_left_dist_add while_zero_2) finally show ?thesis by (metis le_neq_trans less_def while_sub_dist_3) qed lemma while_add_mult_zero: "(x + y ; 0) \ y = x \ y" by (metis less_eq_def while_mult_zero_add while_zero_2 zero_right_mult_decreasing) lemma while_mult_zero_add_2: "(x + y ; 0) \ z = (x \ z) + (x \ (y ; 0))" by (metis add_commutative while_left_dist_add while_mult_zero_add while_zero_2) lemma while_add_zero_star: "(x + y ; 0) \ z = x \ (y ; 0 + z)" by (metis while_mult_zero_add while_zero_2) lemma while_unfold_sum: "(x + y) \ z = (x \ z) + (x \ (y ; ((x + y) \ z)))" apply (rule antisym) apply (smt add_associative less_eq_def while_absorb_1 while_increasing while_mult_star_exchange while_right_unfold while_sub_associative while_sumstar) apply (metis add_least_upper_bound while_decompose_7 while_mult_increasing while_right_isotone while_sub_dist) done lemma while_simulate_left: "x ; z \ z ; y + w \ x \ (z ; v) \ z ; (y \ v) + (x \ (w ; (y \ v)))" by (metis add_left_isotone mult_right_isotone order_trans while_one_increasing while_simulate_left_plus) lemma while_simulate_right: "z ; x \ y ; z + w \ z ; (x \ v) \ y \ (z ; v + w ; (x \ v))" proof - have "y ; z + w \ y ; (y \ z) + w" by (metis add_left_isotone mult_right_isotone while_increasing) thus ?thesis by (smt order_trans while_simulate_right_plus) qed lemma while_simulate: "z ; x \ y ; z \ z ; (x \ v) \ y \ (z ; v)" by (metis add_right_zero mult_left_zero while_simulate_right) -- {* Theorem 9.14 *} lemma while_while_mult: "1 \ (x \ y) = (x \ 1) \ y" proof - have "(x \ 1) \ y \ (x \ 1) ; ((x \ 1) \ y)" by (metis order_refl while_increasing while_sup_one_left_unfold) also have "... \ 1 \ ((x \ 1) ; y)" by (metis mult_left_one order_refl while_mult_upper_bound while_simulate) also have "... \ 1 \ (x \ y)" by (metis while_one_mult_below while_right_isotone) finally show ?thesis by (metis antisym while_sub_dist_3 while_while_add) qed lemma while_simulate_left_1: "x ; z \ z ; y \ x \ (z ; v) \ z ; (y \ v) + (x \ 0)" by (metis add_right_zero mult_left_zero while_simulate_left) -- {* Theorem 9.46 *} lemma while_associative_1: "1 \ z \ x \ (y ; z) = (x \ y) ; z" proof assume 1: "1 \ z" have "x \ (y ; z) \ x \ ((x \ y) ; z)" by (metis less_eq_def mult_right_dist_add while_plus_one while_right_isotone) also have "... \ (x \ y) ; (0 \ z) + (x \ 0)" by (metis mult_associative mult_right_sub_dist_add_right while_left_unfold while_simulate_absorb while_zero) also have "... \ (x \ y) ; z + (x \ 0) ; z" using 1 by (metis add_least_upper_bound add_left_upper_bound add_right_upper_bound case_split_right while_plus_one while_zero) also have "... = (x \ y) ; z" by (metis add_right_zero mult_right_dist_add while_left_dist_add) finally show "x \ (y ; z) = (x \ y) ; z" by (metis antisym while_sub_associative) qed -- {* Theorem 9.29 *} lemma while_associative_while_1: "x \ (y ; (z \ 1)) = (x \ y) ; (z \ 1)" by (metis while_associative_1 while_increasing) -- {* Theorem 9.13 *} lemma while_one_while: "(x \ 1) ; (y \ 1) = x \ (y \ 1)" by (metis mult_left_one while_associative_while_1) lemma while_decompose_5_below: "(x \ (y \ 1)) \ z \ (y \ (x \ 1)) \ z" by (smt add_commutative mult_left_dist_add mult_right_one while_increasing while_left_unfold while_mult_star_exchange while_one_while while_plus_one while_sumstar) -- {* Theorem 9.26 *} lemma while_decompose_5: "(x \ (y \ 1)) \ z = (y \ (x \ 1)) \ z" by (metis antisym while_decompose_5_below) lemma while_decompose_4: "(x \ (y \ 1)) \ z = x \ ((y \ (x \ 1)) \ z)" by (metis while_decompose_5 while_decompose_9 while_transitive) -- {* Theorem 11.7 *} lemma while_simulate_2: "y ; (x \ 1) \ x \ (y \ 1) \ y \ (x \ 1) \ x \ (y \ 1)" proof (rule iffI) assume "y ; (x \ 1) \ x \ (y \ 1)" hence "y ; (x \ 1) \ (x \ 1) ; (y \ 1)" by (metis while_one_while) hence "y \ ((x \ 1) ; 1) \ (x \ 1) ; (y \ 1) + (y \ 0)" by (metis while_simulate_left_plus_1) hence "y \ (x \ 1) \ (x \ (y \ 1)) + (y \ 0)" by (metis mult_right_one while_one_while) also have "... = x \ (y \ 1)" by (metis add_commutative less_eq_def order_trans while_increasing while_right_isotone zero_least) finally show "y \ (x \ 1) \ x \ (y \ 1)" . next assume "y \ (x \ 1) \ x \ (y \ 1)" thus "y ; (x \ 1) \ x \ (y \ 1)" by (metis order_trans while_mult_increasing) qed lemma while_simulate_1: "y ; x \ x ; y \ y \ (x \ 1) \ x \ (y \ 1)" by (metis order_trans while_mult_increasing while_right_isotone while_simulate while_simulate_2) lemma while_simulate_3: "y ; (x \ 1) \ x \ 1 \ y \ (x \ 1) \ x \ (y \ 1)" by (metis add_idempotent case_split_right while_increasing while_mult_upper_bound while_simulate_2) -- {* Theorem 9.28 *} lemma while_extra_while: "(y ; (x \ 1)) \ z = (y ; (y \ (x \ 1))) \ z" proof - have "y ; (y \ (x \ 1)) \ y ; (x \ 1) ; (y ; (x \ 1) \ 1)" by (smt add_commutative add_left_upper_bound mult_right_one order_trans while_back_loop_prefixpoint while_left_isotone while_mult_star_exchange) hence 1: "(y ; (y \ (x \ 1))) \ z \ (y ; (x \ 1)) \ z" by (metis while_simulate_right_plus_1 mult_left_one) have "(y ; (x \ 1)) \ z \ (y ; (y \ (x \ 1))) \ z" by (metis while_increasing while_left_isotone while_mult_star_exchange) thus ?thesis using 1 by (metis antisym) qed -- {* Theorem 11.6 *} lemma while_separate_4: "y ; x \ x ; (x \ (1 + y)) \ (x + y) \ z = x \ (y \ z)" proof assume 1: "y ; x \ x ; (x \ (1 + y))" hence "(1 + y) ; x \ x ; (x \ (1 + y))" by (smt add_associative add_least_upper_bound mult_left_one mult_left_sub_dist_add_left mult_right_dist_add mult_right_one while_left_unfold) hence 2: "(1 + y) ; (x \ 1) \ x \ (1 + y)" by (metis mult_right_one while_simulate_right_plus_1) have "y ; x ; (x \ 1) \ x ; (x \ ((1 + y) ; (x \ 1)))" using 1 by (smt less_eq_def mult_associative mult_right_dist_add while_associative_1 while_increasing) also have "... \ x ; (x \ (1 + y))" using 2 by (metis mult_right_isotone order_refl while_mult_transitive) also have "... \ x ; (x \ 1) ; (y \ 1)" by (metis add_least_upper_bound mult_associative mult_right_isotone while_increasing while_one_increasing while_one_while while_right_isotone) finally have "y \ (x ; (x \ 1)) \ x ; (x \ 1) ; (y \ 1) + (y \ 0)" by (metis mult_associative mult_right_one while_simulate_left_plus_1) hence "(y \ 1) ; (y \ x) \ x ; (x \ y \ 1) + (y \ 0)" by (smt less_eq_def mult_associative mult_right_one order_refl order_trans while_absorb_2 while_left_dist_add while_mult_star_exchange while_one_mult_below while_one_while while_plus_one) hence "(y \ 1) ; ((y \ x) \ (y \ z)) \ x \ ((y \ 1) ; (y \ z) + (y \ 0) ; ((y \ x) \ (y \ z)))" by (metis while_simulate_right_plus) also have "... \ x \ ((y \ z) + (y \ 0))" by (metis add_isotone mult_left_zero order_refl while_absorb_2 while_one_mult_below while_right_isotone while_sub_associative) also have "... = x \ y \ z" by (metis add_right_zero while_left_dist_add) finally show "(x + y) \ z = x \ (y \ z)" by (smt add_commutative less_eq_def mult_left_one mult_right_dist_add while_plus_one while_sub_associative while_sumstar) qed lemma while_separate_5: "y ; x \ x ; (x \ (x + y)) \ (x + y) \ z = x \ (y \ z)" by (smt order_trans while_add_sub_add_one_mult while_separate_4) lemma while_separate_6: "y ; x \ x ; (x + y) \ (x + y) \ z = x \ (y \ z)" by (smt order_trans while_increasing while_mult_star_exchange while_separate_5) -- {* Theorem 11.4 *} lemma while_separate_1: "y ; x \ x ; y \ (x + y) \ z = x \ (y \ z)" by (metis add_least_upper_bound less_eq_def mult_left_sub_dist_add_right while_separate_6) -- {* Theorem 11.2 *} lemma while_separate_mult_1: "y ; x \ x ; y \ (x ; y) \ z \ x \ (y \ z)" by (metis while_mult_sub_add while_separate_1) -- {* Theorem 11.5 *} lemma separation: "y ; x \ x ; (y \ 1) \ (x + y) \ z = x \ (y \ z)" proof assume "y ; x \ x ; (y \ 1)" hence "y \ x \ x ; (y \ 1) + (y \ 0)" by (metis mult_right_one while_simulate_left_plus_1) also have "... \ x ; (x \ y \ 1) + (y \ 0)" by (metis add_left_isotone while_increasing while_mult_star_exchange) finally have "(y \ 1) ; (y \ x) \ x ; (x \ y \ 1) + (y \ 0)" by (metis order_refl order_trans while_absorb_2 while_one_mult_below) hence "(y \ 1) ; ((y \ x) \ (y \ z)) \ x \ ((y \ 1) ; (y \ z) + (y \ 0) ; ((y \ x) \ (y \ z)))" by (metis while_simulate_right_plus) also have "... \ x \ ((y \ z) + (y \ 0))" by (metis add_isotone mult_left_zero order_refl while_absorb_2 while_one_mult_below while_right_isotone while_sub_associative) also have "... = x \ y \ z" by (metis add_right_zero while_left_dist_add) finally show "(x + y) \ z = x \ (y \ z)" by (smt add_commutative less_eq_def mult_left_one mult_right_dist_add while_plus_one while_sub_associative while_sumstar) qed -- {* Theorem 11.5 *} lemma while_separate_left: "y ; x \ x ; (y \ 1) \ y \ (x \ z) \ x \ (y \ z)" by (metis add_commutative separation while_sub_dist_3) -- {* Theorem 11.6 *} lemma while_simulate_4: "y ; x \ x ; (x \ (1 + y)) \ y \ (x \ z) \ x \ (y \ z)" by (metis add_commutative while_separate_4 while_sub_dist_3) lemma while_simulate_5: "y ; x \ x ; (x \ (x + y)) \ y \ (x \ z) \ x \ (y \ z)" by (smt order_trans while_add_sub_add_one_mult while_simulate_4) lemma while_simulate_6: "y ; x \ x ; (x + y) \ y \ (x \ z) \ x \ (y \ z)" by (smt order_trans while_increasing while_mult_star_exchange while_simulate_5) -- {* Theorem 11.3 *} lemma while_simulate_7: "y ; x \ x ; y \ y \ (x \ z) \ x \ (y \ z)" by (metis add_commutative mult_left_sub_dist_add_left order_trans while_simulate_6) lemma while_while_mult_1: "x \ (1 \ y) = 1 \ (x \ y)" by (metis add_commutative mult_left_one mult_right_one order_refl while_separate_1) -- {* Theorem 9.15 *} lemma while_while_mult_2: "x \ (1 \ y) = (x \ 1) \ y" by (metis while_while_mult while_while_mult_1) -- {* Theorem 11.8 *} lemma while_import: "p \ p ; p \ p \ 1 \ p ; x \ x ; p \ p ; (x \ y) = p ; ((p ; x) \ y)" proof assume 1: "p \ p ; p \ p \ 1 \ p ; x \ x ; p" hence "p ; (x \ y) \ (p ; x) \ (p ; y)" by (smt add_commutative less_eq_def mult_associative mult_left_dist_add mult_right_one while_simulate) also have "... \ (p ; x) \ y" using 1 by (metis less_eq_def mult_left_one mult_right_dist_add while_right_isotone) finally have 2: "p ; (x \ y) \ p ; ((p ; x) \ y)" using 1 by (smt add_commutative less_eq_def mult_associative mult_left_dist_add mult_right_one) have "p ; ((p ; x) \ y) \ p ; (x \ y)" using 1 by (metis mult_left_isotone mult_left_one mult_right_isotone while_left_isotone) thus "p ; (x \ y) = p ; ((p ; x) \ y)" using 2 by (metis antisym) qed -- {* Theorem 11.8 *} lemma while_preserve: "p \ p ; p \ p \ 1 \ p ; x \ x ; p \ p ; (x \ y) = p ; (x \ (p ; y))" apply rule apply (rule antisym) apply (metis mult_associative mult_left_isotone mult_right_isotone order_trans while_simulate) apply (metis mult_left_isotone mult_left_one mult_right_isotone while_right_isotone) done lemma while_plus_below_while: "(x \ 1) ; x \ x \ 1" by (metis order_refl while_mult_upper_bound while_one_increasing) -- {* Theorem 9.40 *} lemma while_01: "(w ; (x \ 1)) \ (y ; z) \ (x \ w) \ ((x \ y) ; z)" proof - have "(w ; (x \ 1)) \ (y ; z) = y ; z + w ; (((x \ 1) ; w) \ ((x \ 1) ; y ; z))" by (metis mult_associative while_productstar) also have "... \ y ; z + w ; ((x \ w) \ ((x \ y) ; z))" by (metis add_right_isotone mult_left_isotone mult_right_isotone while_isotone while_one_mult_below) also have "... \ (x \ y) ; z + (x \ w) ; ((x \ w) \ ((x \ y) ; z))" by (metis add_isotone mult_right_sub_dist_add_left while_left_unfold) finally show ?thesis by (metis while_left_unfold) qed -- {* Theorem 9.37 *} lemma while_while_sub_associative: "x \ (y \ z) \ ((x \ y) \ z) + (x \ z)" proof - have 1: "x ; (x \ 1) \ (x \ 1) ; ((x \ y) \ 1)" by (metis add_least_upper_bound order_trans while_back_loop_prefixpoint while_left_plus_below) have "x \ (y \ z) \ x \ ((x \ 1) ; (y \ z))" by (metis mult_left_isotone mult_left_one while_increasing while_right_isotone) also have "... \ (x \ 1) ; ((x \ y) \ (y \ z)) + (x \ 0)" using 1 by (metis while_simulate_left_plus_1) also have "... \ (x \ 1) ; ((x \ y) \ z) + (x \ z)" by (metis add_isotone order_refl while_absorb_2 while_increasing while_right_isotone zero_least) also have "... = (x \ 1) ; z + (x \ 1) ; (x \ y) ; ((x \ y) \ z) + (x \ z)" by (metis mult_associative mult_left_dist_add while_left_unfold) also have "... = (x \ y) ; ((x \ y) \ z) + (x \ z)" by (smt add_associative add_commutative less_eq_def mult_left_one mult_right_dist_add order_refl while_absorb_1 while_plus_one while_sub_associative) also have "... \ ((x \ y) \ z) + (x \ z)" by (metis add_left_isotone while_left_plus_below) finally show ?thesis . qed lemma while_induct: "x ; z \ z \ y \ z \ x \ 1 \ z \ x \ y \ z" by (metis add_commutative add_least_upper_bound add_left_zero less_eq_def while_right_isotone while_simulate_absorb) lemma while_sumstar_4_below: "(x \ y) \ ((x \ 1) ; z) \ x \ ((y ; (x \ 1)) \ z)" oops lemma while_sumstar_2: "(x + y) \ z = x \ ((y ; (x \ 1)) \ z)" oops lemma while_sumstar_3: "(x + y) \ z = ((x \ 1) ; y) \ (x \ z)" oops lemma while_decompose_6: "x \ ((y ; (x \ 1)) \ z) = y \ ((x ; (y \ 1)) \ z)" oops lemma while_finite_associative: "x \ 0 = 0 \ (x \ y) ; z = x \ (y ; z)" oops lemma atomicity_refinement: "s = s ; q \ x = q ; x \ q ; b = 0 \ r ; b \ b ; r \ r ; l \ l ; r \ x ; l \ l ; x \ b ; l \ l ; b \ q ; l \ l ; q \ r \ q \ q ; (r \ 1) \ q \ 1 \ s ; ((x + b + r + l) \ (q ; z)) \ s ; ((x ; (b \ q) + r + l) \ z)" oops lemma while_separate_right_plus: "y ; x \ x ; (x \ (1 + y)) + 1 \ y \ (x \ z) \ x \ (y \ z)" oops lemma while_square_1: "x \ 1 = (x ; x) \ (x + 1)" oops lemma while_absorb_below_one: "y ; x \ x \ y \ x \ 1 \ x" oops lemma "y \ (x \ 1) \ x \ (y \ 1) \ (x + y) \ 1 = x \ (y \ 1)" oops lemma "y ; x \ (1 + x) ; (y \ 1) \ (x + y) \ 1 = x \ (y \ 1)" oops end class bounded_binary_itering = bounded_idempotent_left_zero_semiring + binary_itering begin -- {* Theorem 9 *} lemma while_right_top: "x \ T = T" by (metis add_left_top while_left_unfold) -- {* Theorem 9 *} lemma while_left_top: "T ; (x \ 1) = T" by (metis add_right_top antisym top_greatest while_back_loop_prefixpoint) -- {* Theorem 10.10 counterexamples *} lemma while_sum_below_one: "y ; ((x + y) \ z) \ (y ; (x \ 1)) \ z" nitpick [expect=genuine] oops lemma while_denest_1: "w ; (x \ (y ; z)) \ (w ; (x \ y)) \ z" nitpick [expect=genuine] oops lemma while_denest_2: "w ; ((x \ (y ; w)) \ z) = w ; (((x \ y) ; w) \ z)" nitpick [expect=genuine] oops lemma while_denest_4: "(x \ w) \ (x \ (y ; z)) = (x \ w) \ ((x \ y) ; z)" nitpick [expect=genuine] oops lemma while_denest_6: "(w ; (x \ y)) \ z = z + w ; ((x + y ; w) \ (y ; z))" nitpick [expect=genuine] oops lemma while_decompose_6: "x \ ((y ; (x \ 1)) \ z) = y \ ((x ; (y \ 1)) \ z)" oops lemma while_finite_associative: "x \ 0 = 0 \ (x \ y) ; z = x \ (y ; z)" oops lemma while_sumstar_2: "(x + y) \ z = x \ ((y ; (x \ 1)) \ z)" oops lemma while_sumstar_3: "(x + y) \ z = ((x \ 1) ; y) \ (x \ z)" oops lemma while_sumstar_1: "(x + y) \ z = (x \ y) \ ((x \ 1) ; z)" nitpick [expect=genuine] oops lemma while_mult_zero_zero: "(x ; (y \ 0)) \ 0 = x ; (y \ 0)" nitpick [expect=genuine] oops lemma while_denest_3: "(x \ w) \ (x \ 0) = (x \ w) \ 0" nitpick [expect=genuine] oops lemma while_denest_5: "w ; ((x \ (y ; w)) \ (x \ (y ; z))) = w ; (((x \ y) ; w) \ ((x \ y) ; z))" nitpick [expect=genuine] oops lemma while_separate_unfold: "(y ; (x \ 1)) \ z = (y \ z) + (y \ (y ; x ; (x \ ((y ; (x \ 1)) \ z))))" nitpick [expect=genuine] oops lemma while_02: "x \ ((x \ w) \ ((x \ y) ; z)) = (x \ w) \ ((x \ y) ; z)" nitpick [expect=genuine] oops lemma while_denest_0: "w ; (x \ (y ; z)) \ (w ; (x \ y)) \ (w ; (x \ y) ; z)" nitpick [expect=genuine] oops lemma while_mult_sub_while_while: "x \ (y ; z) \ (x \ y) \ z" nitpick [expect=genuine] oops lemma while_zero_zero: "(x \ 0) \ 0 = x \ 0" nitpick [expect=genuine] oops lemma while_sumstar_3_below: "(x \ y) \ (x \ z) \ (x \ y) \ ((x \ 1) ; z)" nitpick [expect=genuine] oops end class extended_binary_itering = binary_itering + assumes while_denest_0: "w ; (x \ (y ; z)) \ (w ; (x \ y)) \ (w ; (x \ y) ; z)" begin -- {* Theorem 10.2 *} lemma while_denest_1: "w ; (x \ (y ; z)) \ (w ; (x \ y)) \ z" by (metis order_trans while_denest_0 while_right_plus_below) lemma while_mult_sub_while_while: "x \ (y ; z) \ (x \ y) \ z" by (metis mult_left_one while_denest_1) lemma while_zero_zero: "(x \ 0) \ 0 = x \ 0" by (smt less_eq_def mult_left_zero while_left_dist_add while_mult_star_exchange while_mult_sub_while_while while_mult_zero_add_2 while_plus_one while_sumstar) -- {* Theorem 10.11 *} lemma while_mult_zero_zero: "(x ; (y \ 0)) \ 0 = x ; (y \ 0)" apply (rule antisym) apply (metis add_least_upper_bound add_right_zero mult_left_zero mult_right_isotone while_left_dist_add while_slide while_sub_associative) apply (metis mult_left_zero while_denest_1) done -- {* Theorem 10.3 *} lemma while_denest_2: "w ; ((x \ (y ; w)) \ z) = w ; (((x \ y) ; w) \ z)" apply (rule antisym) apply (metis mult_associative while_denest_0 while_simulate_right_plus_1 while_slide) apply (metis mult_right_isotone while_left_isotone while_sub_associative) done -- {* Theorem 10.12 *} lemma while_denest_3: "(x \ w) \ (x \ 0) = (x \ w) \ 0" by (metis while_absorb_2 while_right_isotone while_zero_zero zero_least) -- {* Theorem 10.15 *} lemma while_02: "x \ ((x \ w) \ ((x \ y) ; z)) = (x \ w) \ ((x \ y) ; z)" proof - have "x ; ((x \ w) \ ((x \ y) ; z)) = x ; (x \ y) ; z + x ; (x \ w) ; ((x \ w) \ ((x \ y) ; z))" by (metis mult_associative mult_left_dist_add while_left_unfold) also have "... \ (x \ w) \ ((x \ y) ; z)" by (metis add_isotone mult_right_sub_dist_add_right while_left_unfold) finally have "x \ ((x \ w) \ ((x \ y) ; z)) \ ((x \ w) \ ((x \ y) ; z)) + (x \ 0)" by (metis while_simulate_absorb) also have "... = (x \ w) \ ((x \ y) ; z)" by (metis add_commutative less_eq_def order_trans while_mult_sub_while_while while_right_isotone zero_least) finally show ?thesis by (metis antisym while_increasing) qed lemma while_sumstar_3_below: "(x \ y) \ (x \ z) \ (x \ y) \ ((x \ 1) ; z)" proof - have "(x \ y) \ (x \ z) = (x \ z) + ((x \ y) \ ((x \ y) ; (x \ z)))" by (metis while_right_unfold) also have "... \ (x \ z) + ((x \ y) \ (x \ (y ; (x \ z))))" by (metis add_right_isotone while_right_isotone while_sub_associative) also have "... \ (x \ z) + ((x \ y) \ (x \ ((x \ y) \ (x \ z))))" by (smt add_right_isotone order_trans while_increasing while_mult_upper_bound while_one_increasing while_right_isotone) also have "... \ (x \ z) + ((x \ y) \ (x \ ((x \ y) \ ((x \ 1) ; z))))" by (metis add_right_isotone mult_left_isotone mult_left_one order_trans while_increasing while_right_isotone while_sumstar while_transitive) also have "... = (x \ z) + ((x \ y) \ ((x \ 1) ; z))" by (metis while_02 while_transitive) also have "... = (x \ y) \ ((x \ 1) ; z)" by (smt add_associative mult_left_one mult_right_dist_add while_02 while_left_dist_add while_plus_one) finally show ?thesis . qed lemma while_sumstar_4_below: "(x \ y) \ ((x \ 1) ; z) \ x \ ((y ; (x \ 1)) \ z)" proof - have "(x \ y) \ ((x \ 1) ; z) = (x \ 1) ; z + (x \ y) ; ((x \ y) \ ((x \ 1) ; z))" by (metis while_left_unfold) also have "... \ (x \ z) + (x \ (y ; ((x \ y) \ ((x \ 1) ; z))))" by (metis add_isotone while_one_mult_below while_sub_associative) also have "... = (x \ z) + (x \ (y ; (((x \ 1) ; y) \ ((x \ 1) ; z))))" by (metis mult_left_one while_denest_2) also have "... = x \ ((y ; (x \ 1)) \ z)" by (metis while_left_dist_add while_productstar) finally show ?thesis . qed -- {* Theorem 10.10 *} lemma while_sumstar_1: "(x + y) \ z = (x \ y) \ ((x \ 1) ; z)" by (smt eq_iff order_trans while_add_1_below while_sumstar while_sumstar_3_below while_sumstar_4_below) -- {* Theorem 10.8 *} lemma while_sumstar_2: "(x + y) \ z = x \ ((y ; (x \ 1)) \ z)" by (metis eq_iff while_add_1_below while_sumstar_1 while_sumstar_4_below) -- {* Theorem 10.9 *} lemma while_sumstar_3: "(x + y) \ z = ((x \ 1) ; y) \ (x \ z)" by (metis eq_iff while_sumstar while_sumstar_1_below while_sumstar_2 while_sumstar_2_below) -- {* Theorem 10.6 *} lemma while_decompose_6: "x \ ((y ; (x \ 1)) \ z) = y \ ((x ; (y \ 1)) \ z)" by (metis add_commutative while_sumstar_2) -- {* Theorem 10.4 *} lemma while_denest_4: "(x \ w) \ (x \ (y ; z)) = (x \ w) \ ((x \ y) ; z)" proof - have "(x \ w) \ (x \ (y ; z)) = x \ ((w ; (x \ 1)) \ (y ; z))" by (metis while_sumstar while_sumstar_2) also have "... \ (x \ w) \ ((x \ y) ; z)" by (smt antisym while_01 while_02 while_increasing while_right_isotone) finally show ?thesis by (metis antisym while_right_isotone while_sub_associative) qed -- {* Theorem 10.13 *} lemma while_denest_5: "w ; ((x \ (y ; w)) \ (x \ (y ; z))) = w ; (((x \ y) ; w) \ ((x \ y) ; z))" by (metis while_denest_2 while_denest_4) -- {* Theorem 10.5 *} lemma while_denest_6: "(w ; (x \ y)) \ z = z + w ; ((x + y ; w) \ (y ; z))" by (metis while_denest_5 while_productstar while_sumstar) -- {* Theorem 10.1 *} lemma while_sum_below_one: "y ; ((x + y) \ z) \ (y ; (x \ 1)) \ z" by (metis add_right_divisibility mult_left_one while_denest_6) -- {* Theorem 10.14 *} lemma while_separate_unfold: "(y ; (x \ 1)) \ z = (y \ z) + (y \ (y ; x ; (x \ ((y ; (x \ 1)) \ z))))" proof - have "y \ (y ; x ; (x \ ((y ; (x \ 1)) \ z))) \ y \ (y ; ((x + y) \ z))" by (metis mult_associative mult_right_isotone while_sumstar_2 while_left_plus_below while_right_isotone) also have "... \ (y ; (x \ 1)) \ z" by (metis add_commutative add_left_upper_bound while_absorb_1 while_mult_star_exchange while_sum_below_one) finally have "(y \ z) + (y \ (y ; x ; (x \ ((y ; (x \ 1)) \ z)))) \ (y ; (x \ 1)) \ z" by (metis add_least_upper_bound mult_left_sub_dist_add_left mult_right_one while_left_isotone while_left_unfold) thus ?thesis by (metis antisym while_separate_unfold_below) qed -- {* Theorem 10.7 *} lemma while_finite_associative: "x \ 0 = 0 \ (x \ y) ; z = x \ (y ; z)" by (metis while_denest_4 while_zero) -- {* Theorem 12 *} lemma atomicity_refinement: "s = s ; q \ x = q ; x \ q ; b = 0 \ r ; b \ b ; r \ r ; l \ l ; r \ x ; l \ l ; x \ b ; l \ l ; b \ q ; l \ l ; q \ r \ q \ q ; (r \ 1) \ q \ 1 \ s ; ((x + b + r + l) \ (q ; z)) \ s ; ((x ; (b \ q) + r + l) \ z)" proof assume 1: "s = s ; q \ x = q ; x \ q ; b = 0 \ r ; b \ b ; r \ r ; l \ l ; r \ x ; l \ l ; x \ b ; l \ l ; b \ q ; l \ l ; q \ r \ q \ q ; (r \ 1) \ q \ 1" hence 2: "(x + b + r) ; l \ l ; (x + b + r)" by (smt add_commutative add_least_upper_bound mult_left_sub_dist_add_right mult_right_dist_add order_trans) have "q ; ((x ; (b \ r \ 1) ; q) \ z) \ (x ; (b \ r \ 1) ; q) \ z" using 1 by (smt eq_refl order_trans while_increasing while_mult_upper_bound) also have "... \ (x ; (b \ ((r \ 1) ; q))) \ z" by (metis mult_associative mult_right_isotone while_left_isotone while_sub_associative) also have "... \ (x ; (b \ r \ q)) \ z" by (metis mult_right_isotone while_left_isotone while_one_mult_below while_right_isotone) also have "... \ (x ; (b \ (q ; (r \ 1)))) \ z" using 1 by (metis mult_right_isotone while_left_isotone while_right_isotone) finally have 3: "q ; ((x ; (b \ r \ 1) ; q) \ z) \ (x ; (b \ q) ; (r \ 1)) \ z" by (metis mult_associative while_associative_while_1) have "s ; ((x + b + r + l) \ (q ; z)) = s ; (l \ (x + b + r) \ (q ; z))" using 2 by (metis add_commutative while_separate_1) also have "... = s ; q ; (l \ b \ r \ (q ; x ; (b \ r \ 1)) \ (q ; z))" using 1 by (smt add_associative add_commutative while_sumstar_2 while_separate_1) also have "... = s ; q ; (l \ b \ r \ (q ; ((x ; (b \ r \ 1) ; q) \ z)))" by (smt mult_associative while_slide) also have "... \ s ; q ; (l \ b \ r \ (x ; (b \ q) ; (r \ 1)) \ z)" using 3 by (metis mult_right_isotone while_right_isotone) also have "... \ s ; (l \ q ; (b \ r \ (x ; (b \ q) ; (r \ 1)) \ z))" using 1 by (smt mult_associative mult_right_isotone while_simulate) also have "... = s ; (l \ q ; (r \ (x ; (b \ q) ; (r \ 1)) \ z))" using 1 by (metis while_elimination) also have "... \ s ; (l \ r \ (x ; (b \ q) ; (r \ 1)) \ z)" using 1 by (metis add_left_divisibility mult_left_one mult_right_dist_add mult_right_isotone while_right_isotone) also have "... = s ; (l \ (r + x ; (b \ q)) \ z)" by (metis while_sumstar_2) also have "... \ s ; ((x ; (b \ q) + r + l) \ z)" by (metis add_commutative mult_right_isotone while_sub_dist_3) finally show "s ; ((x + b + r + l) \ (q ; z)) \ s ; ((x ; (b \ q) + r + l) \ z)" . qed end class bounded_extended_binary_itering = bounded_binary_itering + extended_binary_itering begin lemma while_unfold_below: "x = z + y ; x \ x \ y \ z" nitpick [expect=genuine] oops lemma while_unfold_below_1: "x = y ; x \ x \ y \ 1" nitpick [expect=genuine] oops lemma while_loop_is_greatest_postfixpoint: "is_greatest_postfixpoint (\x . y ; x + z) (y \ z)" nitpick [expect=genuine] oops lemma while_loop_is_greatest_fixpoint: "is_greatest_fixpoint (\x . y ; x + z) (y \ z)" nitpick [expect=genuine] oops lemma while_sub_mult_one: "x ; (1 \ y) \ 1 \ x" nitpick [expect=genuine] oops lemma while_sub_while_zero: "x \ z \ (x \ y) \ z" nitpick [expect=genuine] oops lemma while_while_sub_associative: "x \ (y \ z) \ (x \ y) \ z" nitpick [expect=genuine] oops lemma while_top: "T \ x = T" nitpick [expect=genuine] oops lemma while_one_top: "1 \ x = T" nitpick [expect=genuine] oops lemma while_mult_top: "(x ; T) \ z = z + x ; T" nitpick [expect=genuine] oops lemma tarski: "x \ x ; T ; x ; T" nitpick [expect=genuine] oops lemma tarski_mult_top_idempotent: "x ; T = x ; T ; x ; T" nitpick [expect=genuine] oops lemma tarski_top_omega_below: "x ; T \ (x ; T) \ 0" nitpick [expect=genuine] oops lemma tarski_top_omega: "x ; T = (x ; T) \ 0" nitpick [expect=genuine] oops lemma tarski_below_top_omega: "x \ (x ; T) \ 0" nitpick [expect=genuine] oops lemma tarski: "x = 0 \ T ; x ; T = T" nitpick [expect=genuine] oops lemma "(x + y) \ z = ((x \ 1) ; y) \ ((x \ 1) ; z)" nitpick [expect=genuine] oops lemma "1 = (x ; 0) \ 1" nitpick [expect=genuine] oops lemma while_top_2: "T \ z = T ; z" nitpick [expect=genuine] oops lemma while_mult_top_2: "(x ; T) \ z = z + x ; T ; z" nitpick [expect=genuine] oops lemma while_one_mult: "(x \ 1) ; x = x \ x" nitpick [expect=genuine] oops lemma "(x \ 1) ; y = x \ y" nitpick [expect=genuine] oops lemma while_associative: "(x \ y) ; z = x \ (y ; z)" nitpick [expect=genuine] oops lemma while_back_loop_is_fixpoint: "is_fixpoint (\x . x ; y + z) (z ; (y \ 1))" nitpick [expect=genuine] oops lemma "1 + x ; 0 = x \ 1" nitpick [expect=genuine] oops lemma "x = x ; (x \ 1)" nitpick [expect=genuine] oops lemma "x ; (x \ 1) = x \ 1" nitpick [expect=genuine] oops lemma "x \ 1 = x \ (1 \ 1)" nitpick [expect=genuine] oops lemma "(x + y) \ 1 = (x \ (y \ 1)) \ 1" nitpick [expect=genuine] oops lemma "z + y ; x = x \ y \ z \ x" nitpick [expect=genuine] oops lemma "y ; x = x \ y \ x \ x" nitpick [expect=genuine] oops lemma "z + x ; y = x \ z ; (y \ 1) \ x" nitpick [expect=genuine] oops lemma "x ; y = x \ x ; (y \ 1) \ x" nitpick [expect=genuine] oops lemma "x ; z = z ; y \ x \ z \ z ; (y \ 1)" nitpick [expect=genuine] oops lemma "x ; ((y ; x) \ y) \ x ; ((y ; x) \ 1) ; y" nitpick [expect=genuine] oops end (* end theory BinaryIteringStrict imports BinaryItering Itering begin *) class strict_itering = itering + while + assumes while_def: "x \ y = x\<^sup>\ ; y" begin -- {* Theorem 8.1 *} subclass extended_binary_itering apply unfold_locales apply (metis add_commutative circ_loop_fixpoint circ_slide mult_associative while_def) apply (metis circ_add mult_associative while_def) apply (metis mult_left_dist_add while_def) apply (metis mult_associative order_refl while_def) apply (metis circ_simulate_left_plus mult_associative mult_left_isotone mult_right_dist_add mult_right_one while_def) apply (metis circ_simulate_right_plus mult_associative mult_left_isotone mult_right_dist_add while_def) apply (metis add_right_divisibility circ_loop_fixpoint mult_associative while_def) done -- {* Theorem 13.1 *} lemma while_associative: "(x \ y) ; z = x \ (y ; z)" by (metis mult_associative while_def) -- {* Theorem 13.3 *} lemma while_one_mult: "(x \ 1) ; x = x \ x" by (metis mult_right_one while_def) lemma while_back_loop_is_fixpoint: "is_fixpoint (\x . x ; y + z) (z ; (y \ 1))" by (metis circ_back_loop_is_fixpoint mult_right_one while_def) -- {* Theorem 13.4 *} lemma "(x + y) \ z = ((x \ 1) ; y) \ ((x \ 1) ; z)" by (metis mult_right_one while_def while_sumstar) -- {* Theorem 13.2 *} lemma "(x \ 1) ; y = x \ y" by (metis mult_left_one while_associative) lemma "y \ (x \ 1) \ x \ (y \ 1) \ (x + y) \ 1 = x \ (y \ 1)" oops lemma "y ; x \ (1 + x) ; (y \ 1) \ (x + y) \ 1 = x \ (y \ 1)" oops lemma while_square_1: "x \ 1 = (x ; x) \ (x + 1)" oops lemma while_absorb_below_one: "y ; x \ x \ y \ x \ 1 \ x" oops end class bounded_strict_itering = bounded_itering + strict_itering begin subclass bounded_extended_binary_itering .. -- {* Theorem 13.6 *} lemma while_top_2: "T \ z = T ; z" by (metis circ_top while_def) -- {* Theorem 13.5 *} lemma while_mult_top_2: "(x ; T) \ z = z + x ; T ; z" by (metis circ_left_top mult_associative while_def while_left_unfold) -- {* Theorem 13 counterexamples *} lemma while_one_top: "1 \ x = T" nitpick [expect=genuine] oops lemma while_top: "T \ x = T" nitpick [expect=genuine] oops lemma while_sub_mult_one: "x ; (1 \ y) \ 1 \ x" nitpick [expect=genuine] oops lemma while_unfold_below_1: "x = y ; x \ x \ y \ 1" nitpick [expect=genuine] oops lemma while_unfold_below: "x = z + y ; x \ x \ y \ z" nitpick [expect=genuine] oops lemma while_unfold_below: "x \ z + y ; x \ x \ y \ z" nitpick [expect=genuine] oops lemma while_mult_top: "(x ; T) \ z = z + x ; T" nitpick [expect=genuine] oops lemma tarski_mult_top_idempotent: "x ; T = x ; T ; x ; T" nitpick [expect=genuine] oops lemma while_loop_is_greatest_postfixpoint: "is_greatest_postfixpoint (\x . y ; x + z) (y \ z)" nitpick [expect=genuine] oops lemma while_loop_is_greatest_fixpoint: "is_greatest_fixpoint (\x . y ; x + z) (y \ z)" nitpick [expect=genuine] oops lemma while_sub_while_zero: "x \ z \ (x \ y) \ z" nitpick [expect=genuine] oops lemma while_while_sub_associative: "x \ (y \ z) \ (x \ y) \ z" nitpick [expect=genuine] oops lemma tarski: "x \ x ; T ; x ; T" nitpick [expect=genuine] oops lemma tarski_top_omega_below: "x ; T \ (x ; T) \ 0" nitpick [expect=genuine] oops lemma tarski_top_omega: "x ; T = (x ; T) \ 0" nitpick [expect=genuine] oops lemma tarski_below_top_omega: "x \ (x ; T) \ 0" nitpick [expect=genuine] oops lemma tarski: "x = 0 \ T ; x ; T = T" nitpick [expect=genuine] oops lemma "1 = (x ; 0) \ 1" nitpick [expect=genuine] oops lemma "1 + x ; 0 = x \ 1" nitpick [expect=genuine] oops lemma "x = x ; (x \ 1)" nitpick [expect=genuine] oops lemma "x ; (x \ 1) = x \ 1" nitpick [expect=genuine] oops lemma "x \ 1 = x \ (1 \ 1)" nitpick [expect=genuine] oops lemma "(x + y) \ 1 = (x \ (y \ 1)) \ 1" nitpick [expect=genuine] oops lemma "z + y ; x = x \ y \ z \ x" nitpick [expect=genuine] oops lemma "y ; x = x \ y \ x \ x" nitpick [expect=genuine] oops lemma "z + x ; y = x \ z ; (y \ 1) \ x" nitpick [expect=genuine] oops lemma "x ; y = x \ x ; (y \ 1) \ x" nitpick [expect=genuine] oops lemma "x ; z = z ; y \ x \ z \ z ; (y \ 1)" nitpick [expect=genuine] oops end class binary_itering_unary = extended_binary_itering + circ + assumes circ_def: "x\<^sup>\ = x \ 1" begin -- {* Theorem 50.7 *} subclass left_conway_semiring apply unfold_locales apply (metis circ_def while_left_unfold) apply (metis circ_def mult_right_one while_one_mult_below while_slide) apply (metis circ_def while_one_while while_sumstar_2) done end class strict_binary_itering = binary_itering + circ + assumes while_associative: "(x \ y) ; z = x \ (y ; z)" assumes circ_def: "x\<^sup>\ = x \ 1" begin -- {* Theorem 2.8 *} subclass itering apply unfold_locales apply (metis circ_def mult_left_one_1 while_associative while_sumstar) apply (metis circ_def mult_right_one while_associative while_productstar while_slide) apply (metis circ_def mult_right_one while_associative mult_left_one_1 while_slide while_simulate_right_plus) apply (metis circ_def mult_right_one while_associative mult_left_one_1 while_simulate_left_plus mult_right_dist_add) done -- {* Theorem 8.5 *} subclass extended_binary_itering apply unfold_locales apply (metis mult_associative while_associative while_increasing) done end (* end theory Approximation imports Semiring begin *) class apx = fixes apx :: "'a \ 'a \ bool" (infix "\" 50) class apx_order = apx + assumes apx_reflexive: "x \ x" assumes apx_antisymmetric: "x \ y \ y \ x \ x = y" assumes apx_transitive: "x \ y \ y \ z \ x \ z" sublocale apx_order < apx!: order where less_eq = apx and less = "\x y . x \ y \ \ y \ x" apply unfold_locales apply rule apply (rule apx_reflexive) apply (metis apx_transitive) apply (metis apx_antisymmetric) done context apx_order begin abbreviation the_apx_least_fixpoint :: "('a \ 'a) \ 'a" ("\ _" [201] 200) where "\ f \ apx.the_least_fixpoint f" abbreviation the_apx_least_prefixpoint :: "('a \ 'a) \ 'a" ("p\ _" [201] 200) where "p\ f \ apx.the_least_prefixpoint f" definition is_apx_meet :: "'a \ 'a \ 'a \ bool" where "is_apx_meet x y z \ z \ x \ z \ y \ (\w . w \ x \ w \ y \ w \ z)" definition has_apx_meet :: "'a \ 'a \ bool" where "has_apx_meet x y \ (\z . is_apx_meet x y z)" definition the_apx_meet :: "'a \ 'a \ 'a" (infixl "\" 66) where "x \ y = (THE z . is_apx_meet x y z)" lemma apx_meet_unique: "has_apx_meet x y \ (\!z . is_apx_meet x y z)" by (smt apx_antisymmetric has_apx_meet_def is_apx_meet_def) lemma apx_meet: "has_apx_meet x y \ is_apx_meet x y (x \ y)" proof assume "has_apx_meet x y" hence "is_apx_meet x y (THE z . is_apx_meet x y z)" by (smt apx_meet_unique theI') thus "is_apx_meet x y (x \ y)" by (simp add: is_apx_meet_def the_apx_meet_def) qed lemma apx_greatest_lower_bound: "has_apx_meet x y \ (w \ x \ w \ y \ w \ x \ y)" by (smt apx_meet apx_transitive is_apx_meet_def) lemma apx_meet_same: "is_apx_meet x y z \ z = x \ y" by (metis apx_meet apx_meet_unique has_apx_meet_def) lemma apx_meet_char: "is_apx_meet x y z \ has_apx_meet x y \ z = x \ y" by (metis apx_meet_same has_apx_meet_def) end class apx_biorder = apx_order + order begin lemma mu_below_kappa: "has_least_fixpoint f \ apx.has_least_fixpoint f \ \ f \ \ f" by (metis apx.least_fixpoint apx.is_least_fixpoint_def is_least_fixpoint_def least_fixpoint) lemma kappa_below_nu: "has_greatest_fixpoint f \ apx.has_least_fixpoint f \ \ f \ \ f" by (metis apx.least_fixpoint greatest_fixpoint apx.is_least_fixpoint_def is_greatest_fixpoint_def) lemma kappa_apx_below_mu: "has_least_fixpoint f \ apx.has_least_fixpoint f \ \ f \ \ f" by (metis apx.least_fixpoint apx.is_least_fixpoint_def is_least_fixpoint_def least_fixpoint) lemma kappa_apx_below_nu: "has_greatest_fixpoint f \ apx.has_least_fixpoint f \ \ f \ \ f" by (metis apx.least_fixpoint greatest_fixpoint apx.is_least_fixpoint_def is_greatest_fixpoint_def) end class apx_semiring = apx_biorder + idempotent_left_semiring + L + assumes apx_L_least: "L \ x" assumes add_apx_left_isotone: "x \ y \ x + z \ y + z" assumes mult_apx_left_isotone: "x \ y \ x ; z \ y ; z" assumes mult_apx_right_isotone: "x \ y \ z ; x \ z ; y" begin lemma add_apx_right_isotone: "x \ y \ z + x \ z + y" by (metis add_apx_left_isotone add_commutative) lemma add_apx_isotone: "w \ y \ x \ z \ w + x \ y + z" by (metis add_apx_left_isotone add_apx_right_isotone apx_transitive) lemma mult_apx_isotone: "w \ y \ x \ z \ w ; x \ y ; z" by (metis apx_transitive mult_apx_left_isotone mult_apx_right_isotone) lemma affine_apx_isotone: "apx.isotone (\x . y ; x + z)" by (smt add_apx_left_isotone apx.isotone_def mult_apx_right_isotone) end (* end theory LatticeOrderedSemiring imports Semiring begin *) -- {* Many results in this theory are taken from a joint paper with Rudolf Berghammer. *} -- {* M0-algebra *} class lattice_ordered_pre_left_semiring = pre_left_semiring + bounded_distributive_lattice begin subclass bounded_pre_left_semiring apply unfold_locales apply (metis add_right_top_1) done lemma top_mult_right_one: "x ; T = x ; T ; 1" by (metis add_commutative add_left_top less_eq_def mult_semi_associative mult_sub_right_one) lemma mult_left_sub_dist_meet_left: "x ; (y \ z) \ x ; y" by (metis meet.add_left_upper_bound mult_right_isotone) lemma mult_left_sub_dist_meet_right: "x ; (y \ z) \ x ; z" by (metis meet_commutative mult_left_sub_dist_meet_left) lemma mult_right_sub_dist_meet_left: "(x \ y) ; z \ x ; z" by (metis meet.add_left_upper_bound mult_left_isotone) lemma mult_right_sub_dist_meet_right: "(x \ y) ; z \ y ; z" by (metis meet.add_right_upper_bound mult_left_isotone) lemma mult_right_sub_dist_meet: "(x \ y) ; z \ x ; z \ y ; z" by (metis meet.add_least_upper_bound mult_right_sub_dist_meet_left mult_right_sub_dist_meet_right) -- {* Figure 1: fundamental properties *} definition total :: "'a \ bool" where "total x \ x ; T = T" definition co_total :: "'a \ bool" where "co_total x \ x ; 0 = 0" definition transitive :: "'a \ bool" where "transitive x \ x ; x \ x" definition dense :: "'a \ bool" where "dense x \ x \ x ; x" definition reflexive :: "'a \ bool" where "reflexive x \ 1 \ x" definition co_reflexive :: "'a \ bool" where "co_reflexive x \ x \ 1" definition idempotent :: "'a \ bool" where "idempotent x \ x ; x = x" definition up_closed :: "'a \ bool" where "up_closed x \ x ; 1 = x" definition add_distributive :: "'a \ bool" where "add_distributive x \ (\y z . x ; (y + z) = x ; y + x ; z)" definition meet_distributive :: "'a \ bool" where "meet_distributive x \ (\y z . x ; (y \ z) = x ; y \ x ; z)" definition contact :: "'a \ bool" where "contact x \ x ; x + 1 = x" definition kernel :: "'a \ bool" where "kernel x \ x ; x \ 1 = x ; 1" definition add_dist_contact :: "'a \ bool" where "add_dist_contact x \ add_distributive x \ contact x" definition meet_dist_kernel :: "'a \ bool" where "meet_dist_kernel x \ meet_distributive x \ kernel x" definition test :: "'a \ bool" where "test x \ x ; T \ 1 = x" definition co_test :: "'a \ bool" where "co_test x \ x ; 0 + 1 = x" definition co_vector :: "'a \ bool" where "co_vector x \ x ; 0 = x" -- {* CPCP Theorem 5 / Figure 2: relations between properties *} lemma reflexive_total: "reflexive x \ total x" by (metis eq_iff mult_isotone mult_left_one meet.zero_least reflexive_def total_def) lemma reflexive_dense: "reflexive x \ dense x" by (metis mult_left_isotone mult_left_one reflexive_def dense_def) lemma reflexive_transitive_up_closed: "reflexive x \ transitive x \ up_closed x" by (metis antisym_conv mult_isotone mult_sub_right_one reflexive_def reflexive_dense transitive_def dense_def up_closed_def) lemma co_reflexive_co_total: "co_reflexive x \ co_total x" by (metis co_reflexive_def co_total_def eq_iff mult_left_isotone mult_left_one zero_least) lemma co_reflexive_transitive: "co_reflexive x \ transitive x" by (metis co_reflexive_def mult_left_isotone mult_left_one transitive_def) lemma idempotent_transitive_dense: "idempotent x \ transitive x \ dense x" by (metis eq_iff transitive_def dense_def idempotent_def) lemma contact_reflexive: "contact x \ reflexive x" by (metis contact_def add_right_upper_bound reflexive_def) lemma contact_transitive: "contact x \ transitive x" by (metis contact_def add_left_upper_bound transitive_def) lemma contact_dense: "contact x \ dense x" by (metis contact_reflexive reflexive_dense) lemma contact_idempotent: "contact x \ idempotent x" by (metis contact_transitive contact_dense idempotent_transitive_dense) lemma contact_up_closed: "contact x \ up_closed x" by (metis contact_def contact_idempotent dual_order.antisym mult_left_sub_dist_add_right mult_sub_right_one idempotent_def up_closed_def) lemma contact_reflexive_idempotent_up_closed: "contact x \ reflexive x \ idempotent x \ up_closed x" by (metis contact_def contact_idempotent contact_up_closed add_commutative less_eq_def reflexive_def idempotent_def) lemma kernel_co_reflexive: "kernel x \ co_reflexive x" by (metis co_reflexive_def kernel_def meet.add_least_upper_bound mult_sub_right_one) lemma kernel_transitive: "kernel x \ transitive x" by (metis co_reflexive_transitive kernel_co_reflexive) lemma kernel_dense: "kernel x \ dense x" by (metis kernel_def meet.add_least_upper_bound mult_sub_right_one dense_def) lemma kernel_idempotent: "kernel x \ idempotent x" by (metis kernel_transitive kernel_dense idempotent_transitive_dense) lemma kernel_up_closed: "kernel x \ up_closed x" by (metis co_reflexive_def kernel_co_reflexive kernel_def kernel_idempotent meet_less_eq_def idempotent_def up_closed_def) lemma kernel_co_reflexive_idempotent_up_closed: "kernel x \ co_reflexive x \ idempotent x \ up_closed x" by (metis co_reflexive_def kernel_def kernel_idempotent kernel_up_closed meet.less_eq_def meet_commutative idempotent_def up_closed_def) lemma test_co_reflexive: "test x \ co_reflexive x" by (metis co_reflexive_def meet.add_right_upper_bound test_def) lemma test_up_closed: "test x \ up_closed x" by (metis eq_iff mult_left_one mult_sub_right_one mult_right_sub_dist_meet test_def top_mult_right_one up_closed_def) lemma co_test_reflexive: "co_test x \ reflexive x" by (metis co_test_def add_right_upper_bound reflexive_def) lemma co_test_transitive: "co_test x \ transitive x" by (smt2 co_test_def add_associative less_eq_def mult_left_one mult_left_zero mult_right_dist_add mult_semi_associative transitive_def) lemma co_test_idempotent: "co_test x \ idempotent x" by (metis co_test_reflexive co_test_transitive reflexive_dense idempotent_transitive_dense) lemma co_test_up_closed: "co_test x \ up_closed x" by (metis co_test_reflexive co_test_idempotent contact_def contact_up_closed add_commutative less_eq_def reflexive_def idempotent_def) lemma co_test_contact: "co_test x \ contact x" by (metis co_test_reflexive co_test_idempotent co_test_up_closed contact_reflexive_idempotent_up_closed) lemma vector_transitive: "vector x \ transitive x" by (metis mult_right_isotone meet.zero_least vector_def transitive_def) lemma vector_up_closed: "vector x \ up_closed x" by (metis vector_def top_mult_right_one up_closed_def) -- {* CPCP Theorem 8 / Figure 3: closure properties *} -- {* total *} lemma one_total: "total 1" by (metis mult_left_one total_def) lemma top_total: "total T" by (metis top_mult_top total_def) lemma add_total: "total x \ total y \ total (x + y)" by (metis add_left_top mult_right_dist_add total_def) -- {* co-total *} lemma zero_co_total: "co_total 0" by (metis co_total_def mult_left_zero) lemma one_co_total: "co_total 1" by (metis co_total_def mult_left_one) lemma add_co_total: "co_total x \ co_total y \ co_total (x + y)" by (metis co_total_def add_right_zero mult_right_dist_add) lemma meet_co_total: "co_total x \ co_total y \ co_total (x \ y)" by (metis co_total_def add_left_zero antisym_conv less_eq_def mult_right_sub_dist_meet_left) lemma comp_co_total: "co_total x \ co_total y \ co_total (x ; y)" by (metis co_total_def eq_iff mult_semi_associative zero_least) -- {* sub-transitive *} lemma zero_transitive: "transitive 0" by (metis mult_left_zero zero_least transitive_def) lemma one_transitive: "transitive 1" by (metis mult_left_one order_refl transitive_def) lemma top_transitive: "transitive T" by (metis meet.zero_least transitive_def) lemma meet_transitive: "transitive x \ transitive y \ transitive (x \ y)" by (smt2 meet.less_eq_def meet_associative meet_commutative mult_left_sub_dist_meet_left mult_right_sub_dist_meet_left transitive_def) -- {* dense *} lemma zero_dense: "dense 0" by (metis zero_least dense_def) lemma one_dense: "dense 1" by (metis mult_sub_right_one dense_def) lemma top_dense: "dense T" by (metis top_left_mult_increasing dense_def) lemma add_dense: "dense x \ dense y \ dense (x + y)" proof assume "dense x \ dense y" hence "x \ x ; x \ y \ y ; y" by (metis dense_def) hence "x \ (x + y) ; (x + y) \ y \ (x + y) ; (x + y)" by (metis add_left_upper_bound dual_order.trans mult_isotone add_right_upper_bound) hence "x + y \ (x + y) ; (x + y)" by (metis add_least_upper_bound) thus "dense (x + y)" by (metis dense_def) qed -- {* reflexive *} lemma one_reflexive: "reflexive 1" by (metis order_refl reflexive_def) lemma top_reflexive: "reflexive T" by (metis meet.zero_least reflexive_def) lemma add_reflexive: "reflexive x \ reflexive y \ reflexive (x + y)" by (metis add_associative less_eq_def reflexive_def) lemma meet_reflexive: "reflexive x \ reflexive y \ reflexive (x \ y)" by (metis meet.add_least_upper_bound reflexive_def) lemma comp_reflexive: "reflexive x \ reflexive y \ reflexive (x ; y)" by (metis mult_left_isotone mult_left_one order_trans reflexive_def) -- {* co-reflexive *} lemma zero_co_reflexive: "co_reflexive 0" by (metis co_reflexive_def zero_least) lemma one_co_reflexive: "co_reflexive 1" by (metis co_reflexive_def order_refl) lemma add_co_reflexive: "co_reflexive x \ co_reflexive y \ co_reflexive (x + y)" by (metis co_reflexive_def add_least_upper_bound) lemma meet_co_reflexive: "co_reflexive x \ co_reflexive y \ co_reflexive (x \ y)" by (metis co_reflexive_def meet.less_eq_def meet_associative) lemma comp_co_reflexive: "co_reflexive x \ co_reflexive y \ co_reflexive (x ; y)" by (metis co_reflexive_def mult_isotone mult_left_one) -- {* idempotent *} lemma zero_idempotent: "idempotent 0" by (metis mult_left_zero idempotent_def) lemma one_idempotent: "idempotent 1" by (metis mult_left_one idempotent_def) lemma top_idempotent: "idempotent T" by (metis top_mult_top idempotent_def) -- {* up-closed *} lemma zero_up_closed: "up_closed 0" by (metis mult_left_zero up_closed_def) lemma one_up_closed: "up_closed 1" by (metis mult_left_one up_closed_def) lemma top_up_closed: "up_closed T" by (metis top_mult_top vector_def vector_up_closed) lemma add_up_closed: "up_closed x \ up_closed y \ up_closed (x + y)" by (metis mult_right_dist_add up_closed_def) lemma meet_up_closed: "up_closed x \ up_closed y \ up_closed (x \ y)" by (metis dual_order.antisym mult_sub_right_one mult_right_sub_dist_meet up_closed_def) lemma comp_up_closed: "up_closed x \ up_closed y \ up_closed (x ; y)" by (metis dual_order.antisym mult_semi_associative mult_sub_right_one up_closed_def) -- {* add-distributive *} lemma zero_add_distributive: "add_distributive 0" by (metis add_distributive_def add_idempotent mult_left_zero) lemma one_add_distributive: "add_distributive 1" by (metis add_distributive_def mult_left_one) lemma add_add_distributive: "add_distributive x \ add_distributive y \ add_distributive (x + y)" by (smt2 add_distributive_def add_associative add_commutative mult_right_dist_add) -- {* meet-distributive *} lemma zero_meet_distributive: "meet_distributive 0" by (metis meet_left_zero mult_left_zero meet_distributive_def) lemma one_meet_distributive: "meet_distributive 1" by (metis mult_left_one meet_distributive_def) -- {* contact *} lemma one_contact: "contact 1" by (metis contact_def add_idempotent mult_left_one) lemma top_contact: "contact T" by (metis contact_def add_left_top top_mult_top) lemma meet_contact: "contact x \ contact y \ contact (x \ y)" by (smt2 contact_def contact_reflexive contact_transitive contact_up_closed meet.less_eq_def meet_commutative meet_left_dist_add mult_left_sub_dist_add_right meet_transitive meet_up_closed reflexive_def transitive_def up_closed_def) -- {* kernel *} lemma zero_kernel: "kernel 0" by (metis kernel_co_reflexive_idempotent_up_closed zero_co_reflexive zero_idempotent zero_up_closed) lemma one_kernel: "kernel 1" by (metis kernel_def meet_idempotent mult_left_one) lemma add_kernel: "kernel x \ kernel y \ kernel (x + y)" by (metis add_co_reflexive add_dense add_up_closed co_reflexive_transitive kernel_co_reflexive_idempotent_up_closed idempotent_transitive_dense) -- {* add-distributive contact *} lemma one_add_dist_contact: "add_dist_contact 1" by (metis add_dist_contact_def one_add_distributive one_contact) -- {* meet-distributive kernel *} lemma zero_meet_dist_kernel: "meet_dist_kernel 0" by (metis meet_dist_kernel_def zero_kernel zero_meet_distributive) lemma one_meet_dist_kernel: "meet_dist_kernel 1" by (metis meet_dist_kernel_def one_kernel one_meet_distributive) -- {* test *} lemma zero_test: "test 0" by (metis meet_commutative meet_right_zero mult_left_zero test_def) lemma one_test: "test 1" by (metis meet_left_top mult_left_one test_def) lemma add_test: "test x \ test y \ test (x + y)" by (metis (no_types, lifting) meet_commutative meet_left_dist_add mult_right_dist_add test_def) lemma meet_test: "test x \ test y \ test (x \ y)" by (smt2 test_def meet_commutative meet.add_least_upper_bound meet.add_right_isotone mult_right_sub_dist_meet_left meet.add_left_upper_bound top_right_mult_increasing antisym) -- {* co-test *} lemma one_co_test: "co_test 1" by (metis co_test_def co_total_def add_left_zero one_co_total) lemma add_co_test: "co_test x \ co_test y \ co_test (x + y)" by (smt2 co_test_contact co_test_def contact_def add_associative add_commutative add_left_zero mult_left_one mult_right_dist_add) -- {* vector *} lemma zero_vector: "vector 0" by (metis mult_left_zero vector_def) lemma top_vector: "vector T" by (metis top_mult_top vector_def) lemma add_vector: "vector x \ vector y \ vector (x + y)" by (metis mult_right_dist_add vector_def) lemma meet_vector: "vector x \ vector y \ vector (x \ y)" by (metis antisym meet.add_least_upper_bound mult_right_sub_dist_meet_left mult_right_sub_dist_meet_right top_right_mult_increasing vector_def) lemma comp_vector: "vector y \ vector (x ; y)" by (metis antisym_conv mult_semi_associative top_right_mult_increasing vector_def) end class lattice_ordered_pre_left_semiring_1 = non_associative_left_semiring + bounded_distributive_lattice + assumes mult_associative_one: "x ; (y ; z) = (x ; (y ; 1)) ; z" assumes mult_right_dist_meet_one: "(x ; 1 \ y ; 1) ; z = x ; z \ y ; z" begin subclass pre_left_semiring apply unfold_locales apply (metis mult_associative_one mult_left_isotone mult_right_isotone mult_sub_right_one) done subclass lattice_ordered_pre_left_semiring .. lemma mult_zero_associative: "x ; 0 ; y = x ; 0" by (smt mult_left_zero mult_associative_one) lemma mult_zero_add_one_dist: "(x ; 0 + 1) ; z = x ; 0 + z" by (metis mult_left_one mult_right_dist_add mult_zero_associative) lemma mult_zero_add_dist: "(x ; 0 + y) ; z = x ; 0 + y ; z" by (metis mult_right_dist_add mult_zero_associative) lemma vector_zero_meet_one_comp: "(x ; 0 \ 1) ; y = x ; 0 \ y" by (metis mult_left_one mult_right_dist_meet_one mult_zero_associative) -- {* CPCP Theorem 5 / Figure 2: relations between properties *} lemma co_test_meet_distributive: "co_test x \ meet_distributive x" by (metis add_left_dist_meet co_test_def meet_distributive_def mult_zero_add_one_dist) lemma co_test_add_distributive: "co_test x \ add_distributive x" by (smt2 add_associative add_commutative add_distributive_def add_left_upper_bound co_test_def less_eq_def mult_zero_add_one_dist) lemma co_test_add_dist_contact: "co_test x \ add_dist_contact x" by (metis co_test_add_distributive add_dist_contact_def co_test_contact) -- {* CPCP Theorem 8 / Figure 3: closure properties *} -- {* co-test *} lemma meet_co_test: "co_test x \ co_test y \ co_test (x \ y)" by (smt2 add_commutative add_left_dist_meet co_test_def co_test_up_closed up_closed_def mult_right_dist_meet_one) lemma comp_co_test: "co_test x \ co_test y \ co_test (x ; y)" by (metis add_associative co_test_def mult_zero_add_dist mult_zero_add_one_dist) end class lattice_ordered_pre_left_semiring_2 = lattice_ordered_pre_left_semiring + assumes mult_sub_associative_one: "x ; (y ; z) \ (x ; (y ; 1)) ; z" assumes mult_right_dist_meet_one_sub: "x ; z \ y ; z \ (x ; 1 \ y ; 1) ; z" begin subclass lattice_ordered_pre_left_semiring_1 apply unfold_locales apply (metis meet.eq_iff mult_sub_associative_one mult_sup_associative_one) apply (metis meet.antisym_conv mult_one_associative mult_right_dist_meet_one_sub mult_right_sub_dist_meet) done end class multirelation_algebra_1 = lattice_ordered_pre_left_semiring + assumes mult_left_top: "T ; x = T" begin -- {* CPCP Theorem 8 / Figure 3: closure properties *} lemma top_add_distributive: "add_distributive T" by (metis add_distributive_def add_left_top mult_left_top) lemma top_meet_distributive: "meet_distributive T" by (metis meet_idempotent meet_distributive_def mult_left_top) lemma top_add_dist_contact: "add_dist_contact T" by (metis add_dist_contact_def top_add_distributive top_contact) lemma top_co_test: "co_test T" by (metis co_test_def add_left_top mult_left_top) end -- {* M1-algebra *} class multirelation_algebra_2 = multirelation_algebra_1 + lattice_ordered_pre_left_semiring_2 begin lemma mult_top_associative: "x ; T ; y = x ; T" by (metis mult_left_top mult_associative_one) lemma vector_meet_one_comp: "(x ; T \ 1) ; y = x ; T \ y" by (metis mult_left_one mult_left_top mult_associative_one mult_right_dist_meet_one) lemma vector_left_annihilator: "vector x \ x ; y = x" by (metis mult_left_top vector_def mult_associative_one) -- {* properties *} lemma test_comp_meet: "test x \ test y \ x ; y = x \ y" by (smt2 meet_associative meet_commutative meet_idempotent test_def vector_meet_one_comp) -- {* CPCP Theorem 5 / Figure 2: relations between properties *} lemma test_add_distributive: "test x \ add_distributive x" by (metis add_distributive_def meet_left_dist_add test_def vector_meet_one_comp) lemma test_meet_distributive: "test x \ meet_distributive x" by (smt2 meet.less_eq_def meet_associative meet_commutative meet_distributive_def meet.add_right_upper_bound mult_left_one test_def vector_meet_one_comp) lemma test_meet_dist_kernel: "test x \ meet_dist_kernel x" by (metis kernel_co_reflexive_idempotent_up_closed meet_associative meet_dist_kernel_def meet_idempotent test_co_reflexive test_def test_up_closed idempotent_def vector_meet_one_comp test_meet_distributive) lemma vector_idempotent: "vector x \ idempotent x" by (metis idempotent_def vector_left_annihilator) lemma vector_add_distributive: "vector x \ add_distributive x" by (metis add_distributive_def add_idempotent vector_left_annihilator) lemma vector_meet_distributive: "vector x \ meet_distributive x" by (metis meet_distributive_def meet_idempotent vector_left_annihilator) lemma vector_co_vector: "vector x \ co_vector x" by (metis co_vector_def vector_def mult_zero_associative vector_left_annihilator) -- {* CPCP Theorem 8 / Figure 3: closure properties *} -- {* test *} lemma comp_test: "test x \ test y \ test (x ; y)" by (metis meet_associative meet_distributive_def meet.add_right_zero test_def test_up_closed up_closed_def mult_associative_one test_meet_distributive) end class dual = fixes dual :: "'a \ 'a" ("_\<^sup>d" [100] 100) class multirelation_algebra_3 = lattice_ordered_pre_left_semiring + dual + assumes dual_involutive: "x\<^sup>d\<^sup>d = x" assumes dual_dist_add: "(x + y)\<^sup>d = x\<^sup>d \ y\<^sup>d" assumes dual_one: "1\<^sup>d = 1" begin lemma dual_dist_meet: "(x \ y)\<^sup>d = x\<^sup>d + y\<^sup>d" by (metis dual_dist_add dual_involutive) lemma dual_antitone: "x \ y \ y\<^sup>d \ x\<^sup>d" by (metis dual_dist_meet add_left_divisibility meet.add_left_divisibility) lemma dual_zero: "0\<^sup>d = T" by (metis dual_dist_meet add_right_top dual_involutive meet_left_zero) lemma dual_top: "T\<^sup>d = 0" by (metis dual_zero dual_involutive) -- {* CPCP Theorem 8 / Figure 3: closure properties *} lemma reflexive_co_reflexive_dual: "reflexive x \ co_reflexive (x\<^sup>d)" by (metis co_reflexive_def dual_antitone dual_involutive dual_one reflexive_def) end class multirelation_algebra_4 = multirelation_algebra_3 + assumes dual_sub_dist_comp: "(x ; y)\<^sup>d \ x\<^sup>d ; y\<^sup>d" begin subclass multirelation_algebra_1 apply unfold_locales apply (metis dual_zero dual_sub_dist_comp dual_involutive meet.less_eq_def meet_commutative meet_left_top mult_left_zero) done lemma dual_sub_dist_comp_one: "(x ; y)\<^sup>d \ (x ; 1)\<^sup>d ; y\<^sup>d" by (metis dual_sub_dist_comp mult_one_associative) -- {* CPCP Theorem 8 / Figure 3: closure properties *} lemma co_total_total_dual: "co_total x \ total (x\<^sup>d)" by (metis co_total_def dual_sub_dist_comp dual_zero meet.less_eq_def meet_commutative meet_left_top total_def) lemma transitive_dense_dual: "transitive x \ dense (x\<^sup>d)" by (metis dual_antitone dual_sub_dist_comp order_trans transitive_def dense_def) end -- {* M2-algebra *} class multirelation_algebra_5 = multirelation_algebra_3 + assumes dual_dist_comp_one: "(x ; y)\<^sup>d = (x ; 1)\<^sup>d ; y\<^sup>d" begin subclass multirelation_algebra_4 apply unfold_locales apply (metis dual_antitone mult_sub_right_one mult_left_isotone dual_dist_comp_one) done lemma strong_up_closed: "x ; 1 \ x \ x\<^sup>d ; y\<^sup>d \ (x ; y)\<^sup>d" by (metis dual_dist_comp_one eq_iff mult_sub_right_one) lemma strong_up_closed_2: "up_closed x \ (x ; y)\<^sup>d = x\<^sup>d ; y\<^sup>d" by (metis dual_sub_dist_comp eq_iff strong_up_closed up_closed_def) subclass lattice_ordered_pre_left_semiring_2 apply unfold_locales apply (smt2 comp_up_closed dual_antitone dual_dist_comp_one dual_involutive dual_one mult_left_one mult_one_associative mult_semi_associative up_closed_def strong_up_closed_2) apply (smt2 dual_dist_comp_one dual_dist_meet dual_involutive eq_refl mult_one_associative mult_right_dist_add) done -- {* CPCP Theorem 6 *} subclass multirelation_algebra_2 .. -- {* CPCP Theorem 8 / Figure 3: closure properties *} -- {* up-closed *} lemma up_closed_dual: "up_closed x \ up_closed (x\<^sup>d)" by (metis dual_involutive dual_one up_closed_def strong_up_closed_2) -- {* contact *} lemma contact_kernel_dual: "contact x \ kernel (x\<^sup>d)" by (metis contact_def contact_up_closed dual_dist_add dual_involutive dual_one kernel_def kernel_up_closed up_closed_def strong_up_closed_2) -- {* add-distributive contact *} lemma add_dist_contact_meet_dist_kernel_dual: "add_dist_contact x \ meet_dist_kernel (x\<^sup>d)" proof assume 1: "add_dist_contact x" hence 2: "up_closed x" by (metis add_dist_contact_def contact_up_closed) have "add_distributive x" using 1 by (metis add_dist_contact_def) hence "meet_distributive (x\<^sup>d)" using 2 by (smt2 meet_distributive_def add_distributive_def dual_dist_add dual_involutive strong_up_closed_2) thus "meet_dist_kernel (x\<^sup>d)" using 1 by (metis contact_kernel_dual add_dist_contact_def meet_dist_kernel_def) next assume 3: "meet_dist_kernel (x\<^sup>d)" hence 2: "up_closed (x\<^sup>d)" by (metis kernel_up_closed meet_dist_kernel_def) have "meet_distributive (x\<^sup>d)" using 3 by (metis meet_dist_kernel_def) hence "add_distributive (x\<^sup>d\<^sup>d)" using 2 by (smt2 meet_distributive_def add_distributive_def dual_dist_add dual_involutive strong_up_closed_2) thus "add_dist_contact x" using 3 by (metis contact_kernel_dual add_dist_contact_def meet_dist_kernel_def dual_involutive) qed -- {* test *} lemma test_co_test_dual: "test x \ co_test (x\<^sup>d)" by (smt2 co_test_def co_test_up_closed dual_dist_meet dual_involutive dual_one dual_top test_def test_up_closed strong_up_closed_2) -- {* vector *} lemma vector_dual: "vector x \ vector (x\<^sup>d)" by (metis dual_dist_comp_one comp_vector dual_involutive dual_top vector_def zero_vector) end class multirelation_algebra_6 = multirelation_algebra_4 + assumes dual_sub_dist_comp_one: "(x ; 1)\<^sup>d ; y\<^sup>d \ (x ; y)\<^sup>d" begin subclass multirelation_algebra_5 apply unfold_locales apply (metis dual_sub_dist_comp dual_sub_dist_comp_one meet.eq_iff mult_one_associative) done lemma "dense x \ co_reflexive x \ up_closed x" nitpick [expect=genuine] oops lemma "x ; T \ y ; z \ (x ; T \ y) ; z" nitpick [expect=genuine,card=8] oops end -- {* M3-algebra *} class up_closed_multirelation_algebra = multirelation_algebra_3 + assumes dual_dist_comp: "(x ; y)\<^sup>d = x\<^sup>d ; y\<^sup>d" begin lemma mult_right_dist_meet: "(x \ y) ; z = x ; z \ y ; z" by (metis dual_dist_add dual_dist_comp dual_involutive mult_right_dist_add) -- {* CPCP Theorem 7 *} subclass idempotent_left_semiring apply unfold_locales apply (metis antisym dual_antitone dual_dist_comp dual_involutive mult_semi_associative) apply (metis mult_left_one) apply (metis dual_dist_add dual_dist_comp dual_involutive dual_one less_eq_def meet_absorb mult_sub_right_one) done subclass multirelation_algebra_6 apply unfold_locales apply (metis dual_dist_comp eq_iff) apply (metis dual_dist_comp eq_iff mult_right_one) done lemma vector_meet_comp: "(x ; T \ y) ; z = x ; T \ y ; z" by (metis mult_associative mult_left_top mult_right_dist_meet) lemma vector_zero_meet_comp: "(x ; 0 \ y) ; z = x ; 0 \ y ; z" by (metis mult_associative mult_left_zero mult_right_dist_meet) -- {* CPCP Theorem 8 / Figure 3: closure properties *} -- {* total *} lemma meet_total: "total x \ total y \ total (x \ y)" by (metis meet_left_top total_def mult_right_dist_meet) lemma comp_total: "total x \ total y \ total (x ; y)" by (metis mult_associative total_def) lemma total_co_total_dual: "total x \ co_total (x\<^sup>d)" by (metis co_total_def dual_dist_comp dual_involutive dual_top total_def) -- {* dense *} lemma transitive_iff_dense_dual: "transitive x \ dense (x\<^sup>d)" by (metis dense_def dual_antitone dual_dist_comp dual_involutive transitive_def) -- {* idempotent *} lemma idempotent_dual: "idempotent x \ idempotent (x\<^sup>d)" by (metis dual_involutive idempotent_transitive_dense transitive_iff_dense_dual) -- {* add-distributive *} lemma comp_add_distributive: "add_distributive x \ add_distributive y \ add_distributive (x ; y)" by (metis add_distributive_def mult_associative) lemma add_meet_distributive_dual: "add_distributive x \ meet_distributive (x\<^sup>d)" by (metis (no_types, hide_lams) add_distributive_def dual_dist_add dual_dist_comp dual_involutive meet_distributive_def) -- {* meet-distributive *} lemma meet_meet_distributive: "meet_distributive x \ meet_distributive y \ meet_distributive (x \ y)" by (smt2 meet_distributive_def meet_associative meet_commutative mult_right_dist_meet) lemma comp_meet_distributive: "meet_distributive x \ meet_distributive y \ meet_distributive (x ; y)" by (metis meet_distributive_def mult_associative) lemma "co_total x \ transitive x \ up_closed x \ co_reflexive x" nitpick [expect=genuine] oops lemma "total x \ dense x \ up_closed x \ reflexive x" nitpick [expect=genuine] oops lemma "x ; T \ x\<^sup>d ; 0 = 0" nitpick [expect=genuine] oops end class multirelation_algebra_7 = multirelation_algebra_4 + assumes vector_meet_comp: "(x ; T \ y) ; z = x ; T \ y ; z" begin lemma vector_zero_meet_comp: "(x ; 0 \ y) ; z = x ; 0 \ y ; z" by (metis vector_def comp_vector vector_meet_comp zero_vector) lemma test_add_distributive: "test x \ add_distributive x" by (metis add_distributive_def meet_left_dist_add mult_left_one test_def vector_meet_comp) lemma test_meet_distributive: "test x \ meet_distributive x" by (smt2 meet.less_eq_def meet_associative meet_commutative meet_distributive_def meet.add_right_upper_bound mult_left_one test_def vector_meet_comp) lemma test_meet_dist_kernel: "test x \ meet_dist_kernel x" by (metis kernel_co_reflexive_idempotent_up_closed meet_associative meet_dist_kernel_def meet_idempotent mult_left_one test_co_reflexive test_def test_up_closed idempotent_def vector_meet_comp test_meet_distributive) lemma co_test_meet_distributive: "co_test x \ meet_distributive x" proof assume "co_test x" hence "x = x ; 0 + 1" by (metis co_test_def) hence "\y z . x ; y \ x ; z = x ; (y \ z)" by (metis mult_left_one mult_left_top mult_right_dist_add meet.add_right_zero vector_zero_meet_comp add_left_dist_meet) thus "meet_distributive x" by (metis meet_distributive_def) qed lemma co_test_add_distributive: "co_test x \ add_distributive x" proof assume "co_test x" hence 1: "x = x ; 0 + 1" by (metis co_test_def) hence "\y z . x ; (y + z) = x ; y + x ; z" by (metis add_associative add_commutative add_idempotent mult_left_one mult_left_top mult_right_dist_add meet.add_right_zero vector_zero_meet_comp) thus "add_distributive x" by (metis add_distributive_def) qed lemma co_test_add_dist_contact: "co_test x \ add_dist_contact x" by (metis co_test_add_distributive add_dist_contact_def co_test_contact) end (* end theory NAlgebra imports LatticeOrderedSemiring begin *) class C_left_n_algebra = bounded_idempotent_left_semiring + bounded_distributive_lattice + n + L begin abbreviation C :: "'a \ 'a" where "C x \ n(L) ; T \ x" -- {* AACP Theorem 3.38 *} lemma C_isotone: "x \ y \ C x \ C y" by (metis meet.add_right_isotone) -- {* AACP Theorem 3.40 *} lemma C_decreasing: "C x \ x" by (metis meet.add_right_upper_bound) end class left_n_algebra = C_left_n_algebra + assumes n_dist_n_add : "n(x) + n(y) = n(n(x) ; T + y)" assumes n_export : "n(x) ; n(y) = n(n(x) ; y)" assumes n_left_upper_bound : "n(x) \ n(x + y)" assumes n_nL_meet_L_nL0 : "n(L) ; x = (x \ L) + n(L ; 0) ; x" assumes n_n_L_split_n_n_L_L : "x ; n(y) ; L = x ; 0 + n(x ; n(y) ; L) ; L" assumes n_sub_nL : "n(x) \ n(L)" assumes n_L_decreasing : "n(x) ; L \ x" assumes n_L_T_meet_mult_combined: "C (x ; y) ; z \ C x ; y ; C z" assumes n_n_top_split_n_top : "x ; n(y) ; T \ x ; 0 + n(x ; y) ; T" assumes n_top_meet_L_below_L : "x ; T ; y \ L \ x ; L ; y" begin subclass lattice_ordered_pre_left_semiring .. lemma n_L_T_meet_mult_below: "C (x ; y) \ C x ; y" proof - have "C (x ; y) \ C x ; y ; C 1" by (metis mult_right_one mult_associative n_L_T_meet_mult_combined) also have "... \ C x ; y" by (metis mult_right_one C_decreasing mult_right_isotone mult_associative) finally show ?thesis . qed -- {* AACP Theorem 3.41 *} lemma n_L_T_meet_mult_propagate: "C x ; y \ x ; C y" proof - have "C x ; y \ C x ; 1 ; C y" by (metis mult_right_one mult_associative n_L_T_meet_mult_combined mult_right_one) also have "... \ x ; C y" by (metis mult_right_one mult_associative mult_right_one C_decreasing mult_left_isotone) finally show ?thesis . qed -- {* AACP Theorem 3.43 *} lemma C_n_mult_closed: "C (n(x) ; y) = n(x) ; y" by (metis meet.less_eq_def mult_isotone n_sub_nL top_greatest) -- {* AACP Theorem 3.40 *} lemma meet_L_below_C: "x \ L \ C x" by (metis add_left_upper_bound meet.add_left_isotone meet_commutative meet_left_top n_nL_meet_L_nL0) -- {* AACP Theorem 3.42 *} lemma n_L_T_meet_mult: "C (x ; y) = C x ; y" apply (rule antisym) apply (rule n_L_T_meet_mult_below) apply (metis meet.add_least_upper_bound meet.add_left_upper_bound mult_associative mult_isotone mult_right_sub_dist_meet_right top_greatest top_mult_top) done -- {* AACP Theorem 3.42 *} lemma C_mult_propagate: "C x ; y = C x ; C y" by (smt2 C_n_mult_closed meet.less_eq_def meet_associative meet_less_eq_def mult_left_sub_dist_meet_right n_L_T_meet_mult_propagate) -- {* AACP Theorem 3.32 *} lemma meet_L_below_n_L: "x \ L \ n(L) ; x" by (metis add_left_divisibility n_nL_meet_L_nL0) -- {* AACP Theorem 3.27 *} lemma n_vector_meet_L: "x ; T \ L \ x ; L" by (metis mult_right_one n_top_meet_L_below_L) lemma n_right_upper_bound: "n(x) \ n(y + x)" by (metis add_commutative n_left_upper_bound) -- {* AACP Theorem 3.1 *} lemma n_isotone: "x \ y \ n(x) \ n(y)" by (metis less_eq_def n_left_upper_bound) lemma n_add_left_zero: "n(0) + n(x) = n(x)" by (metis less_eq_def n_isotone zero_least) -- {* AACP Theorem 3.13 *} lemma n_mult_right_zero_L: "n(x) ; 0 \ L" by (metis dual_order.trans mult_isotone n_L_decreasing n_sub_nL zero_least) lemma n_add_left_top: "n(T) + n(x) = n(T)" by (metis add_commutative add_right_top_1 n_dist_n_add) -- {* AACP Theorem 3.18 *} lemma n_n_L: "n(n(x) ; L) = n(x)" by (metis add_commutative add_right_zero less_eq_def n_dist_n_add n_export n_sub_nL n_add_left_top n_add_left_zero) lemma n_mult_transitive : "n(x) ; n(x) \ n(x)" by (metis mult_right_isotone n_export n_sub_nL n_n_L) lemma n_mult_left_absorb_add_sub: "n(x) ; (n(x) + n(y)) \ n(x)" by (metis add_associative less_eq_def mult_left_sub_dist_add_left n_export n_sub_nL n_n_L) -- {* AACP Theorem 3.21 *} lemma n_mult_left_lower_bound: "n(x) ; n(y) \ n(x)" by (metis mult_right_isotone n_export n_sub_nL n_n_L) -- {* AACP Theorem 3.20 *} lemma n_mult_left_zero: "n(0) ; n(x) = n(0)" by (metis add_commutative less_eq_def n_export n_add_left_zero n_mult_left_lower_bound) lemma n_mult_right_one: "n(x) ; n(T) = n(x)" by (metis add_left_upper_bound add_right_zero meet.eq_iff n_dist_n_add n_export n_mult_left_lower_bound) lemma n_L_increasing: "n(x) \ n(n(x) ; L)" by (metis order_refl n_n_L) -- {* AACP Theorem 3.2 *} lemma n_galois: "n(x) \ n(y) \ n(x) ; L \ y" by (metis mult_left_isotone n_L_decreasing n_L_increasing n_isotone order_trans) lemma n_add_n_top: "n(x + n(x) ; T) = n(x)" by (metis add_commutative add_idempotent n_dist_n_add) -- {* AACP Theorem 3.6 *} lemma n_L_below_nL_top: "L \ n(L) ; T" by (metis meet.add_right_zero meet_L_below_C meet_left_top) -- {* AACP Theorem 3.4 *} lemma n_less_eq_char_n: "x \ y \ x \ y + L \ C x \ y + n(y) ; T" proof assume "x \ y" thus "x \ y + L \ C x \ y + n(y) ; T" by (metis add_left_upper_bound dual_order.trans meet.add_right_upper_bound) next assume 1: "x \ y + L \ C x \ y + n(y) ; T" hence "x \ y + (x \ L)" by (metis meet_less_eq_def add_left_dist_meet add_right_upper_bound meet.add_left_isotone) also have "... \ y + C x" by (metis add_right_isotone less_eq_def meet_left_dist_add n_L_below_nL_top meet_commutative) also have "... \ y + n(y) ; T" using 1 by (metis add_least_upper_bound add_left_upper_bound) finally have "x \ y + (L \ n(y) ; T)" using 1 by (metis meet.add_least_upper_bound add_left_dist_meet) thus "x \ y" by (metis add_idempotent add_least_upper_bound n_vector_meet_L less_eq_def meet_commutative n_L_decreasing) qed -- {* AACP Theorem 3.31 *} lemma n_L_decreasing_meet_L: "n(x) ; L \ x \ L" by (metis meet.add_least_upper_bound n_L_decreasing n_sub_nL n_galois) -- {* AACP Theorem 3.5 *} lemma n_zero_L_zero: "n(0) ; L = 0" by (metis antisym n_L_decreasing zero_least) lemma n_L_top_below_L: "L ; T \ L" proof - have "n(L ; 0) ; L ; T \ L ; 0" by (metis mult_associative mult_left_isotone n_L_decreasing vector_def zero_vector) hence "n(L ; 0) ; L ; T \ L" by (metis order_trans zero_right_mult_decreasing) hence "n(L) ; L ; T \ L" by (metis add_least_upper_bound meet.add_right_upper_bound mult_associative n_nL_meet_L_nL0) thus "L ; T \ L" by (metis add_least_upper_bound eq_iff meet_idempotent n_nL_meet_L_nL0 top_right_mult_increasing) qed -- {* AACP Theorem 3.9 *} lemma n_L_top_L: "L ; T = L" by (metis antisym n_L_top_below_L top_right_mult_increasing) -- {* AACP Theorem 3.10 *} lemma n_L_below_L: "L ; x \ L" by (metis add_right_top n_L_top_below_L mult_left_sub_dist_add_left order_trans) -- {* AACP Theorem 3.7 *} lemma n_nL_nT: "n(L) = n(T)" by (metis antisym n_isotone n_sub_nL top_greatest) -- {* AACP Theorem 3.8 *} lemma n_L_L: "n(L) ; L = L" by (metis meet.less_eq_def meet_absorb meet_idempotent n_L_decreasing n_nL_meet_L_nL0) lemma n_top_L: "n(T) ; L = L" by (metis n_L_L n_nL_nT) -- {* AACP Theorem 3.23 *} lemma n_n_L_split_n_L: "x ; n(y) ; L \ x ; 0 + n(x ; y) ; L" by (metis n_n_L_split_n_n_L_L n_L_decreasing mult_associative mult_left_isotone mult_right_isotone n_isotone add_right_isotone) -- {* AACP Theorem 3.12 *} lemma n_L_split_n_L_L: "x ; L = x ; 0 + n(x ; L) ; L" apply (rule antisym) apply (metis mult_associative n_n_L_split_n_L n_L_L) apply (metis add_least_upper_bound mult_right_isotone n_L_decreasing zero_least) done -- {* AACP Theorem 3.11 *} lemma n_L_split_L: "x ; L \ x ; 0 + L" by (metis add_commutative add_left_isotone meet.add_least_upper_bound n_L_decreasing_meet_L n_L_split_n_L_L) -- {* AACP Theorem 3.24 *} lemma n_split_top: "x ; n(y) ; T \ x ; y + n(x ; y) ; T" proof - have "x ; 0 + n(x ; y) ; T \ x ; y + n(x ; y) ; T" by (metis add_left_isotone mult_right_isotone zero_least) thus ?thesis by (smt n_n_top_split_n_top order_trans) qed -- {* AACP Theorem 3.9 *} lemma n_L_L_L: "L ; L = L" by (metis antisym n_vector_meet_L n_L_below_L meet_idempotent n_L_top_L) -- {* AACP Theorem 3.9 *} lemma n_L_top_L_L: "L ; T ; L = L" by (metis n_L_L_L n_L_top_L) -- {* AACP Theorem 3.19 *} lemma n_n_nL: "n(x) = n(x) ; n(L)" by (metis n_export n_n_L) lemma n_L_mult_idempotent: "n(L) ; n(L) = n(L)" by (metis n_n_nL) -- {* AACP Theorem 3.22 *} lemma n_n_L_n: "n(x ; n(y) ; L) \ n(x ; y)" by (metis mult_associative mult_right_isotone n_L_decreasing n_isotone) -- {* AACP Theorem 3.3 *} lemma n_less_eq_char: "x \ y \ x \ y + L \ x \ y + n(y) ; T" by (smt add_absorb add_associative add_idempotent n_less_eq_char_n meet_less_eq_def meet_left_dist_add n_add_n_top) -- {* AACP Theorem 3.28 *} lemma n_top_meet_L_split_L: "x ; T ; y \ L \ x ; 0 + L ; y" proof - have "x ; T ; y \ L \ x ; 0 + n(x ; L) ; L ; y" by (smt n_top_meet_L_below_L mult_associative n_L_L_L n_L_split_n_L_L mult_right_dist_add mult_left_zero) also have "... \ x ; 0 + x ; L ; y" by (metis add_right_isotone mult_right_sub_dist_add_right n_L_split_n_L_L) also have "... \ x ; 0 + (x ; 0 + L) ; y" by (metis add_right_isotone mult_left_isotone n_L_split_L) also have "... = x ; 0 + x ; 0 ; y + L ; y" by (metis add_associative mult_right_dist_add) also have "... = x ; 0 + L ; y" by (metis add_idempotent mult_associative mult_left_zero) finally show ?thesis . qed -- {* AACP Theorem 3.29 *} lemma n_top_meet_L_L_meet_L: "x ; T ; y \ L = x ; L ; y \ L" apply (rule antisym) apply (metis meet.add_least_upper_bound meet.add_right_upper_bound n_top_meet_L_below_L) apply (metis meet.add_left_isotone mult_left_isotone mult_right_isotone top_greatest) done lemma n_n_top_below_n_L: "n(x ; T) \ n(x ; L)" by (metis n_vector_meet_L n_L_decreasing_meet_L n_galois order_trans) -- {* AACP Theorem 3.14 *} lemma n_n_top_n_L: "n(x ; T) = n(x ; L)" by (metis antisym mult_right_isotone n_isotone n_n_top_below_n_L top_greatest) -- {* AACP Theorem 3.30 *} lemma n_meet_L_0_below_0_meet_L: "(x \ L) ; 0 \ x ; 0 \ L" by (metis meet.add_least_upper_bound mult_isotone order.refl zero_right_mult_decreasing) -- {* AACP Theorem 3.15 *} lemma n_n_L_below_L: "n(x) ; L \ x ; L" by (metis mult_associative mult_left_isotone n_L_L_L n_L_decreasing) lemma n_n_L_below_n_L_L: "n(x) ; L \ n(x ; L) ; L" by (metis n_n_L_below_L mult_left_isotone n_galois) -- {* AACP Theorem 3.16 *} lemma n_below_n_L: "n(x) \ n(x ; L)" by (metis n_n_L_below_L n_galois) -- {* AACP Theorem 3.17 *} lemma n_below_n_L_mult: "n(x) \ n(L) ; n(x)" by (metis n_export order_trans meet_L_below_n_L n_L_decreasing_meet_L n_isotone n_n_L) -- {* AACP Theorem 3.33 *} lemma n_meet_L_below: "n(x) \ L \ x" by (metis meet.add_left_isotone n_L_decreasing n_vector_meet_L order_trans top_right_mult_increasing) -- {* AACP Theorem 3.35 *} lemma n_meet_L_top_below_n_L: "(n(x) \ L) ; T \ n(x) ; L" proof - have "(n(x) \ L) ; T \ n(x) ; T \ L ; T" by (metis meet.add_least_upper_bound meet.add_left_upper_bound meet_commutative mult_left_isotone) thus ?thesis by (metis n_L_top_L n_vector_meet_L order_trans) qed -- {* AACP Theorem 3.34 *} lemma n_meet_L_top_below: "(n(x) \ L) ; T \ x" by (metis n_L_decreasing n_meet_L_top_below_n_L order_trans) -- {* AACP Theorem 3.36 *} lemma n_n_meet_L: "n(x) = n(x \ L)" by (metis add_absorb add_commutative less_eq_def n_L_decreasing_meet_L n_n_L n_right_upper_bound) lemma n_T_below_n_meet: "n(x) ; T = n(C x) ; T" by (metis C_n_mult_closed meet_associative meet_commutative n_L_L n_n_meet_L) -- {* AACP Theorem 3.44 *} lemma n_C: "n(C x) = n(x)" by (metis n_T_below_n_meet n_export n_mult_right_one) -- {* AACP Theorem 3.37 *} lemma n_T_meet_L: "n(x) ; T \ L = n(x) ; L" by (metis antisym_conv n_L_decreasing_meet_L n_n_L n_n_top_n_L n_vector_meet_L) -- {* AACP Theorem 3.39 *} lemma n_L_top_meet_L: "C L = L" by (metis meet_less_eq_def n_L_below_nL_top meet_commutative) end class n_algebra = left_n_algebra + idempotent_left_zero_semiring begin (* independence of axioms, checked in n_algebra without the respective axiom: lemma n_dist_n_add : "n(x) + n(y) = n(n(x) ; T + y)" nitpick [expect=genuine,card=5] oops lemma n_export : "n(x) ; n(y) = n(n(x) ; y)" nitpick [expect=genuine,card=4] oops lemma n_left_upper_bound : "n(x) \ n(x + y)" nitpick [expect=genuine,card=5] oops lemma n_nL_meet_L_nL0 : "n(L) ; x = (x \ L) + n(L ; 0) ; x" nitpick [expect=genuine,card=2] oops lemma n_n_L_split_n_n_L_L : "x ; n(y) ; L = x ; 0 + n(x ; n(y) ; L) ; L" nitpick [expect=genuine,card=6] oops lemma n_sub_nL : "n(x) \ n(L)" nitpick [expect=genuine,card=2] oops lemma n_L_decreasing : "n(x) ; L \ x" nitpick [expect=genuine,card=3] oops lemma n_L_T_meet_mult_combined: "C (x ; y) ; z \ C x ; y ; C z" nitpick [expect=genuine,card=4] oops lemma n_n_top_split_n_top : "x ; n(y) ; T \ x ; 0 + n(x ; y) ; T" nitpick [expect=genuine,card=4] oops lemma n_top_meet_L_below_L : "x ; T ; y \ L \ x ; L ; y" nitpick [expect=genuine,card=5] oops *) -- {* AACP Theorem 3.25 *} lemma n_top_split_0: "n(x) ; T ; y \ x ; y + n(x ; 0) ; T" proof - have 1: "n(x) ; T ; y \ L \ x ; y" by (metis mult_left_isotone n_L_decreasing n_top_meet_L_below_L order_trans) have "n(x) ; T ; y = n(x) ; n(L) ; T ; y" by (metis n_export n_n_L) also have "... = n(x) ; ((T ; y \ L) + n(L ; 0) ; T ; y)" by (metis mult_associative n_nL_meet_L_nL0) also have "... \ n(x) ; (T ; y \ L) + n(x) ; n(L ; 0) ; T" by (metis add_right_isotone mult_associative mult_left_dist_add mult_right_isotone top_greatest) also have "... \ (n(x) ; T ; y \ L) + n(n(x) ; L ; 0) ; T" by (metis add_left_isotone dual_order.trans meet.add_least_upper_bound mult_associative mult_left_sub_dist_meet_left mult_left_sub_dist_meet_right n_export n_galois n_sub_nL) also have "... \ x ; y + n(n(x) ; L ; 0) ; T" using 1 by (metis add_left_isotone) also have "... \ x ; y + n(x ; 0) ; T" by (metis add_right_isotone mult_left_isotone n_L_decreasing n_isotone) finally show ?thesis . qed -- {* AACP Theorem 3.26 *} lemma n_top_split: "n(x) ; T ; y \ x ; y + n(x ; y) ; T" by (metis add_right_isotone add_right_zero meet.order_trans mult_left_isotone mult_left_sub_dist_add_right n_isotone n_top_split_0) lemma n_n_meet_L_n_zero: "n(x) = (n(x) \ L) + n(x ; 0)" oops lemma n_below_n_zero: "n(x) \ x + n(x ; 0)" oops lemma n_n_top_split_n_L_n_zero_top: "n(x) ; T = n(x) ; L + n(x ; 0) ; T" oops lemma n_zero: "n(0) = 0" nitpick [expect=genuine] oops lemma n_one: "n(1) = 0" nitpick [expect=genuine] oops lemma n_nL_one: "n(L) = 1" nitpick [expect=genuine] oops lemma n_nT_one: "n(T) = 1" nitpick [expect=genuine] oops lemma n_n_zero: "n(x) = n(x ; 0)" nitpick [expect=genuine] oops lemma n_dist_add: "n(x) + n(y) = n(x + y)" nitpick [expect=genuine] oops lemma n_L_split: "x ; n(y) ; L = x ; 0 + n(x ; y) ; L" nitpick [expect=genuine] oops lemma n_split: "x \ x ; 0 + n(x ; L) ; T" nitpick [expect=genuine] oops lemma n_mult_top_1: "n(x ; y) \ n(x ; n(y) ; T)" nitpick [expect=genuine] oops lemma l91_1: "n(L) ; x \ n(x ; T) ; T" nitpick [expect=genuine] oops lemma meet_domain_top: "x \ n(y) ; T = n(y) ; x" nitpick [expect=genuine] oops lemma meet_domain_2: "x \ n(y) ; T \ n(L) ; x" nitpick [expect=genuine] oops lemma n_nL_top_n_top_meet_L_top_2: "n(L) ; x ; T \ n(x ; T \ L) ; T" nitpick [expect=genuine] oops lemma n_nL_top_n_top_meet_L_top_1: "n(x ; T \ L) ; T \ n(L) ; x ; T" nitpick [expect=genuine] oops lemma l9: "x ; 0 \ L \ n(x ; L) ; L" nitpick [expect=genuine] oops lemma l18_2: "n(x ; L) ; L \ n(x) ; L" nitpick [expect=genuine] oops lemma l51_1: "n(x) ; L \ (x \ L) ; 0" nitpick [expect=genuine] oops lemma l51_2: "(x \ L) ; 0 \ n(x) ; L" nitpick [expect=genuine] oops lemma n_split_equal: "x + n(x ; L) ; T = x ; 0 + n(x ; L) ; T" nitpick [expect=genuine] oops lemma n_split_top: "x ; T \ x ; 0 + n(x ; L) ; T" nitpick [expect=genuine] oops lemma n_mult: "n(x ; n(y) ; L) = n(x ; y)" nitpick [expect=genuine] oops lemma n_mult_1: "n(x ; y) \ n(x ; n(y) ; L)" nitpick [expect=genuine] oops lemma n_mult_top: "n(x ; n(y) ; T) = n(x ; y)" nitpick [expect=genuine] oops lemma n_mult_right_upper_bound: "n(x ; y) \ n(z) \ n(x) \ n(z) \ x ; n(y) ; L \ x ; 0 + n(z) ; L" nitpick [expect=genuine] oops lemma meet_domain: "x \ n(y) ; z = n(y) ; (x \ z)" nitpick [expect=genuine] oops lemma meet_domain_1: "x \ n(y) ; z \ n(y) ; x" nitpick [expect=genuine] oops lemma meet_domain_top_3: "x \ n(y) ; T \ n(y) ; x" nitpick [expect=genuine] oops lemma n_n_top_n_top_split_n_n_top_top: "n(x) ; T + x ; n(y) ; T = x ; 0 + n(x ; n(y) ; T) ; T" nitpick [expect=genuine] oops lemma n_n_top_n_top_split_n_n_top_top_1: "x ; 0 + n(x ; n(y) ; T) ; T \ n(x) ; T + x ; n(y) ; T" nitpick [expect=genuine] oops lemma n_n_top_n_top_split_n_n_top_top_2: "n(x) ; T + x ; n(y) ; T \ x ; 0 + n(x ; n(y) ; T) ; T" nitpick [expect=genuine] oops lemma n_nL_top_n_top_meet_L_top: "n(L) ; x ; T = n(x ; T \ L) ; T" nitpick [expect=genuine] oops lemma l18: "n(x) ; L = n(x ; L) ; L" nitpick [expect=genuine] oops lemma l22: "x ; 0 \ L = n(x) ; L" nitpick [expect=genuine] oops lemma l22_1: "x ; 0 \ L = n(x ; L) ; L" nitpick [expect=genuine] oops lemma l22_2: "x \ L = n(x) ; L" nitpick [expect=genuine] oops lemma l22_3: "x \ L = n(x ; L) ; L" nitpick [expect=genuine] oops lemma l22_4: "x \ L \ n(x) ; L" nitpick [expect=genuine] oops lemma l22_5: "x ; 0 \ L \ n(x) ; L" nitpick [expect=genuine] oops lemma l23: "x ; T \ L = n(x) ; L" nitpick [expect=genuine] oops lemma l51: "n(x) ; L = (x \ L) ; 0" nitpick [expect=genuine] oops lemma l91: "x = x ; T \ n(L) ; x \ n(x) ; T" nitpick [expect=genuine] oops lemma l92: "x = x ; T \ n(L) ; x \ n(x \ L) ; T" nitpick [expect=genuine] oops lemma "x \ L \ n(x) ; T" nitpick [expect=genuine] oops lemma n_meet_comp: "n(x) \ n(y) \ n(x) ; n(y)" nitpick [expect=genuine] oops lemma n_meet_L_0_0_meet_L: "(x \ L) ; 0 = x ; 0 \ L" oops end (* end theory Recursion imports Approximation NAlgebra begin *) class n_algebra_apx = n_algebra + apx + assumes apx_def: "x \ y \ x \ y + L \ C y \ x + n(x) ; T" begin lemma apx_transitive_2: "x \ y \ y \ z \ x \ z" proof assume 1: "x \ y \ y \ z" hence "C z \ C (y + n(y) ; T)" by (metis apx_def meet.add_least_upper_bound meet.add_left_upper_bound) also have "... = C y + n(y) ; T" by (metis add_commutative add_left_dist_meet mult_right_dist_add n_add_left_top n_nL_nT) also have "... \ x + n(x) ; T + n(y) ; T" using 1 by (metis add_left_isotone apx_def) also have "... = x + n(x) ; T + n(C y) ; T" by (metis n_T_below_n_meet) also have "... \ x + n(x) ; T" using 1 by (metis add_associative add_idempotent add_right_isotone apx_def mult_left_isotone n_add_n_top n_isotone) finally show "x \ z" using 1 by (smt add_associative add_commutative apx_def less_eq_def) qed lemma apx_meet_L: "y \ x \ x \ L \ y \ L" proof assume 1: "y \ x" have "x \ L = C x \ L" by (metis meet_associative meet_commutative n_L_top_meet_L) also have "... \ (y + n(y) ; T) \ L" using 1 by (metis apx_def meet.add_left_isotone) also have "... = (y \ L) + (n(y) ; T \ L)" using 1 by (metis meet_commutative meet_left_dist_add) also have "... \ (y \ L) + n(y \ L) ; T" by (metis add_least_upper_bound n_vector_meet_L meet.add_least_upper_bound n_L_decreasing order_refl order_trans) finally show "x \ L \ y \ L" by (metis less_eq_def meet.add_right_upper_bound n_less_eq_char) qed -- {* AACP Theorem 4.1 *} subclass apx_biorder apply unfold_locales apply (metis add_absorb add_least_upper_bound add_left_upper_bound apx_def meet_commutative) apply (metis add_same_context antisym apx_def apx_meet_L relative_equality) apply (metis apx_transitive_2) done lemma add_apx_left_isotone_2: "x \ y \ x + z \ y + z" proof assume 1: "x \ y" hence 2: "x + z \ y + z + L" by (smt add_associative add_commutative add_left_isotone apx_def) have "C (y + z) \ x + n(x) ; T + C z" using 1 by (metis add_left_isotone apx_def meet_left_dist_add meet_commutative) also have "... \ x + z + n(x) ; T" by (smt2 add_associative add_commutative add_right_isotone meet.add_right_upper_bound) also have "... \ x + z + n(x + z) ; T" by (metis add_commutative add_right_isotone mult_left_isotone n_right_upper_bound) finally show "x + z \ y + z" using 2 by (metis apx_def) qed lemma mult_apx_left_isotone_2: "x \ y \ x ; z \ y ; z" proof assume 1: "x \ y" hence "x ; z \ y ; z + L ; z" by (metis apx_def mult_left_isotone mult_right_dist_add) hence 2: "x ; z \ y ; z + L" by (metis add_commutative add_left_isotone n_L_below_L order_trans) have "C (y ; z) = C y ; z" by (metis meet_commutative n_L_T_meet_mult) also have "... \ x ; z + n(x) ; T ; z" using 1 by (metis apx_def mult_left_isotone mult_right_dist_add) also have "... \ x ; z + n(x ; z) ; T" by (metis add_least_upper_bound add_left_upper_bound n_top_split) finally show "x ; z \ y ; z" using 2 by (metis apx_def) qed lemma mult_apx_right_isotone_2: "x \ y \ z ; x \ z ; y" proof assume 1: "x \ y" hence "z ; x \ z ; y + z ; L" by (metis apx_def mult_left_dist_add mult_right_isotone) also have "... \ z ; y + z ; 0 + L" by (metis add_associative add_right_isotone n_L_split_L) finally have 2: "z ; x \ z ; y + L" by (metis add_right_zero mult_left_dist_add) have "C (z ; y) \ z ; C y" by (metis meet_commutative n_L_T_meet_mult n_L_T_meet_mult_propagate) also have "... \ z ; (x + n(x) ; T)" using 1 by (metis apx_def mult_right_isotone) also have "... = z ; x + z ; n(x) ; T" by (metis mult_associative mult_left_dist_add) also have "... \ z ; x + n(z ; x) ; T" by (metis add_least_upper_bound add_left_upper_bound n_split_top) finally show "z ; x \ z ; y" using 2 by (metis apx_def) qed -- {* AACP Theorem 4.1 and Theorem 4.2 *} subclass apx_semiring apply unfold_locales apply (metis add_commutative add_left_upper_bound apx_def less_eq_def meet.add_left_upper_bound n_L_below_nL_top) apply (rule add_apx_left_isotone_2) apply (rule mult_apx_left_isotone_2) apply (rule mult_apx_right_isotone_2) done -- {* AACP Theorem 4.2 *} lemma meet_L_apx_isotone: "x \ y \ x \ L \ y \ L" by (smt add_commutative add_idempotent add_left_dist_meet apx_def apx_meet_L n_less_eq_char_n meet_commutative meet.add_right_isotone) -- {* AACP Theorem 4.2 *} lemma n_L_apx_isotone: "x \ y \ n(x) ; L \ n(y) ; L" proof assume 1: "x \ y" have "C (n(y) ; L) \ n(C y) ; L" by (metis meet.add_left_upper_bound n_T_below_n_meet n_n_L n_n_top_n_L meet_commutative) also have "... \ n(x) ; L + n(n(x) ; L) ; T" using 1 by (metis add_left_upper_bound apx_def mult_left_isotone n_add_n_top n_export n_isotone order_trans) finally show "n(x) ; L \ n(y) ; L" by (metis apx_def less_eq_def meet.add_least_upper_bound n_L_decreasing_meet_L) qed definition kappa_apx_meet :: "('a \ 'a) \ bool" where "kappa_apx_meet f \ apx.has_least_fixpoint f \ has_apx_meet (\ f) (\ f) \ \ f = \ f \ \ f" definition kappa_mu_nu :: "('a \ 'a) \ bool" where "kappa_mu_nu f \ apx.has_least_fixpoint f \ \ f = \ f + (\ f \ L)" definition nu_below_mu_nu :: "('a \ 'a) \ bool" where "nu_below_mu_nu f \ C (\ f) \ \ f + (\ f \ L) + n(\ f) ; T" definition nu_below_mu_nu_2 :: "('a \ 'a) \ bool" where "nu_below_mu_nu_2 f \ C (\ f) \ \ f + (\ f \ L) + n(\ f + (\ f \ L)) ; T" definition mu_nu_apx_nu :: "('a \ 'a) \ bool" where "mu_nu_apx_nu f \ \ f + (\ f \ L) \ \ f" definition mu_nu_apx_meet :: "('a \ 'a) \ bool" where "mu_nu_apx_meet f \ has_apx_meet (\ f) (\ f) \ \ f \ \ f = \ f + (\ f \ L)" definition apx_meet_below_nu :: "('a \ 'a) \ bool" where "apx_meet_below_nu f \ has_apx_meet (\ f) (\ f) \ \ f \ \ f \ \ f" lemma mu_below_l: "\ f \ \ f + (\ f \ L)" by (metis add_left_upper_bound) lemma l_below_nu: "has_least_fixpoint f \ has_greatest_fixpoint f \ \ f + (\ f \ L) \ \ f" by (metis add_least_upper_bound meet.add_left_upper_bound mu_below_nu) lemma n_l_nu: "has_least_fixpoint f \ has_greatest_fixpoint f \ (\ f + (\ f \ L)) \ L = \ f \ L" by (smt add_commutative add_left_dist_meet less_eq_def meet_absorb meet_associative meet_commutative mu_below_nu) lemma l_apx_mu: "\ f + (\ f \ L) \ \ f" proof - have 1: "\ f + (\ f \ L) \ \ f + L" by (metis add_right_isotone meet.add_right_upper_bound) have "C (\ f) \ \ f + (\ f \ L) + n(\ f + (\ f \ L)) ; T" by (metis add_left_upper_bound n_less_eq_char_n meet_commutative) thus ?thesis using 1 by (metis apx_def) qed -- {* AACP Theorem 4.8 implies Theorem 4.9 *} lemma nu_below_mu_nu_nu_below_mu_nu_2: "nu_below_mu_nu f \ nu_below_mu_nu_2 f" proof assume 1: "nu_below_mu_nu f" have "C (\ f) = C (C (\ f))" by (metis meet_associative meet_idempotent) also have "... \ C (\ f + (\ f \ L) + n(\ f) ; T)" using 1 by (metis calculation meet.add_least_upper_bound meet.add_left_upper_bound nu_below_mu_nu_def) also have "... = C (\ f + (\ f \ L)) + C (n(\ f) ; T)" by (metis meet_left_dist_add) also have "... = C (\ f + (\ f \ L)) + n(\ f) ; T" by (metis C_n_mult_closed) also have "... \ \ f + (\ f \ L) + n(\ f) ; T" by (metis add_left_isotone meet.add_right_upper_bound) also have "... = \ f + (\ f \ L) + n(\ f \ L) ; T" by (metis n_n_meet_L) also have "... \ \ f + (\ f \ L) + n(\ f + (\ f \ L)) ; T" by (metis add_right_isotone mult_left_isotone n_right_upper_bound) finally show "nu_below_mu_nu_2 f" by (metis nu_below_mu_nu_2_def) qed -- {* AACP Theorem 4.9 implies Theorem 4.8 *} lemma nu_below_mu_nu_2_nu_below_mu_nu: "has_least_fixpoint f \ has_greatest_fixpoint f \ nu_below_mu_nu_2 f \ nu_below_mu_nu f" proof assume 1: "has_least_fixpoint f \ has_greatest_fixpoint f \ nu_below_mu_nu_2 f" hence "C (\ f) \ \ f + (\ f \ L) + n(\ f + (\ f \ L)) ; T" by (metis nu_below_mu_nu_2_def) also have "... \ \ f + (\ f \ L) + n(\ f) ; T" using 1 by (metis add_right_isotone l_below_nu mult_left_isotone n_isotone) finally show "nu_below_mu_nu f" by (metis nu_below_mu_nu_def) qed lemma nu_below_mu_nu_equivalent: "has_least_fixpoint f \ has_greatest_fixpoint f \ (nu_below_mu_nu f \ nu_below_mu_nu_2 f)" by (metis nu_below_mu_nu_2_nu_below_mu_nu nu_below_mu_nu_nu_below_mu_nu_2) -- {* AACP Theorem 4.9 implies Theorem 4.10 *} lemma nu_below_mu_nu_2_mu_nu_apx_nu: "has_least_fixpoint f \ has_greatest_fixpoint f \ nu_below_mu_nu_2 f \ mu_nu_apx_nu f" proof assume 1: "has_least_fixpoint f \ has_greatest_fixpoint f \ nu_below_mu_nu_2 f" hence "\ f + (\ f \ L) \ \ f + L" by (metis add_commutative add_right_upper_bound l_below_nu order_trans) thus "mu_nu_apx_nu f" using 1 by (metis apx_def mu_nu_apx_nu_def nu_below_mu_nu_2_def) qed -- {* AACP Theorem 4.10 implies Theorem 4.11 *} lemma mu_nu_apx_nu_mu_nu_apx_meet: "has_least_fixpoint f \ has_greatest_fixpoint f \ mu_nu_apx_nu f \ mu_nu_apx_meet f" proof let ?l = "\ f + (\ f \ L)" assume "has_least_fixpoint f \ has_greatest_fixpoint f \ mu_nu_apx_nu f" hence "is_apx_meet (\ f) (\ f) ?l" by (smt add_apx_left_isotone add_commutative apx_meet_L is_apx_meet_def l_apx_mu less_eq_def meet.add_least_upper_bound mu_nu_apx_nu_def) thus "mu_nu_apx_meet f" by (smt apx_meet_char mu_nu_apx_meet_def) qed -- {* AACP Theorem 4.11 implies Theorem 4.12 *} lemma mu_nu_apx_meet_apx_meet_below_nu: "has_least_fixpoint f \ has_greatest_fixpoint f \ mu_nu_apx_meet f \ apx_meet_below_nu f" by (metis apx_meet_below_nu_def l_below_nu mu_nu_apx_meet_def) -- {* AACP Theorem 4.12 implies Theorem 4.9 *} lemma apx_meet_below_nu_nu_below_mu_nu_2: "apx_meet_below_nu f \ nu_below_mu_nu_2 f" proof - let ?l = "\ f + (\ f \ L)" have "\m . m \ \ f \ m \ \ f \ m \ \ f \ C (\ f) \ ?l + n(?l) ; T" proof fix m show "m \ \ f \ m \ \ f \ m \ \ f \ C (\ f) \ ?l + n(?l) ; T" proof assume 1: "m \ \ f \ m \ \ f \ m \ \ f" hence "m \ ?l" by (smt add_commutative add_left_dist_meet add_left_upper_bound apx_def meet_less_eq_def meet.add_least_upper_bound) hence "m + n(m) ; T \ ?l + n(?l) ; T" by (metis add_isotone mult_left_isotone n_isotone) thus "C (\ f) \ ?l + n(?l) ; T" using 1 by (smt apx_def order_trans) qed qed thus ?thesis by (smt apx_meet_below_nu_def apx_meet_same apx_meet_unique is_apx_meet_def nu_below_mu_nu_2_def) qed -- {* AACP Theorem 4.5 implies Theorem 4.6 *} lemma has_apx_least_fixpoint_kappa_apx_meet: "has_least_fixpoint f \ has_greatest_fixpoint f \ apx.has_least_fixpoint f \ kappa_apx_meet f" proof assume 1: "has_least_fixpoint f \ has_greatest_fixpoint f \ apx.has_least_fixpoint f" hence 2: "\w . w \ \ f \ w \ \ f \ C (\ f) \ w + n(w) ; T" by (metis apx_def meet.add_right_isotone order_trans kappa_below_nu) have "\w . w \ \ f \ w \ \ f \ w \ \ f + L" using 1 by (metis add_left_isotone apx_def mu_below_kappa order_trans) hence "\w . w \ \ f \ w \ \ f \ w \ \ f" using 2 by (metis apx_def) hence "is_apx_meet (\ f) (\ f) (\ f)" using 1 by (smt apx_meet_char is_apx_meet_def kappa_apx_below_mu kappa_apx_below_nu kappa_apx_meet_def) thus "kappa_apx_meet f" using 1 by (metis apx_meet_char kappa_apx_meet_def) qed -- {* AACP Theorem 4.6 implies Theorem 4.12 *} lemma kappa_apx_meet_apx_meet_below_nu: "has_greatest_fixpoint f \ kappa_apx_meet f \ apx_meet_below_nu f" by (metis apx_meet_below_nu_def kappa_apx_meet_def kappa_below_nu) -- {* AACP Theorem 4.12 implies Theorem 4.7 *} lemma apx_meet_below_nu_kappa_mu_nu: "has_least_fixpoint f \ has_greatest_fixpoint f \ isotone f \ apx.isotone f \ apx_meet_below_nu f \ kappa_mu_nu f" proof let ?l = "\ f + (\ f \ L)" let ?m = "\ f \ \ f" assume 1: "has_least_fixpoint f \ has_greatest_fixpoint f \ isotone f \ apx.isotone f \ apx_meet_below_nu f" hence 2: "?m = ?l" by (metis apx_meet_below_nu_nu_below_mu_nu_2 mu_nu_apx_meet_def mu_nu_apx_nu_mu_nu_apx_meet nu_below_mu_nu_2_mu_nu_apx_nu) have 3: "?l \ f(?l) + L" proof - have "?l \ \ f + L" by (metis add_right_isotone meet.add_right_upper_bound) also have "... = f(\ f) + L" using 1 by (metis is_least_fixpoint_def least_fixpoint) also have "... \ f(?l) + L" using 1 by (metis add_left_isotone add_left_upper_bound isotone_def) finally show "?l \ f(?l) + L" by metis qed have "C (f(?l)) \ ?l + n(?l) ; T" proof - have "C (f(?l)) \ C (f(\ f))" using 1 2 by (metis apx_meet_below_nu_def isotone_def meet.add_right_isotone) also have "... = C (\ f)" using 1 by (metis greatest_fixpoint is_greatest_fixpoint_def) also have "... \ ?l + n(?l) ; T" using 1 by (metis apx_meet_below_nu_nu_below_mu_nu_2 nu_below_mu_nu_2_def) finally show "C (f(?l)) \ ?l + n(?l) ; T" by metis qed hence 4: "?l \ f(?l)" using 3 by (metis apx_def) have 5: "f(?l) \ \ f" proof - have "?l \ \ f" by (metis l_apx_mu) thus "f(?l) \ \ f" using 1 by (metis apx.isotone_def is_least_fixpoint_def least_fixpoint) qed have 6: "f(?l) \ \ f" proof - have "?l \ \ f" using 1 2 by (metis apx_greatest_lower_bound apx_meet_below_nu_def apx_reflexive) thus "f(?l) \ \ f" using 1 by (metis apx.isotone_def greatest_fixpoint is_greatest_fixpoint_def) qed hence "f(?l) \ ?l" using 1 2 5 by (metis apx_greatest_lower_bound apx_meet_below_nu_def) hence 7: "f(?l) = ?l" using 4 by (metis apx_antisymmetric) have "\y . f(y) = y \ ?l \ y" proof fix y show "f(y) = y \ ?l \ y" proof assume 8: "f(y) = y" hence 9: "?l \ y + L" using 1 by (metis add_isotone is_least_fixpoint_def least_fixpoint meet.add_right_upper_bound) have "y \ \ f" using 1 8 by (metis greatest_fixpoint is_greatest_fixpoint_def) hence "C y \ ?l + n(?l) ; T" using 1 4 6 by (smt apx_meet_below_nu_nu_below_mu_nu_2 meet.add_right_isotone nu_below_mu_nu_2_def order_trans) thus "?l \ y" using 9 by (metis apx_def) qed qed thus "kappa_mu_nu f" using 1 2 7 by (smt apx.least_fixpoint_same apx.has_least_fixpoint_def apx.is_least_fixpoint_def kappa_mu_nu_def) qed -- {* AACP Theorem 4.7 implies Theorem 4.5 *} lemma kappa_mu_nu_has_apx_least_fixpoint: "kappa_mu_nu f \ apx.has_least_fixpoint f" by (metis kappa_mu_nu_def) -- {* AACP Theorem 4.8 implies Theorem 4.7 *} lemma nu_below_mu_nu_kappa_mu_nu: "has_least_fixpoint f \ has_greatest_fixpoint f \ isotone f \ apx.isotone f \ nu_below_mu_nu f \ kappa_mu_nu f" by (metis apx_meet_below_nu_kappa_mu_nu mu_nu_apx_meet_apx_meet_below_nu mu_nu_apx_nu_mu_nu_apx_meet nu_below_mu_nu_nu_below_mu_nu_2 nu_below_mu_nu_2_mu_nu_apx_nu) -- {* AACP Theorem 4.7 implies Theorem 4.8 *} lemma kappa_mu_nu_nu_below_mu_nu: "has_least_fixpoint f \ has_greatest_fixpoint f \ kappa_mu_nu f \ nu_below_mu_nu f" by (metis apx_meet_below_nu_nu_below_mu_nu_2 has_apx_least_fixpoint_kappa_apx_meet nu_below_mu_nu_2_nu_below_mu_nu kappa_apx_meet_apx_meet_below_nu kappa_mu_nu_has_apx_least_fixpoint) definition kappa_mu_nu_L :: "('a \ 'a) \ bool" where "kappa_mu_nu_L f \ apx.has_least_fixpoint f \ \ f = \ f + n(\ f) ; L" definition nu_below_mu_nu_L :: "('a \ 'a) \ bool" where "nu_below_mu_nu_L f \ C (\ f) \ \ f + n(\ f) ; T" definition mu_nu_apx_nu_L :: "('a \ 'a) \ bool" where "mu_nu_apx_nu_L f \ \ f + n(\ f) ; L \ \ f" definition mu_nu_apx_meet_L :: "('a \ 'a) \ bool" where "mu_nu_apx_meet_L f \ has_apx_meet (\ f) (\ f) \ \ f \ \ f = \ f + n(\ f) ; L" lemma n_below_l: "x + n(y) ; L \ x + (y \ L)" by (metis add_right_isotone n_L_decreasing_meet_L) lemma n_equal_l: "nu_below_mu_nu_L f \ \ f + n(\ f) ; L = \ f + (\ f \ L)" proof assume "nu_below_mu_nu_L f" hence "\ f \ L \ (\ f + n(\ f) ; T) \ L" by (metis meet_L_below_C meet.add_least_upper_bound meet.add_right_upper_bound nu_below_mu_nu_L_def order_trans) also have "... \ \ f + (n(\ f) ; T \ L)" by (metis add_left_dist_meet add_right_upper_bound meet.add_right_isotone) also have "... \ \ f + n(\ f) ; L" by (metis add_right_isotone n_vector_meet_L) finally have "\ f + (\ f \ L) \ \ f + n(\ f) ; L" by (metis add_least_upper_bound add_left_upper_bound) thus "\ f + n(\ f) ; L = \ f + (\ f \ L)" by (metis antisym n_below_l) qed -- {* AACP Theorem 4.14 implies Theorem 4.8 *} lemma nu_below_mu_nu_L_nu_below_mu_nu: "nu_below_mu_nu_L f \ nu_below_mu_nu f" by (metis add_associative add_right_top mult_left_dist_add n_equal_l nu_below_mu_nu_L_def nu_below_mu_nu_def) -- {* AACP Theorem 4.14 implies Theorem 4.13 *} lemma nu_below_mu_nu_L_kappa_mu_nu_L: "has_least_fixpoint f \ has_greatest_fixpoint f \ isotone f \ apx.isotone f \ nu_below_mu_nu_L f \ kappa_mu_nu_L f" by (metis n_equal_l nu_below_mu_nu_L_nu_below_mu_nu nu_below_mu_nu_kappa_mu_nu kappa_mu_nu_L_def kappa_mu_nu_def) -- {* AACP Theorem 4.14 implies Theorem 4.15 *} lemma nu_below_mu_nu_L_mu_nu_apx_nu_L: "has_least_fixpoint f \ has_greatest_fixpoint f \ nu_below_mu_nu_L f \ mu_nu_apx_nu_L f" by (metis mu_nu_apx_nu_L_def mu_nu_apx_nu_def n_equal_l nu_below_mu_nu_2_mu_nu_apx_nu nu_below_mu_nu_L_nu_below_mu_nu nu_below_mu_nu_nu_below_mu_nu_2) -- {* AACP Theorem 4.14 implies Theorem 4.16 *} lemma nu_below_mu_nu_L_mu_nu_apx_meet_L: "has_least_fixpoint f \ has_greatest_fixpoint f \ nu_below_mu_nu_L f \ mu_nu_apx_meet_L f" by (metis mu_nu_apx_meet_L_def mu_nu_apx_meet_def mu_nu_apx_nu_mu_nu_apx_meet n_equal_l nu_below_mu_nu_2_mu_nu_apx_nu nu_below_mu_nu_L_nu_below_mu_nu nu_below_mu_nu_nu_below_mu_nu_2) -- {* AACP Theorem 4.15 implies Theorem 4.14 *} lemma mu_nu_apx_nu_L_nu_below_mu_nu_L: "has_least_fixpoint f \ has_greatest_fixpoint f \ mu_nu_apx_nu_L f \ nu_below_mu_nu_L f" proof let ?n = "\ f + n(\ f) ; L" let ?l = "\ f + (\ f \ L)" assume 1: "has_least_fixpoint f \ has_greatest_fixpoint f \ mu_nu_apx_nu_L f" hence "C (\ f) \ ?n + n(?n) ; T" by (metis apx_def mu_nu_apx_nu_L_def) also have "... \ ?n + n(?l) ; T" by (metis add_right_isotone n_isotone mult_left_isotone n_below_l) also have "... \ ?n + n(\ f) ; T" using 1 by (metis add_right_isotone n_isotone l_below_nu mult_left_isotone) finally show "nu_below_mu_nu_L f" by (metis add_associative add_right_top mult_left_dist_add nu_below_mu_nu_L_def) qed -- {* AACP Theorem 4.13 implies Theorem 4.15 *} lemma kappa_mu_nu_L_mu_nu_apx_nu_L: "has_greatest_fixpoint f \ kappa_mu_nu_L f \ mu_nu_apx_nu_L f" by (metis mu_nu_apx_nu_L_def kappa_apx_below_nu kappa_mu_nu_L_def) -- {* AACP Theorem 4.16 implies Theorem 4.15 *} lemma mu_nu_apx_meet_L_mu_nu_apx_nu_L: "mu_nu_apx_meet_L f \ mu_nu_apx_nu_L f" by (smt apx_meet_same has_apx_meet_def is_apx_meet_def mu_nu_apx_meet_L_def mu_nu_apx_nu_L_def) -- {* AACP Theorem 4.13 implies Theorem 4.14 *} lemma kappa_mu_nu_L_nu_below_mu_nu_L: "has_least_fixpoint f \ has_greatest_fixpoint f \ kappa_mu_nu_L f \ nu_below_mu_nu_L f" by (metis mu_nu_apx_nu_L_nu_below_mu_nu_L kappa_mu_nu_L_mu_nu_apx_nu_L) lemma nu_below_mu_nu_nu_below_mu_nu_L: "nu_below_mu_nu f \ nu_below_mu_nu_L f" nitpick [expect=genuine] oops lemma unfold_fold_1: "isotone f \ has_least_prefixpoint f \ apx.has_least_fixpoint f \ f(x) \ x \ \ f \ x + L" by (metis add_left_isotone apx_def has_least_fixpoint_def is_least_prefixpoint_def least_prefixpoint_char least_prefixpoint_fixpoint order_trans pmu_mu kappa_apx_below_mu) lemma unfold_fold_2: "isotone f \ apx.isotone f \ has_least_prefixpoint f \ has_greatest_fixpoint f \ apx.has_least_fixpoint f \ f(x) \ x \ \ f \ L \ x \ L \ \ f \ x" proof assume 1: "isotone f \ apx.isotone f \ has_least_prefixpoint f \ has_greatest_fixpoint f \ apx.has_least_fixpoint f \ f(x) \ x \ \ f \ L \ x \ L" hence "\ f \ L = \ f \ L" by (metis apx_meet_L meet.add_left_isotone meet.antisym kappa_apx_below_nu kappa_below_nu) hence "\ f = (\ f \ L) + \ f" using 1 by (metis apx_meet_below_nu_kappa_mu_nu has_apx_least_fixpoint_kappa_apx_meet add_commutative least_fixpoint_char least_prefixpoint_fixpoint kappa_apx_meet_apx_meet_below_nu kappa_mu_nu_def) thus "\ f \ x" using 1 by (metis add_least_upper_bound is_least_prefixpoint_def least_prefixpoint meet.add_least_upper_bound pmu_mu) qed end class n_algebra_apx_2 = n_algebra + apx + assumes apx_def: "x \ y \ x \ y + L \ y \ x + n(x) ; T" begin lemma apx_transitive_2: "x \ y \ y \ z \ x \ z" proof assume 1: "x \ y \ y \ z" hence "z \ y + n(y) ; T" by (metis apx_def) also have "... \ x + n(x) ; T + n(y) ; T" using 1 by (metis add_left_isotone apx_def) also have "... \ x + n(x) ; T" using 1 by (metis add_associative add_idempotent add_right_isotone apx_def mult_left_isotone n_add_n_top n_isotone) finally show "x \ z" using 1 by (smt add_associative add_commutative apx_def less_eq_def) qed lemma apx_meet_L: "y \ x \ x \ L \ y \ L" proof assume 1: "y \ x" have "x \ L \ (y \ L) + (n(y) ; T \ L)" using 1 by (metis apx_def meet_commutative meet_left_dist_add meet.add_left_isotone) also have "... \ (y \ L) + n(y \ L) ; T" by (metis add_least_upper_bound n_vector_meet_L meet.add_least_upper_bound n_L_decreasing order_refl order_trans) finally show "x \ L \ y \ L" by (metis less_eq_def meet.add_right_upper_bound n_less_eq_char) qed -- {* AACP Theorem 4.1 *} subclass apx_biorder apply unfold_locales apply (metis add_left_upper_bound apx_def) apply (metis add_same_context antisym apx_def apx_meet_L relative_equality) apply (metis apx_transitive_2) done lemma add_apx_left_isotone_2: "x \ y \ x + z \ y + z" proof assume 1: "x \ y" hence 2: "x + z \ y + z + L" by (smt add_associative add_commutative add_left_isotone apx_def) have "y + z \ x + n(x) ; T + z" using 1 by (metis add_left_isotone apx_def) also have "... \ x + z + n(x + z) ; T" by (metis add_associative add_commutative add_right_isotone mult_left_isotone n_right_upper_bound) finally show "x + z \ y + z" using 2 by (metis apx_def) qed lemma mult_apx_left_isotone_2: "x \ y \ x ; z \ y ; z" proof assume 1: "x \ y" hence "x ; z \ y ; z + L ; z" by (metis apx_def mult_left_isotone mult_right_dist_add) hence 2: "x ; z \ y ; z + L" by (metis add_commutative add_left_isotone n_L_below_L order_trans) have "y ; z \ x ; z + n(x) ; T ; z" using 1 by (metis apx_def mult_left_isotone mult_right_dist_add) also have "... \ x ; z + n(x ; z) ; T" by (metis add_least_upper_bound add_left_upper_bound n_top_split) finally show "x ; z \ y ; z" using 2 by (metis apx_def) qed lemma mult_apx_right_isotone_2: "x \ y \ z ; x \ z ; y" proof assume 1: "x \ y" hence "z ; x \ z ; y + z ; L" by (metis apx_def mult_left_dist_add mult_right_isotone) also have "... \ z ; y + z ; 0 + L" by (metis add_associative add_right_isotone n_L_split_L) finally have 2: "z ; x \ z ; y + L" by (metis add_right_zero mult_left_dist_add) have "z ; y \ z ; (x + n(x) ; T)" using 1 by (metis apx_def mult_right_isotone) also have "... = z ; x + z ; n(x) ; T" by (metis mult_associative mult_left_dist_add) also have "... \ z ; x + n(z ; x) ; T" by (metis add_least_upper_bound add_left_upper_bound n_split_top) finally show "z ; x \ z ; y" using 2 by (metis apx_def) qed end (* end theory NOmegaAlgebra imports OmegaAlgebra Recursion begin *) class itering_apx = bounded_itering + n_algebra_apx begin lemma circ_L: "L\<^sup>\ = L + 1" by (metis add_commutative mult_top_circ n_L_top_L) lemma C_circ_import: "C (x\<^sup>\) \ (C x)\<^sup>\" proof - have 1: "C x ; x\<^sup>\ \ (C x)\<^sup>\ ; C x" by (metis C_mult_propagate circ_simulate eq_refl) have "C (x\<^sup>\) = C (1 + x ; x\<^sup>\)" by (metis circ_left_unfold) also have "... = C 1 + C (x ; x\<^sup>\)" by (metis meet_left_dist_add) also have "... \ 1 + C (x ; x\<^sup>\)" by (metis add_left_isotone meet.add_right_upper_bound) also have "... = 1 + C x ; x\<^sup>\" by (metis n_L_T_meet_mult) also have "... \ (C x)\<^sup>\" using 1 by (metis add_right_isotone circ_left_unfold circ_plus_same) finally show ?thesis . qed -- {* AACP Theorem 4.3 and Theorem 4.4 *} lemma circ_apx_isotone: "x \ y \ x\<^sup>\ \ y\<^sup>\" proof assume "x \ y" hence 1: "x \ y + L \ C y \ x + n(x) ; T" by (metis apx_def) have "C (y\<^sup>\) \ (C y)\<^sup>\" by (metis C_circ_import) also have "... \ x\<^sup>\ + x\<^sup>\ ; n(x) ; T" using 1 by (metis circ_isotone circ_left_top circ_unfold_sum mult_associative) also have "... \ x\<^sup>\ + (x\<^sup>\ ; 0 + n(x\<^sup>\ ; x) ; T)" by (smt add_right_isotone n_n_top_split_n_top) also have "... \ x\<^sup>\ + (x\<^sup>\ ; 0 + n(x\<^sup>\) ; T)" by (metis add_right_isotone mult_left_isotone n_isotone right_plus_below_circ) also have "... = x\<^sup>\ + n(x\<^sup>\) ; T" by (smt add_associative add_commutative less_eq_def zero_right_mult_decreasing) finally have 2: "C (y\<^sup>\) \ x\<^sup>\ + n(x\<^sup>\) ; T" by metis have "x\<^sup>\ \ y\<^sup>\ ; L\<^sup>\" using 1 by (metis circ_add_1 circ_back_loop_fixpoint circ_isotone n_L_below_L less_eq_def mult_associative) also have "... = y\<^sup>\ + y\<^sup>\ ; L" by (metis add_commutative circ_L mult_left_dist_add mult_right_one) also have "... \ y\<^sup>\ + y\<^sup>\ ; 0 + L" by (metis add_associative add_right_isotone n_L_split_L) finally have "x\<^sup>\ \ y\<^sup>\ + L" by (metis add_commutative less_eq_def zero_right_mult_decreasing) thus "x\<^sup>\ \ y\<^sup>\" using 2 by (metis apx_def) qed end class n_omega_algebra_1 = bounded_left_zero_omega_algebra + n_algebra_apx + Omega + assumes Omega_def: "x\<^sup>\ = n(x\<^sup>\) ; L + x\<^sup>\" begin -- {* AACP Theorem 8.13 *} lemma C_omega_export: "C (x\<^sup>\) = (C x)\<^sup>\" proof - have "C (x\<^sup>\) = C x ; C (x\<^sup>\)" by (metis C_mult_propagate n_L_T_meet_mult omega_unfold) hence 1: "C (x\<^sup>\) \ (C x)\<^sup>\" by (metis omega_induct_mult order.refl) have "(C x)\<^sup>\ = C (x ; (C x)\<^sup>\)" by (metis n_L_T_meet_mult omega_unfold) also have "... \ C (x\<^sup>\)" by (metis calculation meet.add_least_upper_bound meet.add_left_upper_bound meet_commutative omega_isotone) finally show ?thesis using 1 by (metis antisym) qed -- {* AACP Theorem 8.2 *} lemma L_mult_star: "L ; x\<^sup>\ = L" by (metis less_eq_def mult_associative n_L_below_L star.circ_back_loop_fixpoint) -- {* AACP Theorem 8.3 *} lemma mult_L_star: "(x ; L)\<^sup>\ = 1 + x ; L" by (smt L_mult_star mult_associative star.circ_mult) lemma mult_L_omega_below: "(x ; L)\<^sup>\ \ x ; L" by (metis mult_right_isotone n_L_below_L omega_slide) -- {* AACP Theorem 8.5 *} lemma mult_L_add_star: "(x ; L + y)\<^sup>\ = y\<^sup>\ + y\<^sup>\ ; x ; L" by (metis L_mult_star mult_associative star.circ_add_1 star.circ_decompose_6 star.circ_unfold_sum) lemma mult_L_add_omega_below: "(x ; L + y)\<^sup>\ \ y\<^sup>\ + y\<^sup>\ ; x ; L" proof - have "(x ; L + y)\<^sup>\ \ y\<^sup>\ ; x ; L + (y\<^sup>\ ; x ; L)\<^sup>\ ; y\<^sup>\" by (metis add_commutative mult_associative omega_decompose add_left_isotone mult_L_omega_below) also have "... \ y\<^sup>\ + y\<^sup>\ ; x ; L" by (smt add_associative add_commutative less_eq_def mult_L_star mult_associative mult_left_dist_add mult_left_one mult_right_dist_add n_L_below_L order_refl) finally show ?thesis by metis qed lemma n_Omega_isotone: "x \ y \ x\<^sup>\ \ y\<^sup>\" by (metis Omega_def add_isotone mult_left_isotone n_isotone omega_isotone star_isotone) lemma n_star_below_Omega: "x\<^sup>\ \ x\<^sup>\" by (metis add_right_upper_bound Omega_def) lemma mult_L_star_mult_below: "(x ; L)\<^sup>\ ; y \ y + x ; L" by (metis add_right_isotone mult_associative mult_right_isotone n_L_below_L star_left_induct) end sublocale n_omega_algebra_1 < star!: itering_apx where circ = star .. class n_omega_algebra = n_omega_algebra_1 + n_algebra_apx + assumes n_split_omega_mult: "C (x\<^sup>\) \ x\<^sup>\ ; n(x\<^sup>\) ; T" assumes tarski: "x ; L \ x ; L ; x ; L" begin -- {* AACP Theorem 8.4 *} lemma mult_L_omega: "(x ; L)\<^sup>\ = x ; L" apply (rule antisym) apply (rule mult_L_omega_below) apply (metis mult_associative omega_induct_mult tarski) done -- {* AACP Theorem 8.6 *} lemma mult_L_add_omega: "(x ; L + y)\<^sup>\ = y\<^sup>\ + y\<^sup>\ ; x ; L" apply (rule antisym) apply (rule mult_L_add_omega_below) apply (metis add_right_isotone add_right_upper_bound less_eq_def mult_L_omega mult_associative mult_isotone omega_sub_dist star.circ_sub_dist star_mult_omega) done -- {* AACP Theorem 8.1 *} lemma tarski_mult_top_idempotent: "x ; L = x ; L ; x ; L" by (metis mult_L_omega mult_associative omega_unfold) -- {* AACP Theorem 8.7 *} lemma n_below_n_omega: "n(x) \ n(x\<^sup>\)" proof - have "n(x) ; L \ n(x) ; L ; n(x) ; L" by (metis tarski) also have "... \ x ; n(x) ; L" by (metis mult_left_isotone n_L_decreasing) finally have "n(x) ; L \ x\<^sup>\" by (metis mult_associative omega_induct_mult) thus ?thesis by (metis n_galois) qed -- {* AACP Theorem 8.14 *} lemma n_split_omega_add_zero: "C (x\<^sup>\) \ x\<^sup>\ ; 0 + n(x\<^sup>\) ; T" proof - have "n(x\<^sup>\) ; T + x ; (x\<^sup>\ ; 0 + n(x\<^sup>\) ; T) = n(x\<^sup>\) ; T + x ; x\<^sup>\ ; 0 + x ; n(x\<^sup>\) ; T" by (metis add_associative mult_associative mult_left_dist_add) also have "... \ n(x\<^sup>\) ; T + x ; x\<^sup>\ ; 0 + x ; 0 + n(x\<^sup>\) ; T" by (metis add_associative add_right_isotone n_n_top_split_n_top omega_unfold) also have "... = x ; x\<^sup>\ ; 0 + n(x\<^sup>\) ; T" by (smt add_associative add_commutative add_left_top add_right_zero mult_associative mult_left_dist_add) also have "... \ x\<^sup>\ ; 0 + n(x\<^sup>\) ; T" by (metis add_left_isotone mult_left_isotone star.left_plus_below_circ) finally have "x\<^sup>\ ; n(x\<^sup>\) ; T \ x\<^sup>\ ; 0 + n(x\<^sup>\) ; T" by (metis mult_associative star_left_induct) thus ?thesis by (metis n_split_omega_mult order_trans) qed lemma n_split_omega_add: "C (x\<^sup>\) \ x\<^sup>\ + n(x\<^sup>\) ; T" by (metis add_left_isotone n_split_omega_add_zero order_trans zero_right_mult_decreasing) -- {* AACP Theorem 8.12 *} lemma n_dist_omega_star: "n(y\<^sup>\ + y\<^sup>\ ; z) = n(y\<^sup>\) + n(y\<^sup>\ ; z)" proof - have "n(y\<^sup>\ + y\<^sup>\ ; z) = n(C (y\<^sup>\) + C (y\<^sup>\ ; z))" by (metis meet_left_dist_add n_C) also have "... \ n(C (y\<^sup>\) + y\<^sup>\ ; z)" by (smt2 add_least_upper_bound add_left_upper_bound add_right_upper_bound meet.add_left_upper_bound meet_commutative n_isotone order_trans) also have "... \ n(y\<^sup>\ ; 0 + n(y\<^sup>\) ; T + y\<^sup>\ ; z)" by (metis add_commutative add_right_isotone n_isotone n_split_omega_add_zero) also have "... = n(y\<^sup>\) + n(y\<^sup>\ ; z)" by (smt add_associative add_commutative add_right_zero mult_left_dist_add n_dist_n_add) finally show ?thesis by (metis add_least_upper_bound n_left_upper_bound n_right_upper_bound antisym) qed lemma mult_L_add_circ_below: "(x ; L + y)\<^sup>\ \ n(y\<^sup>\) ; L + y\<^sup>\ + y\<^sup>\ ; x ; L" proof - have "(x ; L + y)\<^sup>\ \ n(y\<^sup>\ + y\<^sup>\ ; x ; L) ; L + (x ; L + y)\<^sup>\" by (metis add_left_isotone mult_L_add_omega_below mult_left_isotone n_isotone Omega_def) also have "... = n(y\<^sup>\) ; L + n(y\<^sup>\ ; x ; L) ; L + (x ; L + y)\<^sup>\" by (metis mult_associative mult_right_dist_add n_dist_omega_star) also have "... \ n(y\<^sup>\) ; L + y\<^sup>\ + y\<^sup>\ ; x ; L" by (smt add_associative add_commutative add_idempotent add_right_isotone mult_L_add_star n_L_decreasing) finally show ?thesis by metis qed lemma n_mult_omega_L_below_zero: "n(y ; x\<^sup>\) ; L \ y ; x\<^sup>\ ; 0 + y ; n(x\<^sup>\) ; L" proof - have "n(y ; x\<^sup>\) ; L \ C (y ; x\<^sup>\) \ L" by (metis n_C n_L_increasing n_galois n_n_L n_n_meet_L) also have "... \ y ; C (x\<^sup>\) \ L" by (metis meet.add_left_isotone n_L_T_meet_mult n_L_T_meet_mult_propagate) also have "... \ y ; (x\<^sup>\ ; 0 + n(x\<^sup>\) ; T) \ L" by (metis meet_commutative meet.add_right_isotone mult_right_isotone n_split_omega_add_zero) also have "... = (y ; x\<^sup>\ ; 0 \ L) + (y ; n(x\<^sup>\) ; T \ L)" by (metis meet_commutative meet_left_dist_add mult_associative mult_left_dist_add) also have "... \ (y ; x\<^sup>\ ; 0 \ L) + y ; n(x\<^sup>\) ; L" by (metis add_right_isotone n_vector_meet_L) also have "... \ y ; x\<^sup>\ ; 0 + y ; n(x\<^sup>\) ; L" by (metis add_left_isotone meet.add_left_upper_bound) finally show ?thesis by metis qed -- {* AACP Theorem 8.10 *} lemma n_mult_omega_L_star_zero: "y ; x\<^sup>\ ; 0 + n(y ; x\<^sup>\) ; L = y ; x\<^sup>\ ; 0 + y ; n(x\<^sup>\) ; L" apply (rule antisym) apply (metis add_least_upper_bound mult_associative mult_left_dist_add mult_left_sub_dist_add_left n_mult_omega_L_below_zero) apply (smt add_associative add_commutative add_left_zero add_right_isotone mult_associative mult_left_dist_add n_n_L_split_n_L) done -- {* AACP Theorem 8.11 *} lemma n_mult_omega_L_star: "y ; x\<^sup>\ + n(y ; x\<^sup>\) ; L = y ; x\<^sup>\ + y ; n(x\<^sup>\) ; L" by (metis zero_right_mult_decreasing n_mult_omega_L_star_zero add_relative_same_increasing) lemma n_mult_omega_L_below: "n(y ; x\<^sup>\) ; L \ y ; x\<^sup>\ + y ; n(x\<^sup>\) ; L" by (metis add_right_upper_bound n_mult_omega_L_star) lemma n_omega_L_below_zero: "n(x\<^sup>\) ; L \ x ; x\<^sup>\ ; 0 + x ; n(x\<^sup>\) ; L" by (smt omega_unfold n_mult_omega_L_below_zero add_left_isotone star.left_plus_below_circ order_trans) lemma n_omega_L_below: "n(x\<^sup>\) ; L \ x\<^sup>\ + x ; n(x\<^sup>\) ; L" by (metis omega_unfold n_mult_omega_L_below add_left_isotone star.left_plus_below_circ order_trans) lemma n_omega_L_star_zero: "x ; x\<^sup>\ ; 0 + n(x\<^sup>\) ; L = x ; x\<^sup>\ ; 0 + x ; n(x\<^sup>\) ; L" by (metis n_mult_omega_L_star_zero omega_unfold) -- {* AACP Theorem 8.8 *} lemma n_omega_L_star: "x\<^sup>\ + n(x\<^sup>\) ; L = x\<^sup>\ + x ; n(x\<^sup>\) ; L" by (metis star.circ_mult_upper_bound star.left_plus_below_circ zero_least n_omega_L_star_zero add_relative_same_increasing) -- {* AACP Theorem 8.9 *} lemma n_omega_L_star_zero_star: "x\<^sup>\ ; 0 + n(x\<^sup>\) ; L = x\<^sup>\ ; 0 + x\<^sup>\ ; n(x\<^sup>\) ; L" by (metis n_mult_omega_L_star_zero star_mult_omega mult_associative star.circ_transitive_equal) -- {* AACP Theorem 8.8 *} lemma n_omega_L_star_star: "x\<^sup>\ + n(x\<^sup>\) ; L = x\<^sup>\ + x\<^sup>\ ; n(x\<^sup>\) ; L" by (metis zero_right_mult_decreasing n_omega_L_star_zero_star add_relative_same_increasing) lemma n_Omega_left_unfold: "1 + x ; x\<^sup>\ = x\<^sup>\" by (smt Omega_def add_associative add_commutative mult_associative mult_left_dist_add n_omega_L_star star.circ_left_unfold) lemma n_Omega_left_slide: "(x ; y)\<^sup>\ ; x \ x ; (y ; x)\<^sup>\" proof - have "(x ; y)\<^sup>\ ; x \ x ; y ; n((x ; y)\<^sup>\) ; L + (x ; y)\<^sup>\ ; x" by (smt Omega_def add_commutative add_left_isotone mult_associative mult_right_dist_add mult_right_isotone n_L_below_L n_omega_L_star) also have "... \ x ; (y ; 0 + n(y ; (x ; y)\<^sup>\) ; L) + (x ; y)\<^sup>\ ; x" by (smt add_associative add_commutative less_eq_def mult_associative mult_left_dist_add mult_left_sub_dist_add_left n_n_L_split_n_L star.circ_slide) also have "... = x ; (y ; x)\<^sup>\" by (smt Omega_def add_associative add_commutative less_eq_def mult_associative mult_isotone mult_left_dist_add omega_slide star.circ_increasing star.circ_slide zero_least) finally show ?thesis by metis qed lemma n_Omega_add_1: "(x + y)\<^sup>\ = x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" proof - have 1: "(x + y)\<^sup>\ = n((x\<^sup>\ ; y)\<^sup>\) ; L + n((x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\) ; L + (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\" by (smt Omega_def mult_right_dist_add n_dist_omega_star omega_decompose star.circ_add) have "n((x\<^sup>\ ; y)\<^sup>\) ; L \ (x\<^sup>\ ; y)\<^sup>\ + x\<^sup>\ ; (y ; n((x\<^sup>\ ; y)\<^sup>\) ; L)" by (metis n_omega_L_below mult_associative) also have "... \ (x\<^sup>\ ; y)\<^sup>\ + x\<^sup>\ ; y ; 0 + x\<^sup>\ ; n((y ; x\<^sup>\)\<^sup>\) ; L" by (smt add_associative add_right_isotone mult_associative mult_left_dist_add mult_right_isotone n_n_L_split_n_L omega_slide) also have "... = (x\<^sup>\ ; y)\<^sup>\ + x\<^sup>\ ; n((y ; x\<^sup>\)\<^sup>\) ; L" by (metis add_commutative less_eq_def star.circ_sub_dist_1 zero_right_mult_decreasing) also have "... \ x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ + x\<^sup>\ ; n((y ; x\<^sup>\)\<^sup>\) ; L" by (metis add_left_isotone mult_right_isotone star.circ_increasing star.circ_isotone star_decompose_3) also have "... \ x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" by (metis Omega_def add_commutative mult_associative mult_left_dist_add mult_right_isotone n_Omega_isotone n_star_below_Omega) also have "... \ x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" by (metis n_star_below_Omega mult_left_isotone) finally have 2: "n((x\<^sup>\ ; y)\<^sup>\) ; L \ x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" by metis have "n((x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\) ; L \ n(x\<^sup>\) ; L + x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ + x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ ; y ; n(x\<^sup>\) ; L" by (smt add_associative add_commutative mult_left_one mult_right_dist_add n_mult_omega_L_below star.circ_mult star.circ_slide) also have "... = n(x\<^sup>\) ; L ; (y ; x\<^sup>\)\<^sup>\ + x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" by (smt Omega_def add_associative mult_L_add_star mult_associative mult_left_dist_add L_mult_star) also have "... \ x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" by (metis mult_right_dist_add Omega_def n_star_below_Omega mult_right_isotone) finally have 3: "n((x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\) ; L \ x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" by metis have "(x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ \ x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" by (metis star_slide mult_isotone mult_right_isotone n_star_below_Omega order_trans star_isotone) hence 4: "(x + y)\<^sup>\ \ x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" using 1 2 3 by (metis add_least_upper_bound) have 5: "x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ \ n(x\<^sup>\) ; L + x\<^sup>\ ; n((y ; x\<^sup>\)\<^sup>\) ; L + x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\" by (smt Omega_def add_associative add_left_isotone mult_associative mult_left_dist_add mult_right_dist_add mult_right_isotone n_L_below_L) have "n(x\<^sup>\) ; L \ n((x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\) ; L" by (metis add_commutative add_left_upper_bound mult_left_isotone n_isotone star.circ_loop_fixpoint) hence 6: "n(x\<^sup>\) ; L \ (x + y)\<^sup>\" using 1 by (metis Omega_def add_left_upper_bound n_Omega_isotone order_trans) have "x\<^sup>\ ; n((y ; x\<^sup>\)\<^sup>\) ; L \ x\<^sup>\ ; n((y ; x\<^sup>\)\<^sup>\ + (y ; x\<^sup>\)\<^sup>\ ; y ; n(x\<^sup>\) ; L) ; L" by (metis Omega_def mult_L_add_omega_below mult_associative mult_left_dist_add mult_left_isotone mult_right_isotone n_isotone) also have "... \ x\<^sup>\ ; 0 + n(x\<^sup>\ ; ((y ; x\<^sup>\)\<^sup>\ + (y ; x\<^sup>\)\<^sup>\ ; y ; n(x\<^sup>\) ; L)) ; L" by (metis n_n_L_split_n_L) also have "... \ x\<^sup>\ + n((x\<^sup>\ ; y)\<^sup>\ + x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ ; y ; n(x\<^sup>\) ; L) ; L" by (smt add_left_isotone mult_associative mult_left_dist_add omega_slide zero_right_mult_decreasing) also have "... \ x\<^sup>\ + n((x\<^sup>\ ; y)\<^sup>\ + (x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L) ; L" by (smt add_right_divisibility add_right_isotone mult_left_isotone n_isotone star.circ_mult) also have "... \ x\<^sup>\ + n((x + y)\<^sup>\) ; L" by (metis add_right_isotone mult_associative mult_left_isotone mult_right_isotone n_L_decreasing n_isotone omega_decompose) also have "... \ (x + y)\<^sup>\" by (metis add_left_isotone star.circ_sub_dist Omega_def add_commutative) finally have 7: "x\<^sup>\ ; n((y ; x\<^sup>\)\<^sup>\) ; L \ (x + y)\<^sup>\" by metis have "x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ \ (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ + (x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L" by (smt Omega_def add_right_isotone mult_L_add_star mult_associative mult_left_dist_add mult_left_isotone star.left_plus_below_circ star_slide) also have "... \ (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ + n((x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\) ; L" by (metis add_associative add_right_isotone add_right_zero mult_left_dist_add n_n_L_split_n_L) also have "... \ (x + y)\<^sup>\" by (smt Omega_def add_commutative add_right_isotone mult_left_isotone n_right_upper_bound omega_decompose star.circ_add) finally have "n(x\<^sup>\) ; L + x\<^sup>\ ; n((y ; x\<^sup>\)\<^sup>\) ; L + x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ \ (x + y)\<^sup>\" using 6 7 by (metis add_least_upper_bound) hence "x\<^sup>\ ; (y ; x\<^sup>\)\<^sup>\ \ (x + y)\<^sup>\" using 5 by (smt order_trans) thus ?thesis using 4 by (metis antisym) qed end sublocale n_omega_algebra < nL_omega!: left_zero_conway_semiring where circ = Omega apply unfold_locales apply (metis n_Omega_left_unfold) apply (metis n_Omega_left_slide) apply (metis n_Omega_add_1) done (* circ_plus_same does not hold in the non-strict model using Omega *) context n_omega_algebra begin -- {* AACP Theorem 8.16 *} lemma omega_apx_isotone: "x \ y \ x\<^sup>\ \ y\<^sup>\" proof assume "x \ y" hence 1: "x \ y + L \ C y \ x + n(x) ; T" by (metis apx_def) have "n(x) ; T + x ; (x\<^sup>\ + n(x\<^sup>\) ; T) \ n(x) ; T + x\<^sup>\ + n(x\<^sup>\) ; T" by (smt add_associative mult_associative mult_left_dist_add add_right_isotone n_n_top_split_n_top add_right_zero omega_unfold) also have "... \ x\<^sup>\ + n(x\<^sup>\) ; T" by (metis add_commutative add_right_isotone mult_left_isotone n_below_n_omega add_associative add_idempotent) finally have 2: "x\<^sup>\ ; n(x) ; T \ x\<^sup>\ + n(x\<^sup>\) ; T" by (metis mult_associative star_left_induct) have "C (y\<^sup>\) = (C y)\<^sup>\" by (metis C_omega_export) also have "... \ (x + n(x) ; T)\<^sup>\" using 1 by (metis omega_isotone) also have "... = (x\<^sup>\ ; n(x) ; T)\<^sup>\ + (x\<^sup>\ ; n(x) ; T)\<^sup>\ ; x\<^sup>\" by (metis mult_associative omega_decompose) also have "... \ x\<^sup>\ ; n(x) ; T + (x\<^sup>\ ; n(x) ; T)\<^sup>\ ; x\<^sup>\" by (metis add_left_isotone mult_top_omega) also have "... = x\<^sup>\ ; n(x) ; T + (1 + x\<^sup>\ ; n(x) ; T ; (x\<^sup>\ ; n(x) ; T)\<^sup>\) ; x\<^sup>\" by (metis mult_associative star.circ_left_top star.mult_top_circ) also have "... \ x\<^sup>\ + x\<^sup>\ ; n(x) ; T" by (smt add_isotone add_least_upper_bound mult_associative mult_left_one mult_right_dist_add mult_right_isotone order_refl top_greatest) also have "... \ x\<^sup>\ + n(x\<^sup>\) ; T" using 2 by (metis add_least_upper_bound add_left_upper_bound) finally have 3: "C (y\<^sup>\) \ x\<^sup>\ + n(x\<^sup>\) ; T" by metis have "x\<^sup>\ \ (y + L)\<^sup>\" using 1 by (metis omega_isotone) also have "... = (y\<^sup>\ ; L)\<^sup>\ + (y\<^sup>\ ; L)\<^sup>\ ; y\<^sup>\" by (metis omega_decompose) also have "... = y\<^sup>\ ; L ; (y\<^sup>\ ; L)\<^sup>\ + (y\<^sup>\ ; L)\<^sup>\ ; y\<^sup>\" by (metis omega_unfold) also have "... \ y\<^sup>\ ; L + (y\<^sup>\ ; L)\<^sup>\ ; y\<^sup>\" by (metis add_left_isotone n_L_below_L mult_associative mult_right_isotone) also have "... = y\<^sup>\ ; L + (1 + y\<^sup>\ ; L ; (y\<^sup>\ ; L)\<^sup>\) ; y\<^sup>\" by (metis star.circ_left_unfold) also have "... \ y\<^sup>\ ; L + y\<^sup>\" by (metis add_commutative add_least_upper_bound add_right_upper_bound mult_L_star_mult_below mult_associative star.circ_mult star.circ_slide) also have "... \ y\<^sup>\ ; 0 + L + y\<^sup>\" by (metis add_left_isotone n_L_split_L) finally have "x\<^sup>\ \ y\<^sup>\ + L" by (metis add_associative add_commutative less_eq_def star_zero_below_omega) thus "x\<^sup>\ \ y\<^sup>\" using 3 by (metis apx_def) qed lemma combined_apx_left_isotone: "x \ y \ n(x\<^sup>\) ; L + x\<^sup>\ ; z \ n(y\<^sup>\) ; L + y\<^sup>\ ; z" by (metis add_apx_isotone mult_apx_left_isotone omega_apx_isotone star.circ_apx_isotone n_L_apx_isotone) lemma combined_apx_left_isotone_2: "x \ y \ (x\<^sup>\ \ L) + x\<^sup>\ ; z \ (y\<^sup>\ \ L) + y\<^sup>\ ; z" by (metis add_apx_isotone mult_apx_left_isotone omega_apx_isotone star.circ_apx_isotone meet_L_apx_isotone) lemma combined_apx_right_isotone: "y \ z \ n(x\<^sup>\) ; L + x\<^sup>\ ; y \ n(x\<^sup>\) ; L + x\<^sup>\ ; z" by (metis add_apx_right_isotone mult_apx_right_isotone) lemma combined_apx_right_isotone_2: "y \ z \ (x\<^sup>\ \ L) + x\<^sup>\ ; y \ (x\<^sup>\ \ L) + x\<^sup>\ ; z" by (metis add_apx_right_isotone mult_apx_right_isotone) lemma combined_apx_isotone: "x \ y \ w \ z \ n(x\<^sup>\) ; L + x\<^sup>\ ; w \ n(y\<^sup>\) ; L + y\<^sup>\ ; z" by (metis add_apx_isotone mult_apx_isotone omega_apx_isotone star.circ_apx_isotone n_L_apx_isotone) lemma combined_apx_isotone_2: "x \ y \ w \ z \ (x\<^sup>\ \ L) + x\<^sup>\ ; w \ (y\<^sup>\ \ L) + y\<^sup>\ ; z" by (metis add_apx_isotone mult_apx_isotone omega_apx_isotone star.circ_apx_isotone meet_L_apx_isotone) lemma n_split_nu_mu: "C (y\<^sup>\ + y\<^sup>\ ; z) \ y\<^sup>\ ; z + n(y\<^sup>\ + y\<^sup>\ ; z) ; T" proof - have "C (y\<^sup>\ + y\<^sup>\ ; z) \ C (y\<^sup>\) + y\<^sup>\ ; z" by (metis add_right_isotone meet.add_right_upper_bound meet_left_dist_add) also have "... \ y\<^sup>\ ; 0 + n(y\<^sup>\) ; T + y\<^sup>\ ; z" by (metis add_commutative add_right_isotone n_split_omega_add_zero) also have "... \ y\<^sup>\ ; z + n(y\<^sup>\ + y\<^sup>\ ; z) ; T" by (smt add_associative add_commutative add_right_isotone add_right_zero mult_left_dist_add mult_left_isotone n_left_upper_bound) finally show ?thesis by metis qed lemma n_split_nu_mu_2: "C (y\<^sup>\ + y\<^sup>\ ; z) \ y\<^sup>\ ; z + ((y\<^sup>\ + y\<^sup>\ ; z) \ L) + n(y\<^sup>\ + y\<^sup>\ ; z) ; T" proof - have "C (y\<^sup>\ + y\<^sup>\ ; z) \ C (y\<^sup>\) + y\<^sup>\ ; z" by (metis add_right_isotone meet.add_right_upper_bound meet_left_dist_add) also have "... \ y\<^sup>\ ; 0 + n(y\<^sup>\) ; T + y\<^sup>\ ; z" by (metis add_commutative add_right_isotone n_split_omega_add_zero) also have "... \ y\<^sup>\ ; z + n(y\<^sup>\ + y\<^sup>\ ; z) ; T" by (smt add_associative add_commutative add_right_isotone add_right_zero mult_left_dist_add mult_left_isotone n_left_upper_bound) finally show ?thesis by (smt2 add_associative add_commutative meet_left_dist_add meet_commutative add_right_upper_bound order_trans) qed lemma loop_exists: "C (\ (\x . y ; x + z)) \ \ (\x . y ; x + z) + n(\ (\x . y ; x + z)) ; T" by (metis n_split_nu_mu omega_loop_nu star_loop_mu) lemma loop_exists_2: "C (\ (\x . y ; x + z)) \ \ (\x . y ; x + z) + (\ (\x . y ; x + z) \ L) + n(\ (\x . y ; x + z)) ; T" by (metis n_split_nu_mu_2 omega_loop_nu star_loop_mu) lemma loop_apx_least_fixpoint: "apx.is_least_fixpoint (\x . y ; x + z) (\ (\x . y ; x + z) + n(\ (\x . y ; x + z)) ; L)" proof - have "kappa_mu_nu_L (\x . y ; x + z)" by (metis affine_apx_isotone loop_exists affine_has_greatest_fixpoint affine_has_least_fixpoint affine_isotone nu_below_mu_nu_L_def nu_below_mu_nu_L_kappa_mu_nu_L) thus ?thesis by (smt apx.least_fixpoint_char kappa_mu_nu_L_def) qed lemma loop_apx_least_fixpoint_2: "apx.is_least_fixpoint (\x . y ; x + z) (\ (\x . y ; x + z) + (\ (\x . y ; x + z) \ L))" proof - have "kappa_mu_nu (\x . y ; x + z)" by (metis affine_apx_isotone affine_has_greatest_fixpoint affine_has_least_fixpoint affine_isotone loop_exists_2 nu_below_mu_nu_def nu_below_mu_nu_kappa_mu_nu) thus ?thesis by (smt apx.least_fixpoint_char kappa_mu_nu_def) qed lemma loop_has_apx_least_fixpoint: "apx.has_least_fixpoint (\x . y ; x + z)" by (metis apx.has_least_fixpoint_def loop_apx_least_fixpoint_2) lemma loop_semantics: "\ (\x . y ; x + z) = \ (\x . y ; x + z) + n(\ (\x . y ; x + z)) ; L" by (metis apx.least_fixpoint_char loop_apx_least_fixpoint) lemma loop_semantics_2: "\ (\x . y ; x + z) = \ (\x . y ; x + z) + (\ (\x . y ; x + z) \ L)" by (metis apx.least_fixpoint_char loop_apx_least_fixpoint_2) -- {* AACP Theorem 8.15 *} lemma loop_semantics_kappa_mu_nu: "\ (\x . y ; x + z) = n(y\<^sup>\) ; L + y\<^sup>\ ; z" proof - have "\ (\x . y ; x + z) = y\<^sup>\ ; z + n(y\<^sup>\ + y\<^sup>\ ; z) ; L" by (metis loop_semantics omega_loop_nu star_loop_mu) thus ?thesis by (smt n_dist_omega_star add_associative mult_right_dist_add add_commutative less_eq_def n_L_decreasing) qed -- {* AACP Theorem 8.15 *} lemma loop_semantics_kappa_mu_nu_2: "\ (\x . y ; x + z) = (y\<^sup>\ \ L) + y\<^sup>\ ; z" proof - have "\ (\x . y ; x + z) = y\<^sup>\ ; z + ((y\<^sup>\ + y\<^sup>\ ; z) \ L)" by (metis loop_semantics_2 omega_loop_nu star_loop_mu) thus ?thesis by (smt add_absorb add_associative add_commutative add_left_dist_meet) qed -- {* AACP Theorem 8.16 *} lemma loop_semantics_apx_left_isotone: "w \ y \ \ (\x . w ; x + z) \ \ (\x . y ; x + z)" by (metis loop_semantics_kappa_mu_nu_2 combined_apx_left_isotone_2) -- {* AACP Theorem 8.16 *} lemma loop_semantics_apx_right_isotone: "w \ z \ \ (\x . y ; x + w) \ \ (\x . y ; x + z)" by (metis loop_semantics_kappa_mu_nu_2 combined_apx_right_isotone_2) lemma loop_semantics_apx_isotone: "v \ y \ w \ z \ \ (\x . v ; x + w) \ \ (\x . y ; x + z)" by (metis loop_semantics_kappa_mu_nu_2 combined_apx_isotone_2) end (* end theory NOmegaAlgebraBinaryItering imports NOmegaAlgebra BinaryIteringStrict begin *) sublocale extended_binary_itering < left_zero_conway_semiring where circ = "(\x . x \ 1)" apply unfold_locales apply (metis while_left_unfold) apply (metis mult_right_one while_one_mult_below while_slide) apply (metis while_one_while while_sumstar_2) done class binary_itering_apx = bounded_binary_itering + n_algebra_apx begin lemma C_while_import: "C (x \ z) = C (C x \ z)" proof - have 1: "C x ; (x \ z) \ C x \ (C x ; z)" by (metis C_mult_propagate eq_refl while_simulate) have "C (x \ z) = C z + C x ; (x \ z)" by (metis meet_left_dist_add n_L_T_meet_mult while_left_unfold) also have "... \ C x \ z" using 1 by (metis add_isotone meet.add_right_upper_bound while_right_unfold) finally have "C (x \ z) \ C (C x \ z)" by (metis meet.add_least_upper_bound meet.add_left_upper_bound) thus ?thesis by (smt2 meet.add_left_upper_bound meet.less_eq_def meet_associative meet_commutative while_left_isotone) qed lemma C_while_preserve: "C (x \ z) = C (x \ C z)" proof - have "C x ; (x \ z) \ C x \ (C x ; z)" by (metis C_mult_propagate eq_refl while_simulate) also have "... \ x \ (x ; C z)" by (metis C_decreasing n_L_T_meet_mult_propagate while_isotone) finally have 1: "C x ; (x \ z) \ x \ (x ; C z)" by metis have "C (x \ z) = C z + C x ; (x \ z)" by (metis meet_left_dist_add n_L_T_meet_mult while_left_unfold) also have "... \ x \ C z" using 1 by (metis add_least_upper_bound while_increasing while_mult_increasing while_mult_transitive) finally have "C (x \ z) \ C (x \ C z)" by (metis meet.add_least_upper_bound meet.add_left_upper_bound) thus ?thesis by (smt2 meet.add_left_upper_bound meet.less_eq_def meet_associative meet_commutative while_right_isotone) qed lemma C_while_import_preserve: "C (x \ z) = C (C x \ C z)" by (metis C_while_import C_while_preserve) lemma while_L_L: "L \ L = L" by (metis n_L_top_L while_mult_star_exchange while_right_top) lemma while_L_below_add: "L \ x \ x + L" by (metis while_left_unfold add_right_isotone n_L_below_L) lemma while_L_split: "x \ L \ (x \ y) + L" proof - have "x \ L \ (x \ 0) + L" by (metis add_commutative add_left_zero mult_right_one n_L_split_L while_right_unfold while_simulate_left_plus while_zero) thus ?thesis by (metis add_commutative add_right_isotone order_trans while_right_isotone zero_least) qed lemma while_n_while_top_split: "x \ (n(x \ y) ; T) \ (x \ 0) + n(x \ y) ; T" proof - have "x ; n(x \ y) ; T \ x ; 0 + n(x ; (x \ y)) ; T" by (metis n_n_top_split_n_top) also have "... \ n(x \ y) ; T + x ; 0" by (metis add_commutative add_right_isotone mult_left_isotone n_isotone while_left_plus_below) finally have "x \ (n(x \ y) ; T) \ n(x \ y) ; T + (x \ (x ; 0))" by (metis mult_associative mult_right_one while_simulate_left mult_left_zero while_left_top) also have "... \ (x \ 0) + n(x \ y) ; T" by (metis add_least_upper_bound add_left_isotone while_right_plus_below) finally show ?thesis by metis qed lemma circ_apx_right_isotone: "x \ y \ z \ x \ z \ y" proof assume "x \ y" hence 1: "x \ y + L \ C y \ x + n(x) ; T" by (metis apx_def) hence "z \ x \ (z \ y) + (z \ L)" by (metis while_left_dist_add while_right_isotone) hence 2: "z \ x \ (z \ y) + L" by (smt add_least_upper_bound add_left_upper_bound while_L_split order_trans) have "z \ (n(z \ x) ; T) \ (z \ 0) + n(z \ x) ; T" by (metis while_n_while_top_split) also have "... \ (z \ x) + n(z \ x) ; T" by (metis add_left_isotone while_right_isotone zero_least) finally have 3: "z \ (n(x) ; T) \ (z \ x) + n(z \ x) ; T" by (metis mult_left_isotone n_isotone order_trans while_increasing while_right_isotone) have "C (z \ y) \ z \ C y" by (metis C_while_preserve meet.add_right_divisibility) also have "... \ (z \ x) + (z \ (n(x) ; T))" using 1 by (metis while_left_dist_add while_right_isotone) also have "... \ (z \ x) + n(z \ x) ; T" using 3 by (metis add_least_upper_bound add_left_upper_bound) finally show "z \ x \ z \ y" using 2 by (metis apx_def) qed end class extended_binary_itering_apx = binary_itering_apx + bounded_extended_binary_itering + assumes n_below_while_zero: "n(x) \ n(x \ 0)" begin lemma circ_apx_right_isotone: "x \ y \ x \ z \ y \ z" proof assume "x \ y" hence 1: "x \ y + L \ C y \ x + n(x) ; T" by (metis apx_def) hence "x \ z \ ((y \ 1) ; L) \ (y \ z)" by (metis while_left_isotone while_sumstar_3) also have "... \ (y \ z) + (y \ 1) ; L" by (metis while_productstar add_right_isotone mult_right_isotone n_L_below_L while_slide) also have "... \ (y \ z) + L" by (metis add_commutative add_least_upper_bound add_right_upper_bound order_trans while_L_split while_one_mult_below) finally have 2: "x \ z \ (y \ z) + L" by metis have "C (y \ z) \ C y \ z" by (metis C_while_import meet.add_right_divisibility) also have "... \ ((x \ 1) ; n(x) ; T) \ (x \ z)" using 1 by (metis while_left_isotone mult_associative while_sumstar_3) also have "... \ (x \ z) + (x \ 1) ; n(x) ; T" by (metis while_productstar add_left_top add_right_isotone mult_associative mult_left_sub_dist_add_right while_slide) also have "... \ (x \ z) + (x \ (n(x) ; T))" by (metis add_right_isotone mult_associative while_one_mult_below) also have "... \ (x \ z) + (x \ (n(x \ z) ; T))" by (metis n_below_while_zero zero_least while_right_isotone n_isotone mult_left_isotone add_right_isotone order_trans) also have "... \ (x \ z) + n(x \ z) ; T" by (smt add_associative add_right_isotone while_n_while_top_split add_right_zero while_left_dist_add) finally show "x \ z \ y \ z" using 2 by (metis apx_def) qed lemma while_top: "T \ x = L + T ; x" nitpick [expect=genuine] oops lemma while_one_top: "1 \ x = L + x" nitpick [expect=genuine] oops lemma while_unfold_below_1: "x = y ; x \ x \ y \ 1" nitpick [expect=genuine] oops lemma while_square_1: "x \ 1 = (x ; x) \ (x + 1)" oops lemma while_absorb_below_one: "y ; x \ x \ y \ x \ 1 \ x" oops lemma while_mult_L: "(x ; L) \ z = z + x ; L" oops lemma tarski_top_omega_below_2: "x ; L \ (x ; L) \ 0" oops lemma tarski_top_omega_2: "x ; L = (x ; L) \ 0" oops lemma while_separate_right_plus: "y ; x \ x ; (x \ (1 + y)) + 1 \ y \ (x \ z) \ x \ (y \ z)" oops lemma "y \ (x \ 1) \ x \ (y \ 1) \ (x + y) \ 1 = x \ (y \ 1)" oops lemma "y ; x \ (1 + x) ; (y \ 1) \ (x + y) \ 1 = x \ (y \ 1)" oops end class n_omega_algebra_binary = n_omega_algebra + while + assumes while_def: "x \ y = n(x\<^sup>\) ; L + x\<^sup>\ ; y" begin lemma while_omega_meet_L_star: "x \ y = (x\<^sup>\ \ L) + x\<^sup>\ ; y" by (metis loop_semantics_kappa_mu_nu loop_semantics_kappa_mu_nu_2 while_def) lemma while_one_mult_while_below_1: "(y \ 1) ; (y \ v) \ y \ v" proof - have "(y \ 1) ; (y \ v) \ y \ (y \ v)" by (smt add_left_isotone mult_associative mult_right_dist_add mult_right_isotone n_L_below_L while_def mult_left_one) also have "... = n(y\<^sup>\) ; L + y\<^sup>\ ; n(y\<^sup>\) ; L + y\<^sup>\ ; y\<^sup>\ ; v" by (metis while_def mult_left_dist_add add_associative mult_associative) also have "... = n(y\<^sup>\) ; L + n(y\<^sup>\ ; y\<^sup>\) ; L + y\<^sup>\ ; y\<^sup>\ ; v" by (smt n_mult_omega_L_star_zero add_relative_same_increasing add_associative add_left_zero mult_left_sub_dist_add_left add_commutative) finally show ?thesis by (metis add_idempotent star.circ_transitive_equal star_mult_omega while_def) qed lemma star_below_while: "x\<^sup>\ ; y \ x \ y" by (metis add_right_upper_bound while_def) subclass bounded_binary_itering proof unfold_locales fix x y z have "z + x ; ((y ; x) \ (y ; z)) = x ; (y ; x)\<^sup>\ ; y ; z + x ; n((y ; x)\<^sup>\) ; L + z" by (smt add_associative add_commutative mult_associative mult_left_dist_add while_def) also have "... = x ; (y ; x)\<^sup>\ ; y ; z + n(x ; (y ; x)\<^sup>\) ; L + z" by (metis mult_associative mult_right_isotone zero_least n_mult_omega_L_star_zero add_relative_same_increasing) also have "... = (x ; y)\<^sup>\ ; z + n(x ; (y ; x)\<^sup>\) ; L" by (smt add_associative add_commutative mult_associative star.circ_loop_fixpoint star_slide) also have "... = (x ; y) \ z" by (smt omega_slide while_def add_commutative) finally show "(x ; y) \ z = z + x ; ((y ; x) \ (y ; z))" by metis next fix x y z have "(x \ y) \ (x \ z) = n((n(x\<^sup>\) ; L + x\<^sup>\ ; y)\<^sup>\) ; L + (n(x\<^sup>\) ; L + x\<^sup>\ ; y)\<^sup>\ ; (x \ z)" by (metis while_def) also have "... = n((x\<^sup>\ ; y)\<^sup>\ + (x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L) ; L + ((x\<^sup>\ ; y)\<^sup>\ + (x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L) ; (x \ z)" by (metis mult_L_add_star mult_L_add_omega) also have "... = n((x\<^sup>\ ; y)\<^sup>\) ; L + n((x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L) ; L + (x\<^sup>\ ; y)\<^sup>\ ; (x \ z) + (x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L ; (x \ z)" by (metis mult_associative n_dist_omega_star mult_right_dist_add add_associative) also have "... = n((x\<^sup>\ ; y)\<^sup>\) ; L + n((x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L) ; L + (x\<^sup>\ ; y)\<^sup>\ ; 0 + (x\<^sup>\ ; y)\<^sup>\ ; (x \ z) + (x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L ; (x \ z)" by (smt add_associative add_left_zero mult_left_dist_add) also have "... = n((x\<^sup>\ ; y)\<^sup>\) ; L + ((x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L ; (x \ z) + (x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L + (x\<^sup>\ ; y)\<^sup>\ ; (x \ z))" by (smt n_n_L_split_n_n_L_L add_commutative add_associative) also have "... = n((x\<^sup>\ ; y)\<^sup>\) ; L + ((x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L + (x\<^sup>\ ; y)\<^sup>\ ; (x \ z))" by (smt mult_L_omega omega_sub_vector less_eq_def) also have "... = n((x\<^sup>\ ; y)\<^sup>\) ; L + (x\<^sup>\ ; y)\<^sup>\ ; (x \ z)" by (metis add_left_divisibility mult_associative mult_right_isotone while_def less_eq_def) also have "... = (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ ; z + (x\<^sup>\ ; y)\<^sup>\ ; n(x\<^sup>\) ; L + n((x\<^sup>\ ; y)\<^sup>\) ; L" by (metis add_commutative mult_associative mult_left_dist_add while_def) also have "... = (x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\ ; z + n((x\<^sup>\ ; y)\<^sup>\ ; x\<^sup>\) ; L + n((x\<^sup>\ ; y)\<^sup>\) ; L" by (metis add_right_zero mult_left_dist_add add_associative n_mult_omega_L_star_zero) also have "... = (x + y) \ z" by (metis add_associative add_commutative omega_decompose star.circ_add while_def mult_right_dist_add n_dist_omega_star) finally show "(x + y) \ z = (x \ y) \ (x \ z)" by metis next fix x y z show "x \ (y + z) = (x \ y) + (x \ z)" by (smt add_associative add_commutative add_left_upper_bound less_eq_def mult_left_dist_add while_def) next fix x y z show "(x \ y) ; z \ x \ (y ; z)" by (smt add_left_isotone mult_associative mult_right_dist_add mult_right_isotone n_L_below_L while_def) next fix v w x y z show "x ; z \ z ; (y \ 1) + w \ x \ (z ; v) \ z ; (y \ v) + (x \ (w ; (y \ v)))" proof assume 1: "x ; z \ z ; (y \ 1) + w" have "z ; v + x ; (z ; (y \ v) + x\<^sup>\ ; (w ; (y \ v))) \ z ; v + x ; z ; (y \ v) + x\<^sup>\ ; (w ; (y \ v))" by (metis add_associative add_right_isotone mult_associative mult_left_dist_add mult_left_isotone star.left_plus_below_circ) also have "... \ z ; v + z ; (y \ 1) ; (y \ v) + w ; (y \ v) + x\<^sup>\ ; (w ; (y \ v))" using 1 by (metis add_associative add_left_isotone add_right_isotone mult_left_isotone mult_right_dist_add) also have "... \ z ; v + z ; (y \ v) + x\<^sup>\ ; (w ; (y \ v))" by (smt add_least_upper_bound add_right_upper_bound less_eq_def mult_associative mult_left_dist_add star.circ_loop_fixpoint while_one_mult_while_below_1) also have "... = z ; (y \ v) + x\<^sup>\ ; (w ; (y \ v))" by (metis less_eq_def mult_left_dist_add mult_left_one mult_right_sub_dist_add_left order_trans star.circ_plus_one star_below_while) finally have "x\<^sup>\ ; z ; v \ z ; (y \ v) + x\<^sup>\ ; (w ; (y \ v))" by (metis mult_associative star_left_induct) thus "x \ (z ; v) \ z ; (y \ v) + (x \ (w ; (y \ v)))" by (smt add_associative add_commutative add_right_isotone mult_associative while_def) qed next fix v w x y z show "z ; x \ y ; (y \ z) + w \ z ; (x \ v) \ y \ (z ; v + w ; (x \ v))" proof assume "z ; x \ y ; (y \ z) + w" hence 1: "z ; x \ y ; y\<^sup>\ ; z + (y ; n(y\<^sup>\) ; L + w)" by (smt add_associative add_commutative mult_associative mult_left_dist_add while_def) hence "z ; x\<^sup>\ \ y\<^sup>\ ; (z + (y ; n(y\<^sup>\) ; L + w) ; x\<^sup>\)" by (metis star.circ_simulate_right_plus) also have "... = y\<^sup>\ ; z + y\<^sup>\ ; y ; n(y\<^sup>\) ; L + y\<^sup>\ ; w ; x\<^sup>\" by (smt add_associative mult_associative mult_left_dist_add mult_right_dist_add L_mult_star) also have "... = y\<^sup>\ ; z + n(y\<^sup>\ ; y ; y\<^sup>\) ; L + y\<^sup>\ ; w ; x\<^sup>\" by (metis add_relative_same_increasing mult_isotone n_mult_omega_L_star_zero star.left_plus_below_circ star.right_plus_circ zero_least) also have "... = n(y\<^sup>\) ; L + y\<^sup>\ ; z + y\<^sup>\ ; w ; x\<^sup>\" by (metis add_commutative omega_unfold right_plus_omega) finally have "z ; x\<^sup>\ ; v \ n(y\<^sup>\) ; L ; v + y\<^sup>\ ; z ; v + y\<^sup>\ ; w ; x\<^sup>\ ; v" by (smt less_eq_def mult_right_dist_add) also have "... \ n(y\<^sup>\) ; L + y\<^sup>\ ; (z ; v + w ; x\<^sup>\ ; v)" by (metis n_L_below_L mult_associative mult_right_isotone add_left_isotone mult_left_dist_add add_associative) also have "... \ n(y\<^sup>\) ; L + y\<^sup>\ ; (z ; v + w ; (x \ v))" by (metis add_commutative add_right_isotone mult_associative mult_left_sub_dist_add_left mult_right_isotone while_def) finally have 2: "z ; x\<^sup>\ ; v \ y \ (z ; v + w ; (x \ v))" by (metis while_def) have 3: "y\<^sup>\ ; y ; y\<^sup>\ ; 0 \ y\<^sup>\ ; w ; x\<^sup>\" by (metis add_commutative add_left_zero mult_associative mult_left_sub_dist_add_left star.circ_loop_fixpoint star.circ_transitive_equal) have "z ; x\<^sup>\ \ y ; y\<^sup>\ ; z ; x\<^sup>\ + (y ; n(y\<^sup>\) ; L + w) ; x\<^sup>\" using 1 by (metis mult_associative mult_left_isotone mult_right_dist_add omega_unfold) hence "z ; x\<^sup>\ \ y\<^sup>\ + y\<^sup>\ ; y ; n(y\<^sup>\) ; L ; x\<^sup>\ + y\<^sup>\ ; w ; x\<^sup>\" by (smt add_associative add_commutative left_plus_omega mult_associative mult_left_dist_add mult_right_dist_add omega_induct star.left_plus_circ) also have "... \ y\<^sup>\ + y\<^sup>\ ; y ; n(y\<^sup>\) ; L + y\<^sup>\ ; w ; x\<^sup>\" by (metis add_left_isotone add_right_isotone mult_associative mult_right_isotone n_L_below_L) also have "... = y\<^sup>\ + n(y\<^sup>\ ; y ; y\<^sup>\) ; L + y\<^sup>\ ; w ; x\<^sup>\" using 3 by (smt add_associative add_commutative add_relative_same_increasing n_mult_omega_L_star_zero) also have "... = y\<^sup>\ + y\<^sup>\ ; w ; x\<^sup>\" by (metis mult_associative omega_unfold star_mult_omega add_commutative less_eq_def n_L_decreasing) finally have "n(z ; x\<^sup>\) ; L \ n(y\<^sup>\) ; L + n(y\<^sup>\ ; w ; x\<^sup>\) ; L" by (metis mult_associative mult_left_isotone mult_right_dist_add n_dist_omega_star n_isotone) also have "... \ n(y\<^sup>\) ; L + y\<^sup>\ ; (w ; (n(x\<^sup>\) ; L + x\<^sup>\ ; 0))" by (smt add_commutative add_right_isotone mult_associative mult_left_dist_add n_mult_omega_L_below_zero) also have "... \ n(y\<^sup>\) ; L + y\<^sup>\ ; (w ; (n(x\<^sup>\) ; L + x\<^sup>\ ; v))" by (metis add_right_isotone mult_right_isotone zero_least) also have "... \ n(y\<^sup>\) ; L + y\<^sup>\ ; (z ; v + w ; (n(x\<^sup>\) ; L + x\<^sup>\ ; v))" by (metis add_right_isotone mult_left_sub_dist_add_right) finally have 4: "n(z ; x\<^sup>\) ; L \ y \ (z ; v + w ; (x \ v))" by (metis while_def) have "z ; (x \ v) = z ; n(x\<^sup>\) ; L + z ; x\<^sup>\ ; v" by (metis while_def mult_left_dist_add mult_associative) also have "... = n(z ; x\<^sup>\) ; L + z ; x\<^sup>\ ; v" by (metis add_commutative add_relative_same_increasing mult_right_isotone n_mult_omega_L_star_zero zero_least) finally show "z ; (x \ v) \ y \ (z ; v + w ; (x \ v))" using 2 4 by (metis add_least_upper_bound) qed qed lemma while_top: "T \ x = L + T ; x" by (metis n_top_L star.circ_top star_omega_top while_def) lemma while_one_top: "1 \ x = L + x" by (smt mult_left_one n_top_L omega_one star_one while_def) lemma while_finite_associative: "x\<^sup>\ = 0 \ (x \ y) ; z = x \ (y ; z)" by (metis add_left_zero mult_associative n_zero_L_zero while_def) lemma while_while_one: "y \ (x \ 1) = n(y\<^sup>\) ; L + y\<^sup>\ ; n(x\<^sup>\) ; L + y\<^sup>\ ; x\<^sup>\" by (metis add_associative mult_left_dist_add mult_right_one while_def mult_associative) -- {* AACP Theorem 8.17 *} subclass bounded_extended_binary_itering proof unfold_locales fix w x y z have "w ; (x \ y ; z) = n(w ; n(x\<^sup>\) ; L) ; L + w ; x\<^sup>\ ; y ; z" by (smt add_associative add_commutative add_left_zero mult_associative mult_left_dist_add n_n_L_split_n_n_L_L while_def) also have "... \ n((w ; n(x\<^sup>\) ; L)\<^sup>\) ; L + w ; x\<^sup>\ ; y ; z" by (metis eq_refl mult_L_omega) also have "... \ n((w ; (x \ y))\<^sup>\) ; L + w ; x\<^sup>\ ; y ; z" by (smt add_left_isotone add_left_upper_bound mult_associative mult_left_isotone mult_right_isotone n_isotone omega_isotone while_def) also have "... \ n((w ; (x \ y))\<^sup>\) ; L + w ; (x \ y) ; z" by (metis star_below_while mult_associative mult_left_isotone mult_right_isotone add_right_isotone) also have "... \ n((w ; (x \ y))\<^sup>\) ; L + (w ; (x \ y))\<^sup>\ ; (w ; (x \ y) ; z)" by (metis add_right_isotone add_right_upper_bound star.circ_loop_fixpoint) finally show "w ; (x \ y ; z) \ (w ; (x \ y)) \ (w ; (x \ y) ; z)" by (metis while_def) qed subclass extended_binary_itering_apx apply unfold_locales apply (metis n_below_n_omega n_left_upper_bound n_n_L order_trans while_def) done lemma while_simulate_4_plus: "y ; x \ x ; (x \ (1 + y)) \ y ; x ; x\<^sup>\ \ x ; (x \ (1 + y))" proof assume 1: "y ; x \ x ; (x \ (1 + y))" have "x ; (x \ (1 + y)) = x ; n(x\<^sup>\) ; L + x ; x\<^sup>\ ; (1 + y)" by (metis mult_associative mult_left_dist_add while_def) also have "... = n(x ; x\<^sup>\) ; L + x ; x\<^sup>\ ; (1 + y)" by (smt n_mult_omega_L_star_zero add_relative_same_increasing add_commutative add_right_zero mult_left_sub_dist_add_right) finally have 2: "x ; (x \ (1 + y)) = n(x\<^sup>\) ; L + x ; x\<^sup>\ + x ; x\<^sup>\ ; y" by (metis add_associative mult_left_dist_add mult_right_one omega_unfold) hence "x ; x\<^sup>\ ; y ; x \ x ; x\<^sup>\ ; n(x\<^sup>\) ; L + x ; x\<^sup>\ ; x\<^sup>\ ; x + x ; x\<^sup>\ ; x ; x\<^sup>\ ; y" using 1 by (metis mult_associative mult_right_isotone mult_left_dist_add star_plus) also have "... = n(x ; x\<^sup>\ ; x\<^sup>\) ; L + x ; x\<^sup>\ ; x\<^sup>\ ; x + x ; x\<^sup>\ ; x ; x\<^sup>\ ; y" by (smt n_mult_omega_L_star_zero add_relative_same_increasing add_commutative add_right_zero mult_left_sub_dist_add_right) also have "... = n(x\<^sup>\) ; L + x ; x\<^sup>\ ; x + x ; x ; x\<^sup>\ ; y" by (metis mult_associative omega_unfold star.circ_plus_same star.circ_transitive_equal star_mult_omega) also have "... \ n(x\<^sup>\) ; L + x ; x\<^sup>\ + x ; x\<^sup>\ ; y" by (smt add_associative add_right_upper_bound less_eq_def mult_associative mult_right_dist_add star.circ_increasing star.circ_plus_same star.circ_transitive_equal) finally have 3: "x ; x\<^sup>\ ; y ; x \ n(x\<^sup>\) ; L + x ; x\<^sup>\ + x ; x\<^sup>\ ; y" by metis have "(n(x\<^sup>\) ; L + x ; x\<^sup>\ + x ; x\<^sup>\ ; y) ; x \ n(x\<^sup>\) ; L + x ; x\<^sup>\ ; x + x ; x\<^sup>\ ; y ; x" by (metis mult_right_dist_add n_L_below_L mult_associative mult_right_isotone add_left_isotone) also have "... \ n(x\<^sup>\) ; L + x ; x\<^sup>\ + x ; x\<^sup>\ ; y ; x" by (smt add_commutative add_left_isotone mult_associative mult_right_isotone star.left_plus_below_circ star_plus) also have "... \ n(x\<^sup>\) ; L + x ; x\<^sup>\ + x ; x\<^sup>\ ; y" using 3 by (metis add_least_upper_bound add_left_upper_bound) finally show "y ; x ; x\<^sup>\ \ x ; (x \ (1 + y))" using 1 2 by (metis add_least_upper_bound star_right_induct) qed lemma while_simulate_4_omega: "y ; x \ x ; (x \ (1 + y)) \ y ; x\<^sup>\ \ x\<^sup>\" proof assume 1: "y ; x \ x ; (x \ (1 + y))" have "x ; (x \ (1 + y)) = x ; n(x\<^sup>\) ; L + x ; x\<^sup>\ ; (1 + y)" by (metis mult_associative mult_left_dist_add while_def) also have "... = n(x ; x\<^sup>\) ; L + x ; x\<^sup>\ ; (1 + y)" by (smt n_mult_omega_L_star_zero add_relative_same_increasing add_commutative add_right_zero mult_left_sub_dist_add_right) finally have "x ; (x \ (1 + y)) = n(x\<^sup>\) ; L + x ; x\<^sup>\ + x ; x\<^sup>\ ; y" by (metis add_associative mult_left_dist_add mult_right_one omega_unfold) hence "y ; x\<^sup>\ \ n(x\<^sup>\) ; L ; x\<^sup>\ + x ; x\<^sup>\ ; x\<^sup>\ + x ; x\<^sup>\ ; y ; x\<^sup>\" using 1 by (smt less_eq_def mult_associative mult_right_dist_add omega_unfold) also have "... \ x ; x\<^sup>\ ; (y ; x\<^sup>\) + x\<^sup>\" by (metis add_left_isotone mult_L_omega omega_sub_vector mult_associative omega_unfold star_mult_omega n_L_decreasing less_eq_def add_commutative) finally have "y ; x\<^sup>\ \ (x ; x\<^sup>\)\<^sup>\ + (x ; x\<^sup>\)\<^sup>\ ; x\<^sup>\" by (metis add_commutative omega_induct) thus "y ; x\<^sup>\ \ x\<^sup>\" by (metis add_idempotent left_plus_omega star_mult_omega) qed lemma while_square_1: "x \ 1 = (x ; x) \ (x + 1)" by (metis mult_right_one omega_square star_square_2 while_def) lemma while_absorb_below_one: "y ; x \ x \ y \ x \ 1 \ x" by (metis star_left_induct_mult add_isotone n_galois n_sub_nL while_def while_one_top) lemma while_mult_L: "(x ; L) \ z = z + x ; L" by (metis add_right_zero mult_left_zero while_denest_5 while_one_top while_productstar while_sumstar) lemma tarski_top_omega_below_2: "x ; L \ (x ; L) \ 0" by (metis add_right_divisibility while_mult_L) lemma tarski_top_omega_2: "x ; L = (x ; L) \ 0" by (metis add_left_zero while_mult_L) lemma while_sub_mult_one: "x ; (1 \ y) \ 1 \ x" nitpick [expect=genuine] oops lemma while_unfold_below: "x = z + y ; x \ x \ y \ z" nitpick [expect=genuine] oops lemma while_loop_is_greatest_postfixpoint: "is_greatest_postfixpoint (\x . y ; x + z) (y \ z)" nitpick [expect=genuine] oops lemma while_loop_is_greatest_fixpoint: "is_greatest_fixpoint (\x . y ; x + z) (y \ z)" nitpick [expect=genuine] oops lemma while_denest_3: "(x \ w) \ x\<^sup>\ = (x \ w)\<^sup>\" nitpick [expect=genuine] oops lemma while_mult_top: "(x ; T) \ z = z + x ; T" nitpick [expect=genuine] oops lemma tarski_below_top_omega: "x \ (x ; L)\<^sup>\" nitpick [expect=genuine] oops lemma tarski_mult_omega_omega: "(x ; y\<^sup>\)\<^sup>\ = x ; y\<^sup>\" nitpick [expect=genuine] oops lemma tarski_below_top_omega_2: "x \ (x ; L) \ 0" nitpick [expect=genuine] oops lemma "1 = (x ; 0) \ 1" nitpick [expect=genuine] oops lemma tarski: "x = 0 \ T ; x ; T = T" nitpick [expect=genuine] oops lemma "(x + y) \ z = ((x \ 1) ; y) \ ((x \ 1) ; z)" nitpick [expect=genuine] oops lemma while_top_2: "T \ z = T ; z" nitpick [expect=genuine] oops lemma while_mult_top_2: "(x ; T) \ z = z + x ; T ; z" nitpick [expect=genuine] oops lemma while_one_mult: "(x \ 1) ; x = x \ x" nitpick [expect=genuine] oops lemma "(x \ 1) ; y = x \ y" nitpick [expect=genuine] oops lemma while_associative: "(x \ y) ; z = x \ (y ; z)" nitpick [expect=genuine] oops lemma while_back_loop_is_fixpoint: "is_fixpoint (\x . x ; y + z) (z ; (y \ 1))" nitpick [expect=genuine] oops lemma "1 + x ; 0 = x \ 1" nitpick [expect=genuine] oops lemma "x = x ; (x \ 1)" nitpick [expect=genuine] oops lemma "x ; (x \ 1) = x \ 1" nitpick [expect=genuine] oops lemma "x \ 1 = x \ (1 \ 1)" nitpick [expect=genuine] oops lemma "(x + y) \ 1 = (x \ (y \ 1)) \ 1" nitpick [expect=genuine] oops lemma "z + y ; x = x \ y \ z \ x" nitpick [expect=genuine] oops lemma "y ; x = x \ y \ x \ x" nitpick [expect=genuine] oops lemma "z + x ; y = x \ z ; (y \ 1) \ x" nitpick [expect=genuine] oops lemma "x ; y = x \ x ; (y \ 1) \ x" nitpick [expect=genuine] oops lemma "x ; z = z ; y \ x \ z \ z ; (y \ 1)" nitpick [expect=genuine] oops lemma while_unfold_below_1: "x = y ; x \ x \ y \ 1" nitpick [expect=genuine] oops lemma "x\<^sup>\ \ x\<^sup>\ ; x\<^sup>\" oops lemma tarski_omega_idempotent: "x\<^sup>\\<^sup>\ = x\<^sup>\" oops end class n_omega_algebra_binary_strict = n_omega_algebra_binary + circ + assumes L_left_zero: "L ; x = L" assumes circ_def: "x\<^sup>\ = n(x\<^sup>\) ; L + x\<^sup>\" begin subclass strict_binary_itering apply unfold_locales apply (metis while_def mult_associative L_left_zero mult_right_dist_add) apply (metis circ_def while_def mult_right_one) done end (* end theory CappedOmega imports OmegaAlgebra begin *) class capped_omega = fixes capped_omega :: "'a \ 'a \ 'a" ("_\<^sup>\\<^sub>_" [100,100] 100) class capped_omega_algebra = bounded_left_zero_kleene_algebra + bounded_distributive_lattice + capped_omega + assumes capped_omega_unfold: "y\<^sup>\\<^sub>v = y ; y\<^sup>\\<^sub>v \ v" assumes capped_omega_induct: "x \ (y ; x + z) \ v \ x \ y\<^sup>\\<^sub>v + y\<^sup>\ ; z" -- {* AACP Theorem 6.1 *} sublocale capped_omega_algebra < capped!: bounded_left_zero_omega_algebra where omega = "(\y . y\<^sup>\\<^sub>T)" apply unfold_locales apply (metis capped_omega_unfold meet.add_right_zero) apply (metis add_commutative capped_omega_induct meet.add_right_zero) done context capped_omega_algebra begin -- {* AACP Theorem 6.2 *} lemma capped_omega_below_omega: "y\<^sup>\\<^sub>v \ y\<^sup>\\<^sub>T" by (metis capped.omega_induct_mult capped_omega_unfold meet.add_left_upper_bound) -- {* AACP Theorem 6.3 *} lemma capped_omega_below: "y\<^sup>\\<^sub>v \ v" by (metis capped_omega_unfold meet.add_left_divisibility meet_commutative) -- {* AACP Theorem 6.4 *} lemma capped_omega_one: "1\<^sup>\\<^sub>v = v" proof - have "v \ (1 ; v + 0) \ v" by (metis add_right_zero meet_idempotent mult_left_one order.refl) hence "v \ 1\<^sup>\\<^sub>v + 1\<^sup>\ ; 0" by (metis capped_omega_induct) also have "... = 1\<^sup>\\<^sub>v" by (metis add_right_zero mult_left_one star_one) finally show ?thesis by (metis capped_omega_below antisym) qed -- {* AACP Theorem 6.5 *} lemma capped_omega_zero: "0\<^sup>\\<^sub>v = 0" by (metis capped_omega_unfold meet_commutative meet_right_zero mult_left_zero) lemma star_below_cap: "y \ u \ z \ v \ u ; v \ v \ y\<^sup>\ ; z \ v" by (metis add_least_upper_bound dual_order.trans mult_left_isotone star_left_induct) lemma capped_fix: "y \ u \ z \ v \ u ; v \ v \ (y ; (y\<^sup>\\<^sub>v + y\<^sup>\ ; z) + z) \ v = y\<^sup>\\<^sub>v + y\<^sup>\ ; z" proof assume 1: "y \ u \ z \ v \ u ; v \ v" have "(y ; (y\<^sup>\\<^sub>v + y\<^sup>\ ; z) + z) \ v = (y ; y\<^sup>\\<^sub>v + y\<^sup>\ ; z) \ v" by (metis add_associative mult_left_dist_add star.circ_loop_fixpoint) also have "... = (y ; y\<^sup>\\<^sub>v \ v) + (y\<^sup>\ ; z \ v)" by (metis meet_commutative meet_left_dist_add) also have "... = y\<^sup>\\<^sub>v + y\<^sup>\ ; z" using 1 by (metis capped_omega_unfold meet_less_eq_def star_below_cap) finally show "(y ; (y\<^sup>\\<^sub>v + y\<^sup>\ ; z) + z) \ v = y\<^sup>\\<^sub>v + y\<^sup>\ ; z" . qed lemma capped_fixpoint: "y \ u \ z \ v \ u ; v \ v \ is_fixpoint (\x . (y ; x + z) \ v) (y\<^sup>\\<^sub>v + y\<^sup>\ ; z)" by (metis capped_fix meet.is_fixpoint_def) lemma capped_greatest_fixpoint: "y \ u \ z \ v \ u ; v \ v \ is_greatest_fixpoint (\x . (y ; x + z) \ v) (y\<^sup>\\<^sub>v + y\<^sup>\ ; z)" by (smt2 capped_fix order_refl capped_omega_induct is_greatest_fixpoint_def) lemma capped_postfixpoint: "y \ u \ z \ v \ u ; v \ v \ is_postfixpoint (\x . (y ; x + z) \ v) (y\<^sup>\\<^sub>v + y\<^sup>\ ; z)" by (metis capped_fix eq_refl is_postfixpoint_def) lemma capped_greatest_postfixpoint: "y \ u \ z \ v \ u ; v \ v \ is_greatest_postfixpoint (\x . (y ; x + z) \ v) (y\<^sup>\\<^sub>v + y\<^sup>\ ; z)" by (smt2 capped_fix order_refl capped_omega_induct is_greatest_postfixpoint_def) -- {* AACP Theorem 6.6 *} lemma capped_nu: "y \ u \ z \ v \ u ; v \ v \ \(\x . (y ; x + z) \ v) = y\<^sup>\\<^sub>v + y\<^sup>\ ; z" by (metis capped_greatest_fixpoint greatest_fixpoint_same) lemma capped_pnu: "y \ u \ z \ v \ u ; v \ v \ p\(\x . (y ; x + z) \ v) = y\<^sup>\\<^sub>v + y\<^sup>\ ; z" by (metis capped_greatest_postfixpoint greatest_postfixpoint_same) -- {* AACP Theorem 6.7 *} lemma unfold_capped_omega: "y \ u \ u ; v \ v \ y ; y\<^sup>\\<^sub>v = y\<^sup>\\<^sub>v" by (metis capped_omega_below meet.order_trans mult_isotone capped_omega_unfold meet_less_eq_def) -- {* AACP Theorem 6.8 *} lemma star_mult_capped_omega: "y \ u \ u ; v \ v \ y\<^sup>\ ; y\<^sup>\\<^sub>v = y\<^sup>\\<^sub>v" proof assume "y \ u \ u ; v \ v" hence "y ; y\<^sup>\\<^sub>v = y\<^sup>\\<^sub>v" by (metis unfold_capped_omega) hence "y\<^sup>\ ; y\<^sup>\\<^sub>v \ y\<^sup>\\<^sub>v" by (metis star_left_induct_mult_equal) thus "y\<^sup>\ ; y\<^sup>\\<^sub>v = y\<^sup>\\<^sub>v" by (metis add_right_upper_bound antisym_conv star.circ_loop_fixpoint) qed -- {* AACP Theorem 6.9 *} lemma star_zero_below_capped_omega_zero: "y \ u \ u ; v \ v \ y\<^sup>\ ; 0 \ y\<^sup>\\<^sub>v ; 0" proof assume "y \ u \ u ; v \ v" hence "y ; y\<^sup>\\<^sub>v \ v" by (metis capped_omega_below meet.order_trans mult_isotone) hence "y ; y\<^sup>\\<^sub>v = y\<^sup>\\<^sub>v" by (metis capped_omega_unfold meet_less_eq_def) thus "y\<^sup>\ ; 0 \ y\<^sup>\\<^sub>v ; 0" by (metis add_commutative add_left_zero mult_associative star_loop_least_fixpoint) qed lemma star_zero_below_capped_omega: "y \ u \ u ; v \ v \ y\<^sup>\ ; 0 \ y\<^sup>\\<^sub>v" by (metis order.trans zero_right_mult_decreasing star_zero_below_capped_omega_zero) lemma capped_omega_induct_meet_zero: "x \ y ; x \ v \ x \ y\<^sup>\\<^sub>v + y\<^sup>\ ; 0" by (metis add_commutative add_left_zero capped_omega_induct) -- {* AACP Theorem 6.10 *} lemma capped_omega_induct_meet: "y \ u \ u ; v \ v \ x \ y ; x \ v \ x \ y\<^sup>\\<^sub>v" by (metis capped_omega_induct_meet_zero add_commutative less_eq_def star_zero_below_capped_omega) lemma capped_omega_induct_equal: "x = (y ; x + z) \ v \ x \ y\<^sup>\\<^sub>v + y\<^sup>\ ; z" by (metis capped_omega_induct order_refl) -- {* AACP Theorem 6.11 *} lemma capped_meet_nu: "y \ u \ u ; v \ v \ \(\x . y ; x \ v) = y\<^sup>\\<^sub>v" proof assume 1: "y \ u \ u ; v \ v" hence "y\<^sup>\\<^sub>v + y\<^sup>\ ; 0 = y\<^sup>\\<^sub>v" by (smt star_zero_below_capped_omega less_eq_def add_commutative) hence "\(\x . (y ; x + 0) \ v) = y\<^sup>\\<^sub>v" using 1 by (metis capped_nu zero_least) thus "\(\x . y ; x \ v) = y\<^sup>\\<^sub>v" by (simp add: add_right_zero) qed lemma capped_meet_pnu: "y \ u \ u ; v \ v \ p\(\x . y ; x \ v) = y\<^sup>\\<^sub>v" proof assume 1: "y \ u \ u ; v \ v" hence "y\<^sup>\\<^sub>v + y\<^sup>\ ; 0 = y\<^sup>\\<^sub>v" by (smt star_zero_below_capped_omega less_eq_def add_commutative) hence "p\(\x . (y ; x + 0) \ v) = y\<^sup>\\<^sub>v" using 1 by (metis capped_pnu zero_least) thus "p\(\x . y ; x \ v) = y\<^sup>\\<^sub>v" by (simp add: add_right_zero) qed -- {* AACP Theorem 6.12 *} lemma capped_omega_isotone: "y \ u \ u ; v \ v \ t \ y \ t\<^sup>\\<^sub>v \ y\<^sup>\\<^sub>v" by (metis capped_omega_induct_meet capped_omega_unfold less_eq_def meet.add_left_isotone mult_right_sub_dist_add_left) -- {* AACP Theorem 6.13 *} lemma capped_omega_simulation: "y \ u \ s \ u \ u ; v \ v \ s ; t \ y ; s \ s ; t\<^sup>\\<^sub>v \ y\<^sup>\\<^sub>v" proof assume 1: "y \ u \ s \ u \ u ; v \ v" show "s ; t \ y ; s \ s ; t\<^sup>\\<^sub>v \ y\<^sup>\\<^sub>v" proof assume 2: "s ; t \ y ; s" have "s ; t\<^sup>\\<^sub>v \ s ; t ; t\<^sup>\\<^sub>v \ s ; v" by (metis capped_omega_unfold meet.add_least_upper_bound meet.add_right_upper_bound meet_commutative mult_associative mult_right_isotone) also have "... \ s ; t ; t\<^sup>\\<^sub>v \ v" using 1 by (metis meet.add_left_isotone meet_commutative mult_left_isotone order_trans) also have "... \ y ; s ; t\<^sup>\\<^sub>v \ v" using 2 by (metis meet.add_left_isotone mult_left_isotone) finally show "s ; t\<^sup>\\<^sub>v \ y\<^sup>\\<^sub>v" using 1 by (metis capped_omega_induct_meet mult_associative) qed qed lemma capped_omega_slide_sub: "s \ u \ y \ u \ u ; u \ u \ u ; v \ v \ s ; (y ; s)\<^sup>\\<^sub>v \ (s ; y)\<^sup>\\<^sub>v" proof assume 1: "s \ u \ y \ u \ u ; u \ u \ u ; v \ v" hence "s ; y \ u" by (metis meet.order_trans mult_isotone) thus "s ; (y ; s)\<^sup>\\<^sub>v \ (s ; y)\<^sup>\\<^sub>v" using 1 by (metis mult_associative capped_omega_simulation order_refl) qed -- {* AACP Theorem 6.14 *} lemma capped_omega_slide: "s \ u \ y \ u \ u ; u \ u \ u ; v \ v \ s ; (y ; s)\<^sup>\\<^sub>v = (s ; y)\<^sup>\\<^sub>v" by (smt antisym mult_associative mult_right_isotone capped_omega_unfold capped_omega_slide_sub meet.add_left_upper_bound order_trans) lemma capped_omega_sub_dist: "s \ u \ y \ u \ u ; v \ v \ s\<^sup>\\<^sub>v \ (s + y)\<^sup>\\<^sub>v" by (metis capped_omega_isotone add_least_upper_bound add_left_upper_bound) -- {* AACP Theorem 6.15 *} lemma capped_omega_simulation_2: "s \ u \ y \ u \ u ; u \ u \ u ; v \ v \ y ; s \ s ; y \ (s ; y)\<^sup>\\<^sub>v \ s\<^sup>\\<^sub>v" proof assume 1: "s \ u \ y \ u \ u ; u \ u \ u ; v \ v" hence 2: "s ; y \ u" by (metis mult_isotone order.trans) have 3: "s ; (s ; y)\<^sup>\\<^sub>v \ v" using 1 by (metis capped_omega_below meet.order_trans mult_isotone) show "y ; s \ s ; y \ (s ; y)\<^sup>\\<^sub>v \ s\<^sup>\\<^sub>v" proof assume 4: "y ; s \ s ; y" have "(s ; y)\<^sup>\\<^sub>v = s ; (y ; s)\<^sup>\\<^sub>v" using 1 by (metis capped_omega_slide) also have "... \ s ; (s ; y)\<^sup>\\<^sub>v" using 1 2 4 by (smt less_eq_def mult_right_isotone capped_omega_sub_dist order_trans) also have "... = s ; (s ; y)\<^sup>\\<^sub>v \ v" using 3 by (metis meet.less_eq_def meet_commutative) finally show "(s ; y)\<^sup>\\<^sub>v \ s\<^sup>\\<^sub>v" using 1 by (metis capped_omega_induct_meet) qed qed -- {* AACP Theorem 6.16 *} lemma left_plus_capped_omega: "y \ u \ u ; u \ u \ u ; v \ v \ (y ; y\<^sup>\)\<^sup>\\<^sub>v = y\<^sup>\\<^sub>v" proof assume 1: "y \ u \ u ; u \ u \ u ; v \ v" hence 2: "y ; y\<^sup>\ \ u" by (metis star_plus star_below_cap) hence "y ; y\<^sup>\ ; (y ; y\<^sup>\)\<^sup>\\<^sub>v \ v" using 1 by (metis capped_omega_below meet.order_trans mult_isotone) hence "y ; y\<^sup>\ ; (y ; y\<^sup>\)\<^sup>\\<^sub>v = (y ; y\<^sup>\)\<^sup>\\<^sub>v" by (metis meet.less_eq_def meet_commutative capped_omega_unfold) hence "(y ; y\<^sup>\)\<^sup>\\<^sub>v \ y\<^sup>\\<^sub>v" using 1 2 by (smt2 capped_omega_simulation mult_associative mult_semi_associative star.circ_transitive_equal star_simulation_right_equal) thus "(y ; y\<^sup>\)\<^sup>\\<^sub>v = y\<^sup>\\<^sub>v" using 1 2 by (metis antisym capped_omega_isotone star.circ_mult_increasing) qed -- {* AACP Theorem 6.17 *} lemma capped_omega_sub_vector: "z \ v \ y \ u \ u ; v \ v \ y\<^sup>\\<^sub>u ; z \ y\<^sup>\\<^sub>v" proof assume 1: "z \ v \ y \ u \ u ; v \ v" have "y\<^sup>\\<^sub>u ; z \ y ; y\<^sup>\\<^sub>u ; z \ u ; z" by (metis capped_omega_unfold meet.add_least_upper_bound meet.add_right_upper_bound meet_commutative mult_left_isotone) also have "... \ y ; y\<^sup>\\<^sub>u ; z \ v" using 1 by (metis meet.add_left_isotone meet_commutative mult_right_isotone order_trans) finally show "y\<^sup>\\<^sub>u ; z \ y\<^sup>\\<^sub>v" using 1 by (metis capped_omega_induct_meet mult_associative) qed -- {* AACP Theorem 6.18 *} lemma capped_omega_omega: "y \ u \ u ; v \ v \ (y\<^sup>\\<^sub>u)\<^sup>\\<^sub>v \ y\<^sup>\\<^sub>v" by (metis capped_omega_below capped_omega_sub_vector capped_omega_unfold meet.add_left_upper_bound order_trans) end (* end *) end