src/HOL/IMP/Hoare_Total_EX.thy
author wenzelm
Fri, 10 Feb 2017 10:34:14 +0100
changeset 65020 a4fcaeadc825
parent 63538 d7b5e2a222c2
child 67019 7a3724078363
permissions -rw-r--r--
more operations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
     2
63538
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents: 63070
diff changeset
     3
theory Hoare_Total_EX
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents: 63070
diff changeset
     4
imports Hoare
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents: 63070
diff changeset
     5
begin
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
     6
63538
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents: 63070
diff changeset
     7
subsubsection "Hoare Logic for Total Correctness --- \<open>nat\<close>-Indexed Invariant"
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
     8
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
     9
text{* This is the standard set of rules that you find in many publications.
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    10
The While-rule is different from the one in Concrete Semantics in that the
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    11
invariant is indexed by natural numbers and goes down by 1 with
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    12
every iteration. The completeness proof is easier but the rule is harder
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    13
to apply in program proofs. *}
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    14
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    15
definition hoare_tvalid :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    16
  ("\<Turnstile>\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    17
"\<Turnstile>\<^sub>t {P}c{Q}  \<longleftrightarrow>  (\<forall>s. P s \<longrightarrow> (\<exists>t. (c,s) \<Rightarrow> t \<and> Q t))"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    18
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    19
inductive
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    20
  hoaret :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" ("\<turnstile>\<^sub>t ({(1_)}/ (_)/ {(1_)})" 50)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    21
where
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    22
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    23
Skip:  "\<turnstile>\<^sub>t {P} SKIP {P}"  |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    24
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    25
Assign:  "\<turnstile>\<^sub>t {\<lambda>s. P(s[a/x])} x::=a {P}"  |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    26
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    27
Seq: "\<lbrakk> \<turnstile>\<^sub>t {P\<^sub>1} c\<^sub>1 {P\<^sub>2}; \<turnstile>\<^sub>t {P\<^sub>2} c\<^sub>2 {P\<^sub>3} \<rbrakk> \<Longrightarrow> \<turnstile>\<^sub>t {P\<^sub>1} c\<^sub>1;;c\<^sub>2 {P\<^sub>3}"  |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    28
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    29
If: "\<lbrakk> \<turnstile>\<^sub>t {\<lambda>s. P s \<and> bval b s} c\<^sub>1 {Q}; \<turnstile>\<^sub>t {\<lambda>s. P s \<and> \<not> bval b s} c\<^sub>2 {Q} \<rbrakk>
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    30
  \<Longrightarrow> \<turnstile>\<^sub>t {P} IF b THEN c\<^sub>1 ELSE c\<^sub>2 {Q}"  |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    31
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    32
While:
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    33
  "\<lbrakk> \<And>n::nat. \<turnstile>\<^sub>t {P (Suc n)} c {P n};
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    34
     \<forall>n s. P (Suc n) s \<longrightarrow> bval b s;  \<forall>s. P 0 s \<longrightarrow> \<not> bval b s \<rbrakk>
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    35
   \<Longrightarrow> \<turnstile>\<^sub>t {\<lambda>s. \<exists>n. P n s} WHILE b DO c {P 0}"  |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    36
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    37
conseq: "\<lbrakk> \<forall>s. P' s \<longrightarrow> P s; \<turnstile>\<^sub>t {P}c{Q}; \<forall>s. Q s \<longrightarrow> Q' s  \<rbrakk> \<Longrightarrow>
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    38
           \<turnstile>\<^sub>t {P'}c{Q'}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    39
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    40
text{* Building in the consequence rule: *}
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    41
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    42
lemma strengthen_pre:
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    43
  "\<lbrakk> \<forall>s. P' s \<longrightarrow> P s;  \<turnstile>\<^sub>t {P} c {Q} \<rbrakk> \<Longrightarrow> \<turnstile>\<^sub>t {P'} c {Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    44
by (metis conseq)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    45
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    46
lemma weaken_post:
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    47
  "\<lbrakk> \<turnstile>\<^sub>t {P} c {Q};  \<forall>s. Q s \<longrightarrow> Q' s \<rbrakk> \<Longrightarrow>  \<turnstile>\<^sub>t {P} c {Q'}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    48
by (metis conseq)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    49
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    50
lemma Assign': "\<forall>s. P s \<longrightarrow> Q(s[a/x]) \<Longrightarrow> \<turnstile>\<^sub>t {P} x ::= a {Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    51
by (simp add: strengthen_pre[OF _ Assign])
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    52
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    53
text{* The soundness theorem: *}
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    54
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    55
theorem hoaret_sound: "\<turnstile>\<^sub>t {P}c{Q}  \<Longrightarrow>  \<Turnstile>\<^sub>t {P}c{Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    56
proof(unfold hoare_tvalid_def, induction rule: hoaret.induct)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    57
  case (While P c b)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    58
  {
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    59
    fix n s
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    60
    have "\<lbrakk> P n s \<rbrakk> \<Longrightarrow> \<exists>t. (WHILE b DO c, s) \<Rightarrow> t \<and> P 0 t"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    61
    proof(induction "n" arbitrary: s)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    62
      case 0 thus ?case using While.hyps(3) WhileFalse by blast
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    63
    next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    64
      case (Suc n)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    65
      thus ?case by (meson While.IH While.hyps(2) WhileTrue)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    66
    qed
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    67
  }
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    68
  thus ?case by auto
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    69
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    70
  case If thus ?case by auto blast
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    71
qed fastforce+
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    72
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    73
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    74
definition wpt :: "com \<Rightarrow> assn \<Rightarrow> assn" ("wp\<^sub>t") where
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    75
"wp\<^sub>t c Q  =  (\<lambda>s. \<exists>t. (c,s) \<Rightarrow> t \<and> Q t)"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    76
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    77
lemma [simp]: "wp\<^sub>t SKIP Q = Q"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    78
by(auto intro!: ext simp: wpt_def)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    79
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    80
lemma [simp]: "wp\<^sub>t (x ::= e) Q = (\<lambda>s. Q(s(x := aval e s)))"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    81
by(auto intro!: ext simp: wpt_def)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    82
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    83
lemma [simp]: "wp\<^sub>t (c\<^sub>1;;c\<^sub>2) Q = wp\<^sub>t c\<^sub>1 (wp\<^sub>t c\<^sub>2 Q)"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    84
unfolding wpt_def
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    85
apply(rule ext)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    86
apply auto
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    87
done
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    88
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    89
lemma [simp]:
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    90
 "wp\<^sub>t (IF b THEN c\<^sub>1 ELSE c\<^sub>2) Q = (\<lambda>s. wp\<^sub>t (if bval b s then c\<^sub>1 else c\<^sub>2) Q s)"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    91
apply(unfold wpt_def)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    92
apply(rule ext)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    93
apply auto
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    94
done
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    95
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    96
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    97
text{* Function @{text wpw} computes the weakest precondition of a While-loop
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    98
that is unfolded a fixed number of times. *}
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    99
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   100
fun wpw :: "bexp \<Rightarrow> com \<Rightarrow> nat \<Rightarrow> assn \<Rightarrow> assn" where
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   101
"wpw b c 0 Q s = (\<not> bval b s \<and> Q s)" |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   102
"wpw b c (Suc n) Q s = (bval b s \<and> (\<exists>s'. (c,s) \<Rightarrow> s' \<and>  wpw b c n Q s'))"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   103
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   104
lemma WHILE_Its: "(WHILE b DO c,s) \<Rightarrow> t \<Longrightarrow> Q t \<Longrightarrow> \<exists>n. wpw b c n Q s"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   105
proof(induction "WHILE b DO c" s t rule: big_step_induct)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   106
  case WhileFalse thus ?case using wpw.simps(1) by blast 
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   107
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   108
  case WhileTrue thus ?case using wpw.simps(2) by blast
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   109
qed
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   110
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   111
lemma wpt_is_pre: "\<turnstile>\<^sub>t {wp\<^sub>t c Q} c {Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   112
proof (induction c arbitrary: Q)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   113
  case SKIP show ?case by (auto intro:hoaret.Skip)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   114
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   115
  case Assign show ?case by (auto intro:hoaret.Assign)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   116
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   117
  case Seq thus ?case by (auto intro:hoaret.Seq)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   118
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   119
  case If thus ?case by (auto intro:hoaret.If hoaret.conseq)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   120
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   121
  case (While b c)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   122
  let ?w = "WHILE b DO c"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   123
  have c1: "\<forall>s. wp\<^sub>t ?w Q s \<longrightarrow> (\<exists>n. wpw b c n Q s)"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   124
    unfolding wpt_def by (metis WHILE_Its)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   125
  have c3: "\<forall>s. wpw b c 0 Q s \<longrightarrow> Q s" by simp
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   126
  have w2: "\<forall>n s. wpw b c (Suc n) Q s \<longrightarrow> bval b s" by simp
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   127
  have w3: "\<forall>s. wpw b c 0 Q s \<longrightarrow> \<not> bval b s" by simp
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   128
  { fix n
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   129
    have 1: "\<forall>s. wpw b c (Suc n) Q s \<longrightarrow> (\<exists>t. (c, s) \<Rightarrow> t \<and> wpw b c n Q t)"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   130
      by simp
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   131
    note strengthen_pre[OF 1 While.IH[of "wpw b c n Q", unfolded wpt_def]]
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   132
  }
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   133
  from conseq[OF c1 hoaret.While[OF this w2 w3] c3]
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   134
  show ?case .
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   135
qed
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   136
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   137
theorem hoaret_complete: "\<Turnstile>\<^sub>t {P}c{Q} \<Longrightarrow> \<turnstile>\<^sub>t {P}c{Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   138
apply(rule strengthen_pre[OF _ wpt_is_pre])
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   139
apply(auto simp: hoare_tvalid_def wpt_def)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   140
done
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   141
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   142
corollary hoaret_sound_complete: "\<turnstile>\<^sub>t {P}c{Q} \<longleftrightarrow> \<Turnstile>\<^sub>t {P}c{Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   143
by (metis hoaret_sound hoaret_complete)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   144
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   145
end