theory turing_hoare+ −
imports turing_basic+ −
begin+ −
+ −
+ −
type_synonym assert = "tape \<Rightarrow> bool"+ −
+ −
definition + −
assert_imp :: "assert \<Rightarrow> assert \<Rightarrow> bool" ("_ \<mapsto> _" [0, 0] 100)+ −
where+ −
"P \<mapsto> Q \<equiv> \<forall>l r. P (l, r) \<longrightarrow> Q (l, r)"+ −
+ −
fun + −
holds_for :: "(tape \<Rightarrow> bool) \<Rightarrow> config \<Rightarrow> bool" ("_ holds'_for _" [100, 99] 100)+ −
where+ −
"P holds_for (s, l, r) = P (l, r)" + −
+ −
(* halting case *)+ −
definition+ −
Hoare :: "assert \<Rightarrow> tprog0 \<Rightarrow> assert \<Rightarrow> bool" ("({(1_)}/ (_)/ {(1_)})" [50, 49] 50)+ −
where+ −
"{P} p {Q} \<equiv> + −
(\<forall>tp. P tp \<longrightarrow> (\<exists>n. is_final (steps0 (1, tp) p n) \<and> Q holds_for (steps0 (1, tp) p n)))"+ −
+ −
(* not halting case *)+ −
definition+ −
Hoare_unhalt :: "assert \<Rightarrow> tprog0 \<Rightarrow> bool" ("({(1_)}/ (_)) \<up>" [50, 49] 50)+ −
where+ −
"{P} p \<up> \<equiv> (\<forall>tp. P tp \<longrightarrow> (\<forall> n . \<not> (is_final (steps0 (1, tp) p n))))"+ −
+ −
+ −
lemma HoareI:+ −
assumes "\<And>l r. P (l, r) \<Longrightarrow> \<exists>n. is_final (steps0 (1, (l, r)) p n) \<and> Q holds_for (steps0 (1, (l, r)) p n)"+ −
shows "{P} p {Q}"+ −
unfolding Hoare_def + −
using assms by auto+ −
+ −
lemma Hoare_unhaltI:+ −
assumes "\<And>l r n. P (l, r) \<Longrightarrow> \<not> is_final (steps0 (1, (l, r)) p n)"+ −
shows "{P} p \<up>"+ −
unfolding Hoare_unhalt_def + −
using assms by auto+ −
+ −
lemma is_final_holds[simp]:+ −
assumes "is_final c"+ −
shows "Q holds_for (steps c p n) = Q holds_for c"+ −
using assms + −
apply(induct n)+ −
apply(auto)+ −
apply(case_tac [!] c)+ −
apply(auto)+ −
done+ −
+ −
+ −
text {*+ −
{P1} A {Q1} {P2} B {Q2} Q1 \<mapsto> P2 A well-formed+ −
---------------------------------------------------+ −
{P1} A |+| B {Q2}+ −
*}+ −
+ −
+ −
lemma Hoare_plus_halt: + −
assumes A_halt : "{P1} A {Q1}"+ −
and B_halt : "{P2} B {Q2}"+ −
and a_imp: "Q1 \<mapsto> P2"+ −
and A_wf : "tm_wf (A, 0)"+ −
shows "{P1} A |+| B {Q2}"+ −
proof(rule HoareI)+ −
fix l r+ −
assume h: "P1 (l, r)"+ −
then obtain n1 l' r' + −
where "is_final (steps0 (1, l, r) A n1)" + −
and a1: "Q1 holds_for (steps0 (1, l, r) A n1)"+ −
and a2: "steps0 (1, l, r) A n1 = (0, l', r')"+ −
using A_halt unfolding Hoare_def+ −
by (metis is_final_eq surj_pair) + −
then obtain n2 + −
where "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')"+ −
using A_wf by (rule_tac tm_comp_pre_halt_same) + −
moreover+ −
from a1 a2 a_imp have "P2 (l', r')" by (simp add: assert_imp_def)+ −
then obtain n3 l'' r''+ −
where "is_final (steps0 (1, l', r') B n3)" + −
and b1: "Q2 holds_for (steps0 (1, l', r') B n3)"+ −
and b2: "steps0 (1, l', r') B n3 = (0, l'', r'')"+ −
using B_halt unfolding Hoare_def + −
by (metis is_final_eq surj_pair) + −
then have "steps0 (Suc (length A div 2), l', r') (A |+| B) n3 = (0, l'', r'')"+ −
using A_wf by (rule_tac t_merge_second_halt_same) + −
ultimately show + −
"\<exists>n. is_final (steps0 (1, l, r) (A |+| B) n) \<and> Q2 holds_for (steps0 (1, l, r) (A |+| B) n)"+ −
using b1 b2 by (rule_tac x = "n2 + n3" in exI) (simp)+ −
qed+ −
+ −
lemma Hoare_plus_halt_simple [case_names A_halt B_halt A_wf]: + −
assumes A_halt : "{P1} A {P2}"+ −
and B_halt : "{P2} B {P3}"+ −
and A_wf : "tm_wf (A, 0)"+ −
shows "{P1} A |+| B {P3}"+ −
by (rule Hoare_plus_halt[OF A_halt B_halt _ A_wf])+ −
(simp add: assert_imp_def)+ −
+ −
+ −
+ −
text {*+ −
{P1} A {Q1} {P2} B loops Q1 \<mapsto> P2 A well-formed+ −
------------------------------------------------------+ −
{P1} A |+| B loops+ −
*}+ −
+ −
lemma Hoare_plus_unhalt:+ −
assumes A_halt: "{P1} A {Q1}"+ −
and B_uhalt: "{P2} B \<up>"+ −
and a_imp: "Q1 \<mapsto> P2"+ −
and A_wf : "tm_wf (A, 0)"+ −
shows "{P1} (A |+| B) \<up>"+ −
proof(rule_tac Hoare_unhaltI)+ −
fix n l r + −
assume h: "P1 (l, r)"+ −
then obtain n1 l' r'+ −
where a: "is_final (steps0 (1, l, r) A n1)" + −
and b: "Q1 holds_for (steps0 (1, l, r) A n1)"+ −
and c: "steps0 (1, l, r) A n1 = (0, l', r')"+ −
using A_halt unfolding Hoare_def + −
by (metis is_final_eq surj_pair) + −
then obtain n2 where eq: "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')"+ −
using A_wf by (rule_tac tm_comp_pre_halt_same)+ −
then show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)"+ −
proof(cases "n2 \<le> n")+ −
case True+ −
from b c a_imp have "P2 (l', r')" by (simp add: assert_imp_def)+ −
then have "\<forall> n. \<not> is_final (steps0 (1, l', r') B n) "+ −
using B_uhalt unfolding Hoare_unhalt_def by simp+ −
then have "\<not> is_final (steps0 (Suc 0, l', r') B (n - n2))" by auto+ −
then obtain s'' l'' r'' + −
where "steps0 (Suc 0, l', r') B (n - n2) = (s'', l'', r'')" + −
and "\<not> is_final (s'', l'', r'')" by (metis surj_pair)+ −
then have "steps0 (Suc (length A div 2), l', r') (A |+| B) (n - n2) = (s''+ length A div 2, l'', r'')"+ −
using A_wf by (auto dest: t_merge_second_same simp del: tm_wf.simps)+ −
then have "\<not> is_final (steps0 (1, l, r) (A |+| B) (n2 + (n - n2)))"+ −
using A_wf by (simp only: steps_add eq) (simp add: tm_wf.simps)+ −
then show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)" + −
using `n2 \<le> n` by simp+ −
next + −
case False+ −
then obtain n3 where "n = n2 - n3"+ −
by (metis diff_le_self le_imp_diff_is_add nat_add_commute nat_le_linear)+ −
moreover+ −
with eq show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)"+ −
by (simp add: not_is_final[where ?n1.0="n2"])+ −
qed+ −
qed+ −
+ −
lemma Hoare_plus_unhalt_simple [case_names A_halt B_unhalt A_wf]: + −
assumes A_halt: "{P1} A {P2}"+ −
and B_uhalt: "{P2} B \<up>"+ −
and A_wf : "tm_wf (A, 0)"+ −
shows "{P1} (A |+| B) \<up>"+ −
by (rule Hoare_plus_unhalt[OF A_halt B_uhalt _ A_wf])+ −
(simp add: assert_imp_def)+ −
+ −
+ −
lemma Hoare_weak:+ −
assumes a: "{P} p {Q}"+ −
and b: "P' \<mapsto> P" + −
and c: "Q \<mapsto> Q'"+ −
shows "{P'} p {Q'}"+ −
using assms+ −
unfolding Hoare_def assert_imp_def+ −
by (metis holds_for.simps surj_pair)+ −
+ −
+ −
+ −
end+ −