author | haftmann |
Mon, 13 Sep 2021 14:18:24 +0000 | |
changeset 74309 | 42523fbf643b |
parent 67019 | 7a3724078363 |
permissions | -rw-r--r-- |
43158 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
63538
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
nipkow
parents:
54809
diff
changeset
|
3 |
subsection \<open>Soundness and Completeness\<close> |
43158 | 4 |
|
63538
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
nipkow
parents:
54809
diff
changeset
|
5 |
theory Hoare_Sound_Complete |
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
nipkow
parents:
54809
diff
changeset
|
6 |
imports Hoare |
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
nipkow
parents:
54809
diff
changeset
|
7 |
begin |
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
nipkow
parents:
54809
diff
changeset
|
8 |
|
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
nipkow
parents:
54809
diff
changeset
|
9 |
subsubsection "Soundness" |
43158 | 10 |
|
11 |
lemma hoare_sound: "\<turnstile> {P}c{Q} \<Longrightarrow> \<Turnstile> {P}c{Q}" |
|
45015 | 12 |
proof(induction rule: hoare.induct) |
43158 | 13 |
case (While P b c) |
67019
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63538
diff
changeset
|
14 |
have "(WHILE b DO c,s) \<Rightarrow> t \<Longrightarrow> P s \<Longrightarrow> P t \<and> \<not> bval b t" for s t |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63538
diff
changeset
|
15 |
proof(induction "WHILE b DO c" s t rule: big_step_induct) |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63538
diff
changeset
|
16 |
case WhileFalse thus ?case by blast |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63538
diff
changeset
|
17 |
next |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63538
diff
changeset
|
18 |
case WhileTrue thus ?case |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63538
diff
changeset
|
19 |
using While.IH unfolding hoare_valid_def by blast |
7a3724078363
Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
63538
diff
changeset
|
20 |
qed |
43158 | 21 |
thus ?case unfolding hoare_valid_def by blast |
22 |
qed (auto simp: hoare_valid_def) |
|
23 |
||
24 |
||
63538
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
nipkow
parents:
54809
diff
changeset
|
25 |
subsubsection "Weakest Precondition" |
43158 | 26 |
|
27 |
definition wp :: "com \<Rightarrow> assn \<Rightarrow> assn" where |
|
28 |
"wp c Q = (\<lambda>s. \<forall>t. (c,s) \<Rightarrow> t \<longrightarrow> Q t)" |
|
29 |
||
30 |
lemma wp_SKIP[simp]: "wp SKIP Q = Q" |
|
31 |
by (rule ext) (auto simp: wp_def) |
|
32 |
||
33 |
lemma wp_Ass[simp]: "wp (x::=a) Q = (\<lambda>s. Q(s[a/x]))" |
|
34 |
by (rule ext) (auto simp: wp_def) |
|
35 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52373
diff
changeset
|
36 |
lemma wp_Seq[simp]: "wp (c\<^sub>1;;c\<^sub>2) Q = wp c\<^sub>1 (wp c\<^sub>2 Q)" |
43158 | 37 |
by (rule ext) (auto simp: wp_def) |
38 |
||
39 |
lemma wp_If[simp]: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52373
diff
changeset
|
40 |
"wp (IF b THEN c\<^sub>1 ELSE c\<^sub>2) Q = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52373
diff
changeset
|
41 |
(\<lambda>s. if bval b s then wp c\<^sub>1 Q s else wp c\<^sub>2 Q s)" |
43158 | 42 |
by (rule ext) (auto simp: wp_def) |
43 |
||
44 |
lemma wp_While_If: |
|
45 |
"wp (WHILE b DO c) Q s = |
|
52046
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
47818
diff
changeset
|
46 |
wp (IF b THEN c;;WHILE b DO c ELSE SKIP) Q s" |
43158 | 47 |
unfolding wp_def by (metis unfold_while) |
48 |
||
49 |
lemma wp_While_True[simp]: "bval b s \<Longrightarrow> |
|
52046
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents:
47818
diff
changeset
|
50 |
wp (WHILE b DO c) Q s = wp (c;; WHILE b DO c) Q s" |
43158 | 51 |
by(simp add: wp_While_If) |
52 |
||
53 |
lemma wp_While_False[simp]: "\<not> bval b s \<Longrightarrow> wp (WHILE b DO c) Q s = Q s" |
|
54 |
by(simp add: wp_While_If) |
|
55 |
||
56 |
||
63538
d7b5e2a222c2
added new vcg based on existentially quantified while-rule
nipkow
parents:
54809
diff
changeset
|
57 |
subsubsection "Completeness" |
43158 | 58 |
|
59 |
lemma wp_is_pre: "\<turnstile> {wp c Q} c {Q}" |
|
45015 | 60 |
proof(induction c arbitrary: Q) |
52373 | 61 |
case If thus ?case by(auto intro: conseq) |
43158 | 62 |
next |
63 |
case (While b c) |
|
64 |
let ?w = "WHILE b DO c" |
|
52193 | 65 |
show "\<turnstile> {wp ?w Q} ?w {Q}" |
66 |
proof(rule While') |
|
43158 | 67 |
show "\<turnstile> {\<lambda>s. wp ?w Q s \<and> bval b s} c {wp ?w Q}" |
52193 | 68 |
proof(rule strengthen_pre[OF _ While.IH]) |
43158 | 69 |
show "\<forall>s. wp ?w Q s \<and> bval b s \<longrightarrow> wp c (wp ?w Q) s" by auto |
70 |
qed |
|
71 |
show "\<forall>s. wp ?w Q s \<and> \<not> bval b s \<longrightarrow> Q s" by auto |
|
72 |
qed |
|
73 |
qed auto |
|
74 |
||
52291 | 75 |
lemma hoare_complete: assumes "\<Turnstile> {P}c{Q}" shows "\<turnstile> {P}c{Q}" |
43158 | 76 |
proof(rule strengthen_pre) |
77 |
show "\<forall>s. P s \<longrightarrow> wp c Q s" using assms |
|
78 |
by (auto simp: hoare_valid_def wp_def) |
|
79 |
show "\<turnstile> {wp c Q} c {Q}" by(rule wp_is_pre) |
|
80 |
qed |
|
81 |
||
54809 | 82 |
corollary hoare_sound_complete: "\<turnstile> {P}c{Q} \<longleftrightarrow> \<Turnstile> {P}c{Q}" |
83 |
by (metis hoare_complete hoare_sound) |
|
84 |
||
43158 | 85 |
end |