src/HOL/IMP/VCG_Total_EX2.thy
author wenzelm
Sun, 27 May 2018 22:37:08 +0200
changeset 68300 cd8ab1a7a286
parent 67613 ce654b0e6d69
child 68776 403dd13cf6e9
permissions -rw-r--r--
retain isolated blob nodes (amending deb2fcbda16e): avoid failure of Session.handle_change with "Missing blob", when opening theory with load command later;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
     2
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
     3
theory VCG_Total_EX2
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
     4
imports Hoare_Total_EX2
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
     5
begin
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
     6
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
     7
subsection "Verification Conditions for Total Correctness"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
     8
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
     9
text \<open>
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    10
Theory \<open>VCG_Total_EX\<close> conatins a VCG built on top of a Hoare logic without logical variables.
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    11
As a result the completeness proof runs into a problem. This theory uses a Hoare logic
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    12
with logical variables and proves soundness and completeness.
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    13
\<close>
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    14
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    15
text\<open>Annotated commands: commands where loops are annotated with
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    16
invariants.\<close>
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    17
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    18
datatype acom =
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    19
  Askip                  ("SKIP") |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    20
  Aassign vname aexp     ("(_ ::= _)" [1000, 61] 61) |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    21
  Aseq   acom acom       ("_;;/ _"  [60, 61] 60) |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    22
  Aif bexp acom acom     ("(IF _/ THEN _/ ELSE _)"  [0, 0, 61] 61) |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    23
  Awhile assn2 lvname bexp acom
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    24
    ("({_'/_}/ WHILE _/ DO _)"  [0, 0, 0, 61] 61)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    25
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    26
notation com.SKIP ("SKIP")
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    27
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    28
text\<open>Strip annotations:\<close>
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    29
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    30
fun strip :: "acom \<Rightarrow> com" where
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    31
"strip SKIP = SKIP" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    32
"strip (x ::= a) = (x ::= a)" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    33
"strip (C\<^sub>1;; C\<^sub>2) = (strip C\<^sub>1;; strip C\<^sub>2)" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    34
"strip (IF b THEN C\<^sub>1 ELSE C\<^sub>2) = (IF b THEN strip C\<^sub>1 ELSE strip C\<^sub>2)" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    35
"strip ({_/_} WHILE b DO C) = (WHILE b DO strip C)"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    36
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    37
text\<open>Weakest precondition from annotated commands:\<close>
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    38
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    39
fun pre :: "acom \<Rightarrow> assn2 \<Rightarrow> assn2" where
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    40
"pre SKIP Q = Q" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    41
"pre (x ::= a) Q = (\<lambda>l s. Q l (s(x := aval a s)))" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    42
"pre (C\<^sub>1;; C\<^sub>2) Q = pre C\<^sub>1 (pre C\<^sub>2 Q)" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    43
"pre (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q =
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    44
  (\<lambda>l s. if bval b s then pre C\<^sub>1 Q l s else pre C\<^sub>2 Q l s)" |
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67406
diff changeset
    45
"pre ({I/x} WHILE b DO C) Q = (\<lambda>l s. \<exists>n. I (l(x:=n)) s)"
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    46
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    47
text\<open>Verification condition:\<close>
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    48
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    49
fun vc :: "acom \<Rightarrow> assn2 \<Rightarrow> bool" where
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    50
"vc SKIP Q = True" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    51
"vc (x ::= a) Q = True" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    52
"vc (C\<^sub>1;; C\<^sub>2) Q = (vc C\<^sub>1 (pre C\<^sub>2 Q) \<and> vc C\<^sub>2 Q)" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    53
"vc (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q = (vc C\<^sub>1 Q \<and> vc C\<^sub>2 Q)" |
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    54
"vc ({I/x} WHILE b DO C) Q =
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    55
  (\<forall>l s. (I (l(x:=Suc(l x))) s \<longrightarrow> pre C I l s) \<and>
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    56
       (l x > 0 \<and> I l s \<longrightarrow> bval b s) \<and>
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    57
       (I (l(x := 0)) s \<longrightarrow> \<not> bval b s \<and> Q l s) \<and>
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    58
       vc C I)"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    59
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    60
lemma vc_sound: "vc C Q \<Longrightarrow> \<turnstile>\<^sub>t {pre C Q} strip C {Q}"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    61
proof(induction C arbitrary: Q)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    62
  case (Awhile I x b C)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    63
  show ?case
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    64
  proof(simp, rule weaken_post[OF While[of I x]], goal_cases)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    65
    case 1 show ?case
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    66
      using Awhile.IH[of "I"] Awhile.prems by (auto intro: strengthen_pre)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    67
  next
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    68
    case 3 show ?case
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    69
      using Awhile.prems by (simp) (metis fun_upd_triv)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    70
  qed (insert Awhile.prems, auto)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    71
