| author | blanchet | 
| Tue, 09 Sep 2014 20:51:36 +0200 | |
| changeset 58255 | 9dfe8506c04d | 
| parent 58249 | 180f1b3508ed | 
| child 58310 | 91ea607a34d8 | 
| permissions | -rw-r--r-- | 
| 50421 | 1  | 
(* Author: Tobias Nipkow *)  | 
| 43158 | 2  | 
|
| 52269 | 3  | 
theory VCG imports Hoare begin  | 
| 43158 | 4  | 
|
| 50421 | 5  | 
subsection "Verification Conditions"  | 
| 43158 | 6  | 
|
7  | 
text{* Annotated commands: commands where loops are annotated with
 | 
|
8  | 
invariants. *}  | 
|
9  | 
||
| 
58249
 
180f1b3508ed
use 'datatype_new' (soon to be renamed 'datatype') in Isabelle's libraries
 
blanchet 
parents: 
55003 
diff
changeset
 | 
10  | 
datatype_new acom =  | 
| 52193 | 11  | 
  Askip                  ("SKIP") |
 | 
| 45840 | 12  | 
  Aassign vname aexp     ("(_ ::= _)" [1000, 61] 61) |
 | 
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
50421 
diff
changeset
 | 
13  | 
  Aseq   acom acom       ("_;;/ _"  [60, 61] 60) |
 | 
| 45840 | 14  | 
  Aif bexp acom acom     ("(IF _/ THEN _/ ELSE _)"  [0, 0, 61] 61) |
 | 
15  | 
  Awhile assn bexp acom  ("({_}/ WHILE _/ DO _)"  [0, 0, 61] 61)
 | 
|
| 43158 | 16  | 
|
| 54941 | 17  | 
notation com.SKIP ("SKIP")
 | 
| 
52267
 
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
 
nipkow 
parents: 
52193 
diff
changeset
 | 
18  | 
|
| 
 
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
 
nipkow 
parents: 
52193 
diff
changeset
 | 
19  | 
text{* Strip annotations: *}
 | 
| 
 
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
 
nipkow 
parents: 
52193 
diff
changeset
 | 
20  | 
|
| 
 
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
 
nipkow 
parents: 
52193 
diff
changeset
 | 
21  | 
fun strip :: "acom \<Rightarrow> com" where  | 
| 54941 | 22  | 
"strip SKIP = SKIP" |  | 
| 
52267
 
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
 
nipkow 
parents: 
52193 
diff
changeset
 | 
23  | 
"strip (x ::= a) = (x ::= a)" |  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52371 
diff
changeset
 | 
24  | 
"strip (C\<^sub>1;; C\<^sub>2) = (strip C\<^sub>1;; strip C\<^sub>2)" |  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52371 
diff
changeset
 | 
25  | 
"strip (IF b THEN C\<^sub>1 ELSE C\<^sub>2) = (IF b THEN strip C\<^sub>1 ELSE strip C\<^sub>2)" |  | 
| 52332 | 26  | 
"strip ({_} WHILE b DO C) = (WHILE b DO strip C)"
 | 
| 
52267
 
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
 
nipkow 
parents: 
52193 
diff
changeset
 | 
27  | 
|
| 43158 | 28  | 
text{* Weakest precondition from annotated commands: *}
 | 
29  | 
||
30  | 
fun pre :: "acom \<Rightarrow> assn \<Rightarrow> assn" where  | 
|
| 
52267
 
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
 
nipkow 
parents: 
52193 
diff
changeset
 | 
31  | 
"pre SKIP Q = Q" |  | 
| 
 
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
 
nipkow 
parents: 
52193 
diff
changeset
 | 
32  | 
"pre (x ::= a) Q = (\<lambda>s. Q(s(x := aval a s)))" |  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52371 
diff
changeset
 | 
33  | 
"pre (C\<^sub>1;; C\<^sub>2) Q = pre C\<^sub>1 (pre C\<^sub>2 Q)" |  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52371 
diff
changeset
 | 
34  | 
"pre (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: 
52371 
diff
changeset
 | 
35  | 
(\<lambda>s. if bval b s then pre C\<^sub>1 Q s else pre C\<^sub>2 Q s)" |  | 
| 52332 | 36  | 
"pre ({I} WHILE b DO C) Q = I"
 | 
| 43158 | 37  | 
|
38  | 
text{* Verification condition: *}
 | 
|
39  | 
||
| 55003 | 40  | 
fun vc :: "acom \<Rightarrow> assn \<Rightarrow> bool" where  | 
41  | 
"vc SKIP Q = True" |  | 
|
42  | 
"vc (x ::= a) Q = True" |  | 
|
43  | 
"vc (C\<^sub>1;; C\<^sub>2) Q = (vc C\<^sub>1 (pre C\<^sub>2 Q) \<and> vc C\<^sub>2 Q)" |  | 
|
44  | 
"vc (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q = (vc C\<^sub>1 Q \<and> vc C\<^sub>2 Q)" |  | 
|
| 52332 | 45  | 
"vc ({I} WHILE b DO C) Q =
 | 
| 55003 | 46  | 
((\<forall>s. (I s \<and> bval b s \<longrightarrow> pre C I s) \<and>  | 
47  | 
(I s \<and> \<not> bval b s \<longrightarrow> Q s)) \<and>  | 
|
48  | 
vc C I)"  | 
|
| 43158 | 49  | 
|
50  | 
||
| 50421 | 51  | 
text {* Soundness: *}
 | 
| 43158 | 52  | 
|
| 55003 | 53  | 
lemma vc_sound: "vc C Q \<Longrightarrow> \<turnstile> {pre C Q} strip C {Q}"
 | 
