| author | wenzelm | 
| Thu, 10 Oct 2024 12:19:50 +0200 | |
| changeset 81144 | 6e6766cddf73 | 
| parent 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 50421 | 1  | 
(* Author: Tobias Nipkow *)  | 
| 43158 | 2  | 
|
| 
63538
 
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
 
nipkow 
parents: 
55132 
diff
changeset
 | 
3  | 
subsection "Hoare Logic for Total Correctness"  | 
| 43158 | 4  | 
|
| 
68776
 
403dd13cf6e9
avoid session qualification because no tex is generated when used;
 
nipkow 
parents: 
67406 
diff
changeset
 | 
5  | 
subsubsection "Separate Termination Relation"  | 
| 
 
403dd13cf6e9
avoid session qualification because no tex is generated when used;
 
nipkow 
parents: 
67406 
diff
changeset
 | 
6  | 
|
| 
63538
 
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
 
nipkow 
parents: 
55132 
diff
changeset
 | 
7  | 
theory Hoare_Total  | 
| 
 
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
 
nipkow 
parents: 
55132 
diff
changeset
 | 
8  | 
imports Hoare_Examples  | 
| 
 
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
 
nipkow 
parents: 
55132 
diff
changeset
 | 
9  | 
begin  | 
| 
 
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
 
nipkow 
parents: 
55132 
diff
changeset
 | 
10  | 
|
| 69505 | 11  | 
text\<open>Note that this definition of total validity \<open>\<Turnstile>\<^sub>t\<close> only  | 
| 67406 | 12  | 
works if execution is deterministic (which it is in our case).\<close>  | 
| 43158 | 13  | 
|
14  | 
definition hoare_tvalid :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool"  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
74371 
diff
changeset
 | 
15  | 
  (\<open>\<Turnstile>\<^sub>t {(1_)}/ (_)/ {(1_)}\<close> 50) where
 | 
| 52281 | 16  | 
"\<Turnstile>\<^sub>t {P}c{Q}  \<longleftrightarrow>  (\<forall>s. P s \<longrightarrow> (\<exists>t. (c,s) \<Rightarrow> t \<and> Q t))"
 | 
| 43158 | 17  | 
|
| 67406 | 18  | 
text\<open>Provability of Hoare triples in the proof system for total  | 
| 69505 | 19  | 
correctness is written \<open>\<turnstile>\<^sub>t {P}c{Q}\<close> and defined
 | 
20  | 
inductively. The rules for \<open>\<turnstile>\<^sub>t\<close> differ from those for  | 
|
21  | 
\<open>\<turnstile>\<close> only in the one place where nontermination can arise: the  | 
|
| 69597 | 22  | 
\<^term>\<open>While\<close>-rule.\<close>  | 
| 43158 | 23  | 
|
24  | 
inductive  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
74371 
diff
changeset
 | 
25  | 
  hoaret :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" (\<open>\<turnstile>\<^sub>t ({(1_)}/ (_)/ {(1_)})\<close> 50)
 | 
| 43158 | 26  | 
where  | 
| 52281 | 27  | 
|
28  | 
Skip:  "\<turnstile>\<^sub>t {P} SKIP {P}"  |
 | 
|
29  | 
||
30  | 
Assign:  "\<turnstile>\<^sub>t {\<lambda>s. P(s[a/x])} x::=a {P}"  |
 | 
|
31  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52529 
diff
changeset
 | 
32  | 
Seq: "\<lbrakk> \<turnstile>\<^sub>t {P\<^sub>1} c\<^sub>1 {P\<^sub>2}; \<turnstile>\<^sub>t {P\<^sub>2} c\<^sub>2 {P\<^sub>3} \<rbrakk> \<Longrightarrow> \<turnstile>\<^sub>t {P\<^sub>1} c\<^sub>1;;c\<^sub>2 {P\<^sub>3}"  |
 | 
| 52281 | 33  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52529 
diff
changeset
 | 
34  | 
If: "\<lbrakk> \<turnstile>\<^sub>t {\<lambda>s. P s \<and> bval b s} c\<^sub>1 {Q}; \<turnstile>\<^sub>t {\<lambda>s. P s \<and> \<not> bval b s} c\<^sub>2 {Q} \<rbrakk>
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52529 
diff
changeset
 | 