qed (auto intro: conseq Seq If simp: Skip Assign)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    72
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    73
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    74
text\<open>Completeness:\<close>
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    75
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    76
lemma pre_mono:
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    77
  "\<forall>l s. P l s \<longrightarrow> P' l s \<Longrightarrow> pre C P l s \<Longrightarrow> pre C P' l s"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    78
proof (induction C arbitrary: P P' l s)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    79
  case Aseq thus ?case by simp metis
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    80
qed simp_all
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    81
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    82
lemma vc_mono:
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    83
  "\<forall>l s. P l s \<longrightarrow> P' l s \<Longrightarrow> vc C P \<Longrightarrow> vc C P'"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    84
proof(induction C arbitrary: P P')
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    85
  case Aseq thus ?case by simp (metis pre_mono)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    86
qed simp_all
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    87
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    88
lemma vc_complete:
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    89
 "\<turnstile>\<^sub>t {P}c{Q} \<Longrightarrow> \<exists>C. strip C = c \<and> vc C Q \<and> (\<forall>l s. P l s \<longrightarrow> pre C Q l s)"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    90
  (is "_ \<Longrightarrow> \<exists>C. ?G P c Q C")
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    91
proof (induction rule: hoaret.induct)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    92
  case Skip
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    93
  show ?case (is "\<exists>C. ?C C")
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    94
  proof show "?C Askip" by simp qed
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    95
next
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    96
  case (Assign P a x)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    97
  show ?case (is "\<exists>C. ?C C")
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    98
  proof show "?C(Aassign x a)" by simp qed
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
    99
next
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   100
  case (Seq P c1 Q c2 R)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   101
  from Seq.IH obtain C1 where ih1: "?G P c1 Q C1" by blast
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   102
  from Seq.IH obtain C2 where ih2: "?G Q c2 R C2" by blast
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   103
  show ?case (is "\<exists>C. ?C C")
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   104
  proof
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   105
    show "?C(Aseq C1 C2)"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   106
      using ih1 ih2 by (fastforce elim!: pre_mono vc_mono)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   107
  qed
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   108
next
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   109
  case (If P b c1 Q c2)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   110
  from If.IH obtain C1 where ih1: "?G (\<lambda>l s. P l s \<and> bval b s) c1 Q C1"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   111
    by blast
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   112
  from If.IH obtain C2 where ih2: "?G (\<lambda>l s. P l s \<and> \<not>bval b s) c2 Q C2"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   113
    by blast
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   114
  show ?case (is "\<exists>C. ?C C")
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   115
  proof
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   116
    show "?C(Aif b C1 C2)" using ih1 ih2 by simp
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   117
  qed
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   118
next
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   119
  case (While P x c b)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   120
  from While.IH obtain C where
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   121
    ih: "?G (\<lambda>l s. P (l(x:=Suc(l x))) s \<and> bval b s) c P C"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   122
    by blast
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   123
  show ?case (is "\<exists>C. ?C C")
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   124
  proof
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   125
    have "vc ({P/x} WHILE b DO C) (\<lambda>l. P (l(x := 0)))"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   126
      using ih While.hyps(2,3)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   127
      by simp (metis fun_upd_same zero_less_Suc)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   128
    thus "?C(Awhile P x b C)" using ih by simp
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   129
 qed
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   130
next
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   131
  case conseq thus ?case by(fast elim!: pre_mono vc_mono)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   132
qed
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   133
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents:
diff changeset
   134
end