| 52332 | 54  | 
proof(induction C arbitrary: Q)  | 
55  | 
case (Awhile I b C)  | 
|
| 43158 | 56  | 
show ?case  | 
57  | 
proof(simp, rule While')  | 
|
| 55003 | 58  | 
from `vc (Awhile I b C) Q`  | 
59  | 
have vc: "vc C I" and IQ: "\<forall>s. I s \<and> \<not> bval b s \<longrightarrow> Q s" and  | 
|
| 52332 | 60  | 
pre: "\<forall>s. I s \<and> bval b s \<longrightarrow> pre C I s" by simp_all  | 
61  | 
    have "\<turnstile> {pre C I} strip C {I}" by(rule Awhile.IH[OF vc])
 | 
|
62  | 
    with pre show "\<turnstile> {\<lambda>s. I s \<and> bval b s} strip C {I}"
 | 
|
| 43158 | 63  | 
by(rule strengthen_pre)  | 
64  | 
show "\<forall>s. I s \<and> \<not>bval b s \<longrightarrow> Q s" by(rule IQ)  | 
|
65  | 
qed  | 
|
66  | 
qed (auto intro: hoare.conseq)  | 
|
67  | 
||
68  | 
corollary vc_sound':  | 
|
| 55003 | 69  | 
  "\<lbrakk> vc C Q; \<forall>s. P s \<longrightarrow> pre C Q s \<rbrakk> \<Longrightarrow> \<turnstile> {P} strip C {Q}"
 | 
| 43158 | 70  | 
by (metis strengthen_pre vc_sound)  | 
71  | 
||
72  | 
||
| 50421 | 73  | 
text{* Completeness: *}
 | 
| 43158 | 74  | 
|
75  | 
lemma pre_mono:  | 
|
| 52332 | 76  | 
"\<forall>s. P s \<longrightarrow> P' s \<Longrightarrow> pre C P s \<Longrightarrow> pre C P' s"  | 
77  | 
proof (induction C arbitrary: P P' s)  | 
|
| 47818 | 78  | 
case Aseq thus ?case by simp metis  | 
| 43158 | 79  | 
qed simp_all  | 
80  | 
||
81  | 
lemma vc_mono:  | 
|
| 55003 | 82  | 
"\<forall>s. P s \<longrightarrow> P' s \<Longrightarrow> vc C P \<Longrightarrow> vc C P'"  | 
| 52332 | 83  | 
proof(induction C arbitrary: P P')  | 
| 47818 | 84  | 
case Aseq thus ?case by simp (metis pre_mono)  | 
| 43158 | 85  | 
qed simp_all  | 
86  | 
||
87  | 
lemma vc_complete:  | 
|
| 55003 | 88  | 
 "\<turnstile> {P}c{Q} \<Longrightarrow> \<exists>C. strip C = c \<and> vc C Q \<and> (\<forall>s. P s \<longrightarrow> pre C Q s)"
 | 
| 52332 | 89  | 
(is "_ \<Longrightarrow> \<exists>C. ?G P c Q C")  | 
| 45015 | 90  | 
proof (induction rule: hoare.induct)  | 
| 43158 | 91  | 
case Skip  | 
| 52332 | 92  | 
show ?case (is "\<exists>C. ?C C")  | 
| 52193 | 93  | 
proof show "?C Askip" by simp qed  | 
| 43158 | 94  | 
next  | 
95  | 
case (Assign P a x)  | 
|
| 52332 | 96  | 
show ?case (is "\<exists>C. ?C C")  | 
| 43158 | 97  | 
proof show "?C(Aassign x a)" by simp qed  | 
98  | 
next  | 
|
| 47818 | 99  | 
case (Seq P c1 Q c2 R)  | 
| 52332 | 100  | 
from Seq.IH obtain C1 where ih1: "?G P c1 Q C1" by blast  | 
101  | 
from Seq.IH obtain C2 where ih2: "?G Q c2 R C2" by blast  | 
|
102  | 
show ?case (is "\<exists>C. ?C C")  | 
|
| 43158 | 103  | 
proof  | 
| 52332 | 104  | 
show "?C(Aseq C1 C2)"  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
43158 
diff
changeset
 | 
105  | 
using ih1 ih2 by (fastforce elim!: pre_mono vc_mono)  | 
| 43158 | 106  | 
qed  | 
107  | 
next  | 
|
108  | 
case (If P b c1 Q c2)  | 
|
| 52332 | 109  | 
from If.IH obtain C1 where ih1: "?G (\<lambda>s. P s \<and> bval b s) c1 Q C1"  | 
| 43158 | 110  | 
by blast  | 
| 52332 | 111  | 
from If.IH obtain C2 where ih2: "?G (\<lambda>s. P s \<and> \<not>bval b s) c2 Q C2"  | 
| 43158 | 112  | 
by blast  | 
| 52332 | 113  | 
show ?case (is "\<exists>C. ?C C")  | 
| 43158 | 114  | 
proof  | 
| 52332 | 115  | 
show "?C(Aif b C1 C2)" using ih1 ih2 by simp  | 
| 43158 | 116  | 
qed  | 
117  | 
next  | 
|
118  | 
case (While P b c)  | 
|
| 52332 | 119  | 
from While.IH obtain C where ih: "?G (\<lambda>s. P s \<and> bval b s) c P C" by blast  | 
120  | 
show ?case (is "\<exists>C. ?C C")  | 
|
121  | 
proof show "?C(Awhile P b C)" using ih by simp qed  | 
|
| 43158 | 122  | 
next  | 
123  | 
case conseq thus ?case by(fast elim!: pre_mono vc_mono)  | 
|
124  | 
qed  | 
|
125  | 
||
126  | 
end  |