src/HOL/IMP/VCG_Total_EX.thy
author nipkow
Sat, 23 Jul 2016 13:25:44 +0200
changeset 63538 d7b5e2a222c2
child 67019 7a3724078363
permissions -rw-r--r--
added new vcg based on existentially quantified while-rule
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63538
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
     2
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
     3
theory VCG_Total_EX
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
     4
imports "~~/src/HOL/IMP/Hoare_Total_EX"
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
     5
begin
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
     6
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
     7
subsection "Verification Conditions for Total Correctness"
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
     8
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
     9
text{* Annotated commands: commands where loops are annotated with
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    10
invariants. *}
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    11
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    12
datatype acom =
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    13
  Askip                  ("SKIP") |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    14
  Aassign vname aexp     ("(_ ::= _)" [1000, 61] 61) |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    15
  Aseq   acom acom       ("_;;/ _"  [60, 61] 60) |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    16
  Aif bexp acom acom     ("(IF _/ THEN _/ ELSE _)"  [0, 0, 61] 61) |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    17
  Awhile "nat \<Rightarrow> assn" bexp acom
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    18
    ("({_}/ WHILE _/ DO _)"  [0, 0, 61] 61)
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    19
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    20
notation com.SKIP ("SKIP")
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    21
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    22
text{* Strip annotations: *}
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    23
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    24
fun strip :: "acom \<Rightarrow> com" where
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    25
"strip SKIP = SKIP" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    26
"strip (x ::= a) = (x ::= a)" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    27
"strip (C\<^sub>1;; C\<^sub>2) = (strip C\<^sub>1;; strip C\<^sub>2)" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    28
"strip (IF b THEN C\<^sub>1 ELSE C\<^sub>2) = (IF b THEN strip C\<^sub>1 ELSE strip C\<^sub>2)" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    29
"strip ({_} WHILE b DO C) = (WHILE b DO strip C)"
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    30
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    31
text{* Weakest precondition from annotated commands: *}
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    32
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    33
fun pre :: "acom \<Rightarrow> assn \<Rightarrow> assn" where
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    34
"pre SKIP Q = Q" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    35
"pre (x ::= a) Q = (\<lambda>s. Q(s(x := aval a s)))" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    36
"pre (C\<^sub>1;; C\<^sub>2) Q = pre C\<^sub>1 (pre C\<^sub>2 Q)" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    37
"pre (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q =
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    38
  (\<lambda>s. if bval b s then pre C\<^sub>1 Q s else pre C\<^sub>2 Q s)" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    39
"pre ({I} WHILE b DO C) Q = (\<lambda>s. EX n. I n s)"
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    40
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    41
text{* Verification condition: *}
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    42
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    43
fun vc :: "acom \<Rightarrow> assn \<Rightarrow> bool" where
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    44
"vc SKIP Q = True" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    45
"vc (x ::= a) Q = True" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    46
"vc (C\<^sub>1;; C\<^sub>2) Q = (vc C\<^sub>1 (pre C\<^sub>2 Q) \<and> vc C\<^sub>2 Q)" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    47
"vc (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q = (vc C\<^sub>1 Q \<and> vc C\<^sub>2 Q)" |
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    48
"vc ({I} WHILE b DO C) Q =
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    49
  (\<forall>s n. (I (Suc n) s \<longrightarrow> pre C (I n) s) \<and>
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    50
       (I (Suc n) s \<longrightarrow> bval b s) \<and>
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    51
       (I 0 s \<longrightarrow> \<not> bval b s \<and> Q s) \<and>
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    52
       vc C (I n))"
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    53
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    54
lemma vc_sound: "vc C Q \<Longrightarrow> \<turnstile>\<^sub>t {pre C Q} strip C {Q}"
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    55
proof(induction C arbitrary: Q)
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    56
  case (Awhile I b C)
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    57
  show ?case
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    58
  proof(simp, rule conseq[OF _ While[of I]], goal_cases)
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    59
    case (2 n) show ?case
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    60
      using Awhile.IH[of "I n"] Awhile.prems
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    61
      by (auto intro: strengthen_pre)
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    62
  qed (insert Awhile.prems, auto)
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    63
qed (auto intro: conseq Seq If simp: Skip Assign)
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    64
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    65
text\<open>When trying to extend the completeness proof of the VCG for partial correctness
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    66
to total correctness one runs into the following problem.
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    67
In the case of the while-rule, the universally quantified \<open>n\<close> in the first premise
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    68
means that for that premise the induction hypothesis does not yield a single
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    69
annotated command \<open>C\<close> but merely that for every \<open>n\<close> such a \<open>C\<close> exists.\<close>
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    70
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents:
diff changeset
    71
end