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