35  | 
  \<Longrightarrow> \<turnstile>\<^sub>t {P} IF b THEN c\<^sub>1 ELSE c\<^sub>2 {Q}"  |
 | 
| 52281 | 36  | 
|
| 43158 | 37  | 
While:  | 
| 52281 | 38  | 
"(\<And>n::nat.  | 
| 52333 | 39  | 
    \<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')})
 | 
| 52281 | 40  | 
   \<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}"  |
 | 
41  | 
||
| 43158 | 42  | 
conseq: "\<lbrakk> \<forall>s. P' s \<longrightarrow> P s; \<turnstile>\<^sub>t {P}c{Q}; \<forall>s. Q s \<longrightarrow> Q' s  \<rbrakk> \<Longrightarrow>
 | 
43  | 
           \<turnstile>\<^sub>t {P'}c{Q'}"
 | 
|
44  | 
||
| 69597 | 45  | 
text\<open>The \<^term>\<open>While\<close>-rule is like the one for partial correctness but it  | 
| 43158 | 46  | 
requires additionally that with every execution of the loop body some measure  | 
| 52281 | 47  | 
relation @{term[source]"T :: state \<Rightarrow> nat \<Rightarrow> bool"} decreases.
 | 
| 67406 | 48  | 
The following functional version is more intuitive:\<close>  | 
| 52281 | 49  | 
|
50  | 
lemma While_fun:  | 
|
51  | 
  "\<lbrakk> \<And>n::nat. \<turnstile>\<^sub>t {\<lambda>s. P s \<and> bval b s \<and> n = f s} c {\<lambda>s. P s \<and> f s < n}\<rbrakk>
 | 
|
52  | 
   \<Longrightarrow> \<turnstile>\<^sub>t {P} WHILE b DO c {\<lambda>s. P s \<and> \<not>bval b s}"
 | 
|
53  | 
by (rule While [where T="\<lambda>s n. n = f s", simplified])  | 
|
54  | 
||
| 67406 | 55  | 
text\<open>Building in the consequence rule:\<close>  | 
| 43158 | 56  | 
|
57  | 
lemma strengthen_pre:  | 
|
58  | 
  "\<lbrakk> \<forall>s. P' s \<longrightarrow> P s;  \<turnstile>\<^sub>t {P} c {Q} \<rbrakk> \<Longrightarrow> \<turnstile>\<^sub>t {P'} c {Q}"
 | 
|
59  | 
by (metis conseq)  | 
|
60  | 
||
61  | 
lemma weaken_post:  | 
|
62  | 
  "\<lbrakk> \<turnstile>\<^sub>t {P} c {Q};  \<forall>s. Q s \<longrightarrow> Q' s \<rbrakk> \<Longrightarrow>  \<turnstile>\<^sub>t {P} c {Q'}"
 | 
|
63  | 
by (metis conseq)  | 
|
64  | 
||
65  | 
lemma Assign': "\<forall>s. P s \<longrightarrow> Q(s[a/x]) \<Longrightarrow> \<turnstile>\<^sub>t {P} x ::= a {Q}"
 | 
|
66  | 
by (simp add: strengthen_pre[OF _ Assign])  | 
|
67  | 
||
| 52281 | 68  | 
lemma While_fun':  | 
69  | 
assumes "\<And>n::nat. \<turnstile>\<^sub>t {\<lambda>s. P s \<and> bval b s \<and> n = f s} c {\<lambda>s. P s \<and> f s < n}"
 | 
|
| 43158 | 70  | 
and "\<forall>s. P s \<and> \<not> bval b s \<longrightarrow> Q s"  | 
| 52281 | 71  | 
shows "\<turnstile>\<^sub>t {P} WHILE b DO c {Q}"
 | 
72  | 
by(blast intro: assms(1) weaken_post[OF While_fun assms(2)])  | 
|
| 43158 | 73  | 
|
| 52227 | 74  | 
|
| 67406 | 75  | 
text\<open>Our standard example:\<close>  | 
| 43158 | 76  | 
|
| 52228 | 77  | 
lemma "\<turnstile>\<^sub>t {\<lambda>s. s ''x'' = i} ''y'' ::= N 0;; wsum {\<lambda>s. s ''y'' = sum i}"
 | 
