author | wenzelm |
Thu, 12 Dec 2024 15:45:29 +0100 | |
changeset 81583 | b6df83045178 |
parent 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
50421 | 1 |
(* Author: Tobias Nipkow *) |
43158 | 2 |
|
68776
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
3 |
subsection "Verification Condition Generation" |
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
4 |
|
52269 | 5 |
theory VCG imports Hoare begin |
43158 | 6 |
|
68776
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
7 |
subsubsection "Annotated Commands" |
43158 | 8 |
|
68776
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
9 |
text\<open>Commands where loops are annotated with invariants.\<close> |
43158 | 10 |
|
58310 | 11 |
datatype acom = |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
68776
diff
changeset
|
12 |
Askip (\<open>SKIP\<close>) | |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
68776
diff
changeset
|
13 |
Aassign vname aexp (\<open>(_ ::= _)\<close> [1000, 61] 61) | |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
68776
diff
changeset
|
14 |
Aseq acom acom (\<open>_;;/ _\<close> [60, 61] 60) | |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
68776
diff
changeset
|
15 |
Aif bexp acom acom (\<open>(IF _/ THEN _/ ELSE _)\<close> [0, 0, 61] 61) | |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
68776
diff
changeset
|
16 |
Awhile assn bexp acom (\<open>({_}/ WHILE _/ DO _)\<close> [0, 0, 61] 61) |
43158 | 17 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
68776
diff
changeset
|
18 |
notation com.SKIP (\<open>SKIP\<close>) |
52267
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
nipkow
parents:
52193
diff
changeset
|
19 |
|
67406 | 20 |
text\<open>Strip annotations:\<close> |
52267
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
nipkow
parents:
52193
diff
changeset
|
21 |
|
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
nipkow
parents:
52193
diff
changeset
|
22 |
fun strip :: "acom \<Rightarrow> com" where |
54941 | 23 |
"strip SKIP = SKIP" | |
52267
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
nipkow
parents:
52193
diff
changeset
|
24 |
"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
|
25 |
"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
|
26 |
"strip (IF b THEN C\<^sub>1 ELSE C\<^sub>2) = (IF b THEN strip C\<^sub>1 ELSE strip C\<^sub>2)" | |
52332 | 27 |
"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
|
28 |
|
68776
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
29 |
subsubsection "Weeakest Precondistion and Verification Condition" |
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
30 |
|
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
31 |
text\<open>Weakest precondition:\<close> |
43158 | 32 |
|
33 |
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
|
34 |
"pre SKIP Q = Q" | |
2a16957cd379
used nice syntax, removed lemma because it makes a nice exercise.
nipkow
parents:
52193
diff
changeset
|
35 |
"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
|
36 |
"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
|
37 |
"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
|
38 |
(\<lambda>s. if bval b s then pre C\<^sub>1 Q s else pre C\<^sub>2 Q s)" | |
52332 | 39 |
"pre ({I} WHILE b DO C) Q = I" |
43158 | 40 |
|
67406 | 41 |
text\<open>Verification condition:\<close> |
43158 | 42 |
|
55003 | 43 |
fun vc :: "acom \<Rightarrow> assn \<Rightarrow> bool" where |
44 |
"vc SKIP Q = True" | |
|
45 |
"vc (x ::= a) Q = True" | |
|
46 |
"vc (C\<^sub>1;; C\<^sub>2) Q = (vc C\<^sub>1 (pre C\<^sub>2 Q) \<and> vc C\<^sub>2 Q)" | |
|
47 |
"vc (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q = (vc C\<^sub>1 Q \<and> vc C\<^sub>2 Q)" | |
|
52332 | 48 |
"vc ({I} WHILE b DO C) Q = |
55003 | 49 |
((\<forall>s. (I s \<and> bval b s \<longrightarrow> pre C I s) \<and> |
50 |
(I s \<and> \<not> bval b s \<longrightarrow> Q s)) \<and> |
|
51 |
vc C I)" |
|
43158 | 52 |
|
53 |
||
68776
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
54 |
subsubsection "Soundness" |
43158 | 55 |
|
55003 | 56 |
lemma vc_sound: "vc C Q \<Longrightarrow> \<turnstile> {pre C Q} strip C {Q}" |
52332 | 57 |
proof(induction C arbitrary: Q) |
58 |
case (Awhile I b C) |
|
43158 | 59 |
show ?case |
60 |
proof(simp, rule While') |
|
67406 | 61 |
from \<open>vc (Awhile I b C) Q\<close> |
55003 | 62 |
have vc: "vc C I" and IQ: "\<forall>s. I s \<and> \<not> bval b s \<longrightarrow> Q s" and |
52332 | 63 |
pre: "\<forall>s. I s \<and> bval b s \<longrightarrow> pre C I s" by simp_all |
64 |
have "\<turnstile> {pre C I} strip C {I}" by(rule Awhile.IH[OF vc]) |
|
65 |
with pre show "\<turnstile> {\<lambda>s. I s \<and> bval b s} strip C {I}" |
|
43158 | 66 |
by(rule strengthen_pre) |
67 |
show "\<forall>s. I s \<and> \<not>bval b s \<longrightarrow> Q s" by(rule IQ) |
|
68 |
qed |
|
69 |
qed (auto intro: hoare.conseq) |
|
70 |
||
71 |
corollary vc_sound': |
|
55003 | 72 |
"\<lbrakk> vc C Q; \<forall>s. P s \<longrightarrow> pre C Q s \<rbrakk> \<Longrightarrow> \<turnstile> {P} strip C {Q}" |
43158 | 73 |
by (metis strengthen_pre vc_sound) |
74 |
||
75 |
||
68776
403dd13cf6e9
avoid session qualification because no tex is generated when used;
nipkow
parents:
67406
diff
changeset
|
76 |
subsubsection "Completeness" |
43158 | 77 |
|
78 |
lemma pre_mono: |
|
52332 | 79 |
"\<forall>s. P s \<longrightarrow> P' s \<Longrightarrow> pre C P s \<Longrightarrow> pre C P' s" |
80 |
proof (induction C arbitrary: P P' s) |
|
47818 | 81 |
case Aseq thus ?case by simp metis |
43158 | 82 |
qed simp_all |
83 |
||
84 |
lemma vc_mono: |
|
55003 | 85 |
"\<forall>s. P s \<longrightarrow> P' s \<Longrightarrow> vc C P \<Longrightarrow> vc C P'" |
52332 | 86 |
proof(induction C arbitrary: P P') |
47818 | 87 |
case Aseq thus ?case by simp (metis pre_mono) |
43158 | 88 |
qed simp_all |
89 |
||
90 |
lemma vc_complete: |
|
55003 | 91 |
"\<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 | 92 |
(is "_ \<Longrightarrow> \<exists>C. ?G P c Q C") |
45015 | 93 |
proof (induction rule: hoare.induct) |
43158 | 94 |
case Skip |
52332 | 95 |
show ?case (is "\<exists>C. ?C C") |
52193 | 96 |
proof show "?C Askip" by simp qed |
43158 | 97 |
next |
98 |
case (Assign P a x) |
|
52332 | 99 |
show ?case (is "\<exists>C. ?C C") |
43158 | 100 |
proof show "?C(Aassign x a)" by simp qed |
101 |
next |
|
47818 | 102 |
case (Seq P c1 Q c2 R) |
52332 | 103 |
from Seq.IH obtain C1 where ih1: "?G P c1 Q C1" by blast |
104 |
from Seq.IH obtain C2 where ih2: "?G Q c2 R C2" by blast |
|
105 |
show ?case (is "\<exists>C. ?C C") |
|
43158 | 106 |
proof |
52332 | 107 |
show "?C(Aseq C1 C2)" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
43158
diff
changeset
|
108 |
using ih1 ih2 by (fastforce elim!: pre_mono vc_mono) |
43158 | 109 |
qed |
110 |
next |
|
111 |
case (If P b c1 Q c2) |
|
52332 | 112 |
from If.IH obtain C1 where ih1: "?G (\<lambda>s. P s \<and> bval b s) c1 Q C1" |
43158 | 113 |
by blast |
52332 | 114 |
from If.IH obtain C2 where ih2: "?G (\<lambda>s. P s \<and> \<not>bval b s) c2 Q C2" |
43158 | 115 |
by blast |
52332 | 116 |
show ?case (is "\<exists>C. ?C C") |
43158 | 117 |
proof |
52332 | 118 |
show "?C(Aif b C1 C2)" using ih1 ih2 by simp |
43158 | 119 |
qed |
120 |
next |
|
121 |
case (While P b c) |
|
52332 | 122 |
from While.IH obtain C where ih: "?G (\<lambda>s. P s \<and> bval b s) c P C" by blast |
123 |
show ?case (is "\<exists>C. ?C C") |
|
124 |
proof show "?C(Awhile P b C)" using ih by simp qed |
|
43158 | 125 |
next |
126 |
case conseq thus ?case by(fast elim!: pre_mono vc_mono) |
|
127 |
qed |
|
128 |
||
129 |
end |