(* Author: Tobias Nipkow *)
subsection "Hoare Logic for Total Correctness"
subsubsection "Separate Termination Relation"
theory Hoare_Total
imports Hoare_Examples
begin
text\<open>Note that this definition of total validity \<open>\<Turnstile>\<^sub>t\<close> only
works if execution is deterministic (which it is in our case).\<close>
definition hoare_tvalid :: "assn \ com \ assn \ bool"
("\\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where
"\\<^sub>t {P}c{Q} \ (\s. P s \ (\t. (c,s) \ t \ Q t))"
text\<open>Provability of Hoare triples in the proof system for total
correctness is written \<open>\<turnstile>\<^sub>t {P}c{Q}\<close> and defined
inductively. The rules for \<open>\<turnstile>\<^sub>t\<close> differ from those for
\<open>\<turnstile>\<close> only in the one place where nontermination can arise: the
\<^term>\<open>While\<close>-rule.\<close>
inductive
hoaret :: "assn \ com \ assn \ bool" ("\\<^sub>t ({(1_)}/ (_)/ {(1_)})" 50)
where
Skip: "\\<^sub>t {P} SKIP {P}" |
Assign: "\\<^sub>t {\s. P(s[a/x])} x::=a {P}" |
Seq: "\ \\<^sub>t {P\<^sub>1} c\<^sub>1 {P\<^sub>2}; \\<^sub>t {P\<^sub>2} c\<^sub>2 {P\<^sub>3} \ \ \\<^sub>t {P\<^sub>1} c\<^sub>1;;c\<^sub>2 {P\<^sub>3}" |
If: "\ \\<^sub>t {\s. P s \ bval b s} c\<^sub>1 {Q}; \\<^sub>t {\s. P s \ \ bval b s} c\<^sub>2 {Q} \
\<Longrightarrow> \<turnstile>\<^sub>t {P} IF b THEN c\<^sub>1 ELSE c\<^sub>2 {Q}" |
While:
"(\n::nat.
\<turnstile>\<^sub>t {\<lambda>s. P s \<and> bval b s \<and> T s n} c {\<lambda>s. P s \<and> (\<exists>n'<n. T s n')})
\<Longrightarrow> \<turnstile>\<^sub>t {\<lambda>s. P s \<and> (\<exists>n. T s n)} WHILE b DO c {\<lambda>s. P s \<and> \<not>bval b s}" |
conseq: "\ \s. P' s \ P s; \\<^sub>t {P}c{Q}; \s. Q s \ Q' s \ \
\<turnstile>\<^sub>t {P'}c{Q'}"
text\<open>The \<^term>\<open>While\<close>-rule is like the one for partial correctness but it
requires additionally that with every execution of the loop body some measure
relation @{term[source]"T :: state \ nat \ bool"} decreases.
The following functional version is more intuitive:\<close>
lemma While_fun:
"\ \n::nat. \\<^sub>t {\s. P s \ bval b s \ n = f s} c {\s. P s \ f s < n}\
\<Longrightarrow> \<turnstile>\<^sub>t {P} WHILE b DO c {\<lambda>s. P s \<and> \<not>bval b s}"
by (rule While [where T="\s n. n = f s", simplified])
text\<open>Building in the consequence rule:\<close>
lemma strengthen_pre:
"\ \s. P' s \ P s; \\<^sub>t {P} c {Q} \ \ \\<^sub>t {P'} c {Q}"
by (metis conseq)
lemma weaken_post:
"\ \\<^sub>t {P} c {Q}; \s. Q s \ Q' s \ \ \\<^sub>t {P} c {Q'}"
by (metis conseq)
lemma Assign': "\s. P s \ Q(s[a/x]) \ \\<^sub>t {P} x ::= a {Q}"
by (simp add: strengthen_pre[OF _ Assign])
lemma While_fun':
assumes "\n::nat. \\<^sub>t {\s. P s \ bval b s \ n = f s} c {\s. P s \ f s < n}"
and "\s. P s \ \ bval b s \ Q s"
shows "\\<^sub>t {P} WHILE b DO c {Q}"
by(blast intro: assms(1) weaken_post[OF While_fun assms(2)])
text\<open>Our standard example:\<close>
lemma "\\<^sub>t {\s. s ''x'' = i} ''y'' ::= N 0;; wsum {\s. s ''y'' = sum i}"
apply(rule Seq)
prefer 2
apply(rule While_fun' [where P = "\s. (s ''y'' = sum i - sum(s ''x''))"
and f = "\s. nat(s ''x'')"])
apply(rule Seq)
prefer 2
apply(rule Assign)
apply(rule Assign')
apply simp
apply(simp)
apply(rule Assign')
apply simp
done
text\<open>The soundness theorem:\<close>
theorem hoaret_sound: "\\<^sub>t {P}c{Q} \ \\<^sub>t {P}c{Q}"
proof(unfold hoare_tvalid_def, induction rule: hoaret.induct)
case (While P b T c)
have "\ P s; T s n \ \ \t. (WHILE b DO c, s) \ t \ P t \ \ bval b t" for s n
proof(induction "n" arbitrary: s rule: less_induct)
case (less n) thus ?case by (metis While.IH WhileFalse WhileTrue)
qed
thus ?case by auto
next
case If thus ?case by auto blast
qed fastforce+
text\<open>
The completeness proof proceeds along the same lines as the one for partial
correctness. First we have to strengthen our notion of weakest precondition
to take termination into account:\<close>
definition wpt :: "com \ assn \ assn" ("wp\<^sub>t") where
"wp\<^sub>t c Q = (\s. \t. (c,s) \ t \ Q t)"
lemma [simp]: "wp\<^sub>t SKIP Q = Q"
by(auto intro!: ext simp: wpt_def)
lemma [simp]: "wp\<^sub>t (x ::= e) Q = (\s. Q(s(x := aval e s)))"
by(auto intro!: ext simp: wpt_def)
lemma [simp]: "wp\<^sub>t (c\<^sub>1;;c\<^sub>2) Q = wp\<^sub>t c\<^sub>1 (wp\<^sub>t c\<^sub>2 Q)"
unfolding wpt_def
apply(rule ext)
apply auto
done
lemma [simp]:
"wp\<^sub>t (IF b THEN c\<^sub>1 ELSE c\<^sub>2) Q = (\s. wp\<^sub>t (if bval b s then c\<^sub>1 else c\<^sub>2) Q s)"
apply(unfold wpt_def)
apply(rule ext)
apply auto
done
text\<open>Now we define the number of iterations \<^term>\<open>WHILE b DO c\<close> needs to
terminate when started in state \<open>s\<close>. Because this is a truly partial
function, we define it as an (inductive) relation first:\<close>
inductive Its :: "bexp \ com \ state \ nat \ bool" where
Its_0: "\ bval b s \ Its b c s 0" |
Its_Suc: "\ bval b s; (c,s) \ s'; Its b c s' n \ \ Its b c s (Suc n)"
text\<open>The relation is in fact a function:\<close>
lemma Its_fun: "Its b c s n \ Its b c s n' \ n=n'"
proof(induction arbitrary: n' rule:Its.induct)
case Its_0 thus ?case by(metis Its.cases)
next
case Its_Suc thus ?case by(metis Its.cases big_step_determ)
qed
text\<open>For all terminating loops, \<^const>\<open>Its\<close> yields a result:\<close>
lemma WHILE_Its: "(WHILE b DO c,s) \ t \ \n. Its b c s n"
proof(induction "WHILE b DO c" s t rule: big_step_induct)
case WhileFalse thus ?case by (metis Its_0)
next
case WhileTrue thus ?case by (metis Its_Suc)
qed
lemma wpt_is_pre: "\\<^sub>t {wp\<^sub>t c Q} c {Q}"
proof (induction c arbitrary: Q)
case SKIP show ?case by (auto intro:hoaret.Skip)
next
case Assign show ?case by (auto intro:hoaret.Assign)
next
case Seq thus ?case by (auto intro:hoaret.Seq)
next
case If thus ?case by (auto intro:hoaret.If hoaret.conseq)
next
case (While b c)
let ?w = "WHILE b DO c"
let ?T = "Its b c"
have 1: "\s. wp\<^sub>t ?w Q s \ wp\<^sub>t ?w Q s \ (\n. Its b c s n)"
unfolding wpt_def by (metis WHILE_Its)
let ?R = "\n s'. wp\<^sub>t ?w Q s' \ (\n'
have "\s. wp\<^sub>t ?w Q s \ bval b s \ ?T s n \ wp\<^sub>t c (?R n) s" for n
proof -
have "wp\<^sub>t c (?R n) s" if "bval b s" and "?T s n" and "(?w, s) \ t" and "Q t" for s t
proof -
from \<open>bval b s\<close> and \<open>(?w, s) \<Rightarrow> t\<close> obtain s' where
"(c,s) \ s'" "(?w,s') \ t" by auto
from \<open>(?w, s') \<Rightarrow> t\<close> obtain n' where "?T s' n'"
by (blast dest: WHILE_Its)
with \<open>bval b s\<close> and \<open>(c, s) \<Rightarrow> s'\<close> have "?T s (Suc n')" by (rule Its_Suc)
with \<open>?T s n\<close> have "n = Suc n'" by (rule Its_fun)
with \<open>(c,s) \<Rightarrow> s'\<close> and \<open>(?w,s') \<Rightarrow> t\<close> and \<open>Q t\<close> and \<open>?T s' n'\<close>
show ?thesis by (auto simp: wpt_def)
qed
thus ?thesis
unfolding wpt_def by auto
(* by (metis WhileE Its_Suc Its_fun WHILE_Its lessI) *)
qed
note 2 = hoaret.While[OF strengthen_pre[OF this While.IH]]
have "\s. wp\<^sub>t ?w Q s \ \ bval b s \ Q s"
by (auto simp add:wpt_def)
with 1 2 show ?case by (rule conseq)
qed
text\<open>\noindent In the \<^term>\<open>While\<close>-case, \<^const>\<open>Its\<close> provides the obvious
termination argument.
The actual completeness theorem follows directly, in the same manner
as for partial correctness:\<close>
theorem hoaret_complete: "\\<^sub>t {P}c{Q} \ \\<^sub>t {P}c{Q}"
apply(rule strengthen_pre[OF _ wpt_is_pre])
apply(auto simp: hoare_tvalid_def wpt_def)
done
corollary hoaret_sound_complete: "\\<^sub>t {P}c{Q} \ \\<^sub>t {P}c{Q}"
by (metis hoaret_sound hoaret_complete)
end
¤ Dauer der Verarbeitung: 0.16 Sekunden
(vorverarbeitet)
¤
|
Haftungshinweis
Die Informationen auf dieser Webseite wurden
nach bestem Wissen sorgfältig zusammengestellt. Es wird jedoch weder Vollständigkeit, noch Richtigkeit,
noch Qualität der bereit gestellten Informationen zugesichert.
Bemerkung:
Die farbliche Syntaxdarstellung ist noch experimentell.
|