| 47818 | 78  | 
apply(rule Seq)  | 
| 52228 | 79  | 
prefer 2  | 
| 52281 | 80  | 
apply(rule While_fun' [where P = "\<lambda>s. (s ''y'' = sum i - sum(s ''x''))"  | 
81  | 
and f = "\<lambda>s. nat(s ''x'')"])  | 
|
| 52228 | 82  | 
apply(rule Seq)  | 
83  | 
prefer 2  | 
|
84  | 
apply(rule Assign)  | 
|
85  | 
apply(rule Assign')  | 
|
86  | 
apply simp  | 
|
87  | 
apply(simp)  | 
|
| 43158 | 88  | 
apply(rule Assign')  | 
89  | 
apply simp  | 
|
90  | 
done  | 
|
91  | 
||
| 74371 | 92  | 
text \<open>Nested loops. This poses a problem for VCGs because the proof of the inner loop needs to  | 
93  | 
refer to outer loops. This works here because the invariant is not written down statically but  | 
|
94  | 
created in the context of a proof that has already introduced/fixed outer \<open>n\<close>s that can be  | 
|
95  | 
referred to.\<close>  | 
|
96  | 
||
97  | 
lemma  | 
|
98  | 
 "\<turnstile>\<^sub>t {\<lambda>_. True}
 | 
|
99  | 
WHILE Less (N 0) (V ''x'')  | 
|
100  | 
  DO (''x'' ::= Plus (V ''x'') (N(-1));;
 | 
|
101  | 
''y'' ::= V ''x'';;  | 
|
102  | 
WHILE Less (N 0) (V ''y'') DO ''y'' ::= Plus (V ''y'') (N(-1)))  | 
|
103  | 
  {\<lambda>_. True}"
 | 
|
104  | 
apply(rule While_fun'[where f = "\<lambda>s. nat(s ''x'')"])  | 
|
105  | 
prefer 2 apply simp  | 
|
106  | 
apply(rule_tac P\<^sub>2 = "\<lambda>s. nat(s ''x'') < n" in Seq)  | 
|
107  | 
apply(rule_tac P\<^sub>2 = "\<lambda>s. nat(s ''x'') < n" in Seq)  | 
|
108  | 
apply(rule Assign')  | 
|
109  | 
apply simp  | 
|
110  | 
apply(rule Assign')  | 
|
111  | 
apply simp  | 
|
112  | 
(* note that the invariant refers to the outer \<open>n\<close>: *)  | 
|
113  | 
apply(rule While_fun'[where f = "\<lambda>s. nat(s ''y'')"])  | 
|
114  | 
prefer 2 apply simp  | 
|
115  | 
apply(rule Assign')  | 
|
116  | 
apply simp  | 
|
117  | 
done  | 
|
| 43158 | 118  | 
|
| 67406 | 119  | 
text\<open>The soundness theorem:\<close>  | 
| 43158 | 120  | 
|
121  | 
theorem hoaret_sound: "\<turnstile>\<^sub>t {P}c{Q}  \<Longrightarrow>  \<Turnstile>\<^sub>t {P}c{Q}"
 | 
|
| 52282 | 122  | 
proof(unfold hoare_tvalid_def, induction rule: hoaret.induct)  | 
| 52227 | 123  | 
case (While P b T c)  | 
| 
67019
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
124  | 
have "\<lbrakk> P s; T s n \<rbrakk> \<Longrightarrow> \<exists>t. (WHILE b DO c, s) \<Rightarrow> t \<and> P t \<and> \<not> bval b t" for s n  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
125  | 
proof(induction "n" arbitrary: s rule: less_induct)  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
126  | 
case (less n) thus ?case by (metis While.IH WhileFalse WhileTrue)  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
127  | 
qed  | 
| 52227 | 128  | 
thus ?case by auto  | 
| 43158 | 129  | 
next  | 
130  | 
case If thus ?case by auto blast  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44177 
diff
changeset
 | 
131  | 
qed fastforce+  | 
| 43158 | 132  | 
|
133  | 
||
| 67406 | 134  | 
text\<open>  | 
| 43158 | 135  | 
The completeness proof proceeds along the same lines as the one for partial  | 
136  | 
correctness. First we have to strengthen our notion of weakest precondition  | 
|
| 67406 | 137  | 
to take termination into account:\<close>  | 
| 43158 | 138  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
74371 
diff
changeset
 | 
139  | 
definition wpt :: "com \<Rightarrow> assn \<Rightarrow> assn" (\<open>wp\<^sub>t\<close>) where  | 
| 52290 | 140  | 
"wp\<^sub>t c Q = (\<lambda>s. \<exists>t. (c,s) \<Rightarrow> t \<and> Q t)"  | 
| 43158 | 141  | 
|
142  | 
lemma [simp]: "wp\<^sub>t SKIP Q = Q"  | 
|
143  | 
by(auto intro!: ext simp: wpt_def)  | 
|
144  | 
||
145  | 
lemma [simp]: "wp\<^sub>t (x ::= e) Q = (\<lambda>s. Q(s(x := aval e s)))"  | 
|
146  | 
by(auto intro!: ext simp: wpt_def)  | 
|
147  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52529 
diff
changeset
 | 
148  | 
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)"  | 
| 43158 | 149  | 
unfolding wpt_def  | 
150  | 
apply(rule ext)  | 
|
151  | 
apply auto  | 
|
152  | 
done  | 
|
153  | 
||
154  | 
lemma [simp]:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52529 
diff
changeset
 | 
155  | 
"wp\<^sub>t (IF b THEN c\<^sub>1 ELSE c\<^sub>2) Q = (\<lambda>s. wp\<^sub>t (if bval b s then c\<^sub>1 else c\<^sub>2) Q s)"  | 
| 43158 | 156  | 
apply(unfold wpt_def)  | 
157  | 
apply(rule ext)  | 
|
158  | 
apply auto  | 
|
159  | 
done  | 
|
160  | 
||
161  | 
||
| 69597 | 162  | 
text\<open>Now we define the number of iterations \<^term>\<open>WHILE b DO c\<close> needs to  | 
| 69505 | 163  | 
terminate when started in state \<open>s\<close>. Because this is a truly partial  | 
| 67406 | 164  | 
function, we define it as an (inductive) relation first:\<close>  | 
| 43158 | 165  | 
|
166  | 
inductive Its :: "bexp \<Rightarrow> com \<Rightarrow> state \<Rightarrow> nat \<Rightarrow> bool" where  | 
|
167  | 
Its_0: "\<not> bval b s \<Longrightarrow> Its b c s 0" |  | 
|
168  | 
Its_Suc: "\<lbrakk> bval b s; (c,s) \<Rightarrow> s'; Its b c s' n \<rbrakk> \<Longrightarrow> Its b c s (Suc n)"  | 
|
169  | 
||
| 67406 | 170  | 
text\<open>The relation is in fact a function:\<close>  | 
| 43158 | 171  | 
|
172  | 
lemma Its_fun: "Its b c s n \<Longrightarrow> Its b c s n' \<Longrightarrow> n=n'"  | 
|
| 45015 | 173  | 
proof(induction arbitrary: n' rule:Its.induct)  | 
| 43158 | 174  | 
case Its_0 thus ?case by(metis Its.cases)  | 
175  | 
next  | 
|
176  | 
case Its_Suc thus ?case by(metis Its.cases big_step_determ)  | 
|
177  | 
qed  | 
|
178  | 
||
| 69597 | 179  | 
text\<open>For all terminating loops, \<^const>\<open>Its\<close> yields a result:\<close>  | 
| 43158 | 180  | 
|
181  | 
lemma WHILE_Its: "(WHILE b DO c,s) \<Rightarrow> t \<Longrightarrow> \<exists>n. Its b c s n"  | 
|
| 45015 | 182  | 
proof(induction "WHILE b DO c" s t rule: big_step_induct)  | 
| 43158 | 183  | 
case WhileFalse thus ?case by (metis Its_0)  | 
184  | 
next  | 
|
185  | 
case WhileTrue thus ?case by (metis Its_Suc)  | 
|
186  | 
qed  | 
|
187  | 
||
188  | 
lemma wpt_is_pre: "\<turnstile>\<^sub>t {wp\<^sub>t c Q} c {Q}"
 | 
|
| 45015 | 189  | 
proof (induction c arbitrary: Q)  | 
| 52373 | 190  | 
case SKIP show ?case by (auto intro:hoaret.Skip)  | 
| 43158 | 191  | 
next  | 
| 52373 | 192  | 
case Assign show ?case by (auto intro:hoaret.Assign)  | 
| 43158 | 193  | 
next  | 
| 52373 | 194  | 
case Seq thus ?case by (auto intro:hoaret.Seq)  | 
| 43158 | 195  | 
next  | 
| 52373 | 196  | 
case If thus ?case by (auto intro:hoaret.If hoaret.conseq)  | 
| 43158 | 197  | 
next  | 
198  | 
case (While b c)  | 
|
199  | 
let ?w = "WHILE b DO c"  | 
|
| 52228 | 200  | 
let ?T = "Its b c"  | 
| 
67019
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
201  | 
have 1: "\<forall>s. wp\<^sub>t ?w Q s \<longrightarrow> wp\<^sub>t ?w Q s \<and> (\<exists>n. Its b c s n)"  | 
| 52227 | 202  | 
unfolding wpt_def by (metis WHILE_Its)  | 
| 
67019
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
203  | 
let ?R = "\<lambda>n s'. wp\<^sub>t ?w Q s' \<and> (\<exists>n'<n. ?T s' n')"  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
204  | 
have "\<forall>s. wp\<^sub>t ?w Q s \<and> bval b s \<and> ?T s n \<longrightarrow> wp\<^sub>t c (?R n) s" for n  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
205  | 
proof -  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
206  | 
have "wp\<^sub>t c (?R n) s" if "bval b s" and "?T s n" and "(?w, s) \<Rightarrow> t" and "Q t" for s t  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
207  | 
proof -  | 
| 67406 | 208  | 
from \<open>bval b s\<close> and \<open>(?w, s) \<Rightarrow> t\<close> obtain s' where  | 
| 52290 | 209  | 
"(c,s) \<Rightarrow> s'" "(?w,s') \<Rightarrow> t" by auto  | 
| 67406 | 210  | 
from \<open>(?w, s') \<Rightarrow> t\<close> obtain n' where "?T s' n'"  | 
| 52290 | 211  | 
by (blast dest: WHILE_Its)  | 
| 67406 | 212  | 
with \<open>bval b s\<close> and \<open>(c, s) \<Rightarrow> s'\<close> have "?T s (Suc n')" by (rule Its_Suc)  | 
213  | 
with \<open>?T s n\<close> have "n = Suc n'" by (rule Its_fun)  | 
|
214  | 
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>  | 
|
| 
67019
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
215  | 
show ?thesis by (auto simp: wpt_def)  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
216  | 
qed  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
217  | 
thus ?thesis  | 
| 52227 | 218  | 
unfolding wpt_def by auto  | 
219  | 
(* by (metis WhileE Its_Suc Its_fun WHILE_Its lessI) *)  | 
|
| 
67019
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
220  | 
qed  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
221  | 
note 2 = hoaret.While[OF strengthen_pre[OF this While.IH]]  | 
| 
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
222  | 
have "\<forall>s. wp\<^sub>t ?w Q s \<and> \<not> bval b s \<longrightarrow> Q s"  | 
| 52290 | 223  | 
by (auto simp add:wpt_def)  | 
| 
67019
 
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
 
nipkow 
parents: 
63538 
diff
changeset
 | 
224  | 
with 1 2 show ?case by (rule conseq)  | 
| 43158 | 225  | 
qed  | 
226  | 
||
227  | 
||
| 69597 | 228  | 
text\<open>\noindent In the \<^term>\<open>While\<close>-case, \<^const>\<open>Its\<close> provides the obvious  | 
| 43158 | 229  | 
termination argument.  | 
230  | 
||
231  | 
The actual completeness theorem follows directly, in the same manner  | 
|
| 67406 | 232  | 
as for partial correctness:\<close>  | 
| 43158 | 233  | 
|
234  | 
theorem hoaret_complete: "\<Turnstile>\<^sub>t {P}c{Q} \<Longrightarrow> \<turnstile>\<^sub>t {P}c{Q}"
 | 
|
235  | 
apply(rule strengthen_pre[OF _ wpt_is_pre])  | 
|
| 52290 | 236  | 
apply(auto simp: hoare_tvalid_def wpt_def)  | 
| 43158 | 237  | 
done  | 
238  | 
||
| 55132 | 239  | 
corollary hoaret_sound_complete: "\<turnstile>\<^sub>t {P}c{Q} \<longleftrightarrow> \<Turnstile>\<^sub>t {P}c{Q}"
 | 
240  | 
by (metis hoaret_sound hoaret_complete)  | 
|
241  | 
||
| 43158 | 242  | 
end  |