src/HOL/IMP/Abs_Int1.thy
author nipkow
Wed, 28 Sep 2011 09:55:11 +0200
changeset 45963 054a9ac0d7ef
child 45998 d2eb07a1e01b
permissions -rw-r--r--
Added Hoare-like Abstract Interpretation
nipkow@45963
     1
(* Author: Tobias Nipkow *)
nipkow@45963
     2
nipkow@45963
     3
theory Abs_Int1
nipkow@45963
     4
imports Abs_Int0_const
nipkow@45963
     5
begin
nipkow@45963
     6
nipkow@45963
     7
instantiation prod :: (preord,preord) preord
nipkow@45963
     8
begin
nipkow@45963
     9
nipkow@45963
    10
definition "le_prod p1 p2 = (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)"
nipkow@45963
    11
nipkow@45963
    12
instance
nipkow@45963
    13
proof
nipkow@45963
    14
  case goal1 show ?case by(simp add: le_prod_def)
nipkow@45963
    15
next
nipkow@45963
    16
  case goal2 thus ?case unfolding le_prod_def by(metis le_trans)
nipkow@45963
    17
qed
nipkow@45963
    18
nipkow@45963
    19
end
nipkow@45963
    20
nipkow@45963
    21
nipkow@45963
    22
subsection "Backward Analysis of Expressions"
nipkow@45963
    23
nipkow@45963
    24
hide_const bot
nipkow@45963
    25
nipkow@45963
    26
class L_top_bot = SL_top +
nipkow@45963
    27
fixes meet :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 65)
nipkow@45963
    28
and bot :: "'a" ("\<bottom>")
nipkow@45963
    29
assumes meet_le1 [simp]: "x \<sqinter> y \<sqsubseteq> x"
nipkow@45963
    30
and meet_le2 [simp]: "x \<sqinter> y \<sqsubseteq> y"
nipkow@45963
    31
and meet_greatest: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<sqinter> z"
nipkow@45963
    32
assumes bot[simp]: "\<bottom> \<sqsubseteq> x"
nipkow@45963
    33
nipkow@45963
    34
locale Rep1 = Rep rep for rep :: "'a::L_top_bot \<Rightarrow> 'b set" +
nipkow@45963
    35
assumes inter_rep_subset_rep_meet: "rep a1 \<inter> rep a2 \<subseteq> rep(a1 \<sqinter> a2)"
nipkow@45963
    36
  -- "this means the meet is precise"
nipkow@45963
    37
and rep_Bot: "rep \<bottom> = {}"
nipkow@45963
    38
begin
nipkow@45963
    39
nipkow@45963
    40
lemma in_rep_meet: "x <: a1 \<Longrightarrow> x <: a2 \<Longrightarrow> x <: a1 \<sqinter> a2"
nipkow@45963
    41
by (metis IntI inter_rep_subset_rep_meet set_mp)
nipkow@45963
    42
nipkow@45963
    43
lemma rep_meet[simp]: "rep(a1 \<sqinter> a2) = rep a1 \<inter> rep a2"
nipkow@45963
    44
by (metis equalityI inter_rep_subset_rep_meet le_inf_iff le_rep meet_le1 meet_le2)
nipkow@45963
    45
nipkow@45963
    46
lemma mono_meet: "x \<sqsubseteq> x' \<Longrightarrow> y \<sqsubseteq> y' \<Longrightarrow> x \<sqinter> y \<sqsubseteq> x' \<sqinter> y'"
nipkow@45963
    47
by (metis meet_greatest meet_le1 meet_le2 le_trans)
nipkow@45963
    48
nipkow@45963
    49
end
nipkow@45963
    50
nipkow@45963
    51
nipkow@45963
    52
locale Val_abs1 = Val_abs rep num' plus' + Rep1 rep
nipkow@45963
    53
  for rep :: "'a::L_top_bot \<Rightarrow> int set" and num' plus' +
nipkow@45963
    54
fixes filter_plus' :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a * 'a"
nipkow@45963
    55
and filter_less' :: "bool \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a * 'a"
nipkow@45963
    56
assumes filter_plus': "filter_plus' a a1 a2 = (a1',a2') \<Longrightarrow>
nipkow@45963
    57
  n1 <: a1 \<Longrightarrow> n2 <: a2 \<Longrightarrow> n1+n2 <: a \<Longrightarrow> n1 <: a1' \<and> n2 <: a2'"
nipkow@45963
    58
and filter_less': "filter_less' (n1<n2) a1 a2 = (a1',a2') \<Longrightarrow>
nipkow@45963
    59
  n1 <: a1 \<Longrightarrow> n2 <: a2 \<Longrightarrow> n1 <: a1' \<and> n2 <: a2'"
nipkow@45963
    60
and mono_filter_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> r \<sqsubseteq> r' \<Longrightarrow>
nipkow@45963
    61
  filter_plus' r a1 a2 \<sqsubseteq> filter_plus' r' b1 b2"
nipkow@45963
    62
and mono_filter_less': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow>
nipkow@45963
    63
  filter_less' bv a1 a2 \<sqsubseteq> filter_less' bv b1 b2"
nipkow@45963
    64
nipkow@45963
    65
locale Abs_Int1 = Val_abs1 +
nipkow@45963
    66
fixes pfp :: "('a st up acom \<Rightarrow> 'a st up acom) \<Rightarrow> 'a st up acom \<Rightarrow> 'a st up acom"
nipkow@45963
    67
assumes pfp: "\<forall>c. strip(f c) = strip c \<Longrightarrow> mono f \<Longrightarrow> f(pfp f c) \<sqsubseteq> pfp f c"
nipkow@45963
    68
and strip_pfp: "\<forall>c. strip(f c) = strip c \<Longrightarrow> strip(pfp f c) = strip c"
nipkow@45963
    69
begin
nipkow@45963
    70
nipkow@45963
    71
lemma in_rep_join_UpI: "s <:up S1 | s <:up S2 \<Longrightarrow> s <:up S1 \<squnion> S2"
nipkow@45963
    72
by (metis join_ge1 join_ge2 up_fun_in_rep_le)
nipkow@45963
    73
nipkow@45963
    74
fun aval' :: "aexp \<Rightarrow> 'a st up \<Rightarrow> 'a" where
nipkow@45963
    75
"aval' _ Bot = \<bottom>" |
nipkow@45963
    76
"aval' (N n) _ = num' n" |
nipkow@45963
    77
"aval' (V x) (Up S) = lookup S x" |
nipkow@45963
    78
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
nipkow@45963
    79
nipkow@45963
    80
lemma aval'_sound: "s <:up S \<Longrightarrow> aval a s <: aval' a S"
nipkow@45963
    81
by (induct a)(auto simp: rep_num' rep_plus' in_rep_up_iff lookup_def rep_st_def)
nipkow@45963
    82
nipkow@45963
    83
fun afilter :: "aexp \<Rightarrow> 'a \<Rightarrow> 'a st up \<Rightarrow> 'a st up" where
nipkow@45963
    84
"afilter (N n) a S = (if n <: a then S else Bot)" |
nipkow@45963
    85
"afilter (V x) a S = (case S of Bot \<Rightarrow> Bot | Up S \<Rightarrow>
nipkow@45963
    86
  let a' = lookup S x \<sqinter> a in
nipkow@45963
    87
  if a' \<sqsubseteq> \<bottom> then Bot else Up(update S x a'))" |
nipkow@45963
    88
"afilter (Plus e1 e2) a S =
nipkow@45963
    89
 (let (a1,a2) = filter_plus' a (aval' e1 S) (aval' e2 S)
nipkow@45963
    90
  in afilter e1 a1 (afilter e2 a2 S))"
nipkow@45963
    91
nipkow@45963
    92
text{* The test for @{const Bot} in the @{const V}-case is important: @{const
nipkow@45963
    93
Bot} indicates that a variable has no possible values, i.e.\ that the current
nipkow@45963
    94
program point is unreachable. But then the abstract state should collapse to
nipkow@45963
    95
@{const bot}. Put differently, we maintain the invariant that in an abstract
nipkow@45963
    96
state all variables are mapped to non-@{const Bot} values. Otherwise the
nipkow@45963
    97
(pointwise) join of two abstract states, one of which contains @{const Bot}
nipkow@45963
    98
values, may produce too large a result, thus making the analysis less
nipkow@45963
    99
precise. *}
nipkow@45963
   100
nipkow@45963
   101
nipkow@45963
   102
fun bfilter :: "bexp \<Rightarrow> bool \<Rightarrow> 'a st up \<Rightarrow> 'a st up" where
nipkow@45963
   103
"bfilter (B bv) res S = (if bv=res then S else Bot)" |
nipkow@45963
   104
"bfilter (Not b) res S = bfilter b (\<not> res) S" |
nipkow@45963
   105
"bfilter (And b1 b2) res S =
nipkow@45963
   106
  (if res then bfilter b1 True (bfilter b2 True S)
nipkow@45963
   107
   else bfilter b1 False S \<squnion> bfilter b2 False S)" |
nipkow@45963
   108
"bfilter (Less e1 e2) res S =
nipkow@45963
   109
  (let (res1,res2) = filter_less' res (aval' e1 S) (aval' e2 S)
nipkow@45963
   110
   in afilter e1 res1 (afilter e2 res2 S))"
nipkow@45963
   111
nipkow@45963
   112
lemma afilter_sound: "s <:up S \<Longrightarrow> aval e s <: a \<Longrightarrow> s <:up afilter e a S"
nipkow@45963
   113
proof(induction e arbitrary: a S)
nipkow@45963
   114
  case N thus ?case by simp
nipkow@45963
   115
next
nipkow@45963
   116
  case (V x)
nipkow@45963
   117
  obtain S' where "S = Up S'" and "s <:f S'" using `s <:up S`
nipkow@45963
   118
    by(auto simp: in_rep_up_iff)
nipkow@45963
   119
  moreover hence "s x <: lookup S' x" by(simp add: rep_st_def)
nipkow@45963
   120
  moreover have "s x <: a" using V by simp
nipkow@45963
   121
  ultimately show ?case using V(1)
nipkow@45963
   122
    by(simp add: lookup_update Let_def rep_st_def)
nipkow@45963
   123
      (metis le_rep emptyE in_rep_meet rep_Bot subset_empty)
nipkow@45963
   124
next
nipkow@45963
   125
  case (Plus e1 e2) thus ?case
nipkow@45963
   126
    using filter_plus'[OF _ aval'_sound[OF Plus(3)] aval'_sound[OF Plus(3)]]
nipkow@45963
   127
    by (auto split: prod.split)
nipkow@45963
   128
qed
nipkow@45963
   129
nipkow@45963
   130
lemma bfilter_sound: "s <:up S \<Longrightarrow> bv = bval b s \<Longrightarrow> s <:up bfilter b bv S"
nipkow@45963
   131
proof(induction b arbitrary: S bv)
nipkow@45963
   132
  case B thus ?case by simp
nipkow@45963
   133
next
nipkow@45963
   134
  case (Not b) thus ?case by simp
nipkow@45963
   135
next
nipkow@45963
   136
  case (And b1 b2) thus ?case by(fastforce simp: in_rep_join_UpI)
nipkow@45963
   137
next
nipkow@45963
   138
  case (Less e1 e2) thus ?case
nipkow@45963
   139
    by (auto split: prod.split)
nipkow@45963
   140
       (metis afilter_sound filter_less' aval'_sound Less)
nipkow@45963
   141
qed
nipkow@45963
   142
nipkow@45963
   143
nipkow@45963
   144
fun step :: "'a st up \<Rightarrow> 'a st up acom \<Rightarrow> 'a st up acom" where
nipkow@45963
   145
"step S (SKIP {P}) = (SKIP {S})" |
nipkow@45963
   146
"step S (x ::= e {P}) =
nipkow@45963
   147
  x ::= e {case S of Bot \<Rightarrow> Bot
nipkow@45963
   148
           | Up S \<Rightarrow> Up(update S x (aval' e (Up S)))}" |
nipkow@45963
   149
"step S (c1; c2) = step S c1; step (post c1) c2" |
nipkow@45963
   150
"step S (IF b THEN c1 ELSE c2 {P}) =
nipkow@45963
   151
  (let c1' = step (bfilter b True S) c1; c2' = step (bfilter b False S) c2
nipkow@45963
   152
   in IF b THEN c1' ELSE c2' {post c1 \<squnion> post c2})" |
nipkow@45963
   153
"step S ({Inv} WHILE b DO c {P}) =
nipkow@45963
   154
   {S \<squnion> post c}
nipkow@45963
   155
   WHILE b DO step (bfilter b True Inv) c
nipkow@45963
   156
   {bfilter b False Inv}"
nipkow@45963
   157
nipkow@45963
   158
lemma strip_step[simp]: "strip(step S c) = strip c"
nipkow@45963
   159
by(induct c arbitrary: S) (simp_all add: Let_def)
nipkow@45963
   160
nipkow@45963
   161
nipkow@45963
   162
definition AI :: "com \<Rightarrow> 'a st up acom" where
nipkow@45963
   163
"AI c = pfp (step \<top>) (\<bottom>\<^sub>c c)"
nipkow@45963
   164
nipkow@45963
   165
nipkow@45963
   166
subsubsection "Monotonicity"
nipkow@45963
   167
nipkow@45963
   168
lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'"
nipkow@45963
   169
apply(cases S)
nipkow@45963
   170
 apply simp
nipkow@45963
   171
apply(cases S')
nipkow@45963
   172
 apply simp
nipkow@45963
   173
apply simp
nipkow@45963
   174
by(induction e) (auto simp: le_st_def lookup_def mono_plus')
nipkow@45963
   175
nipkow@45963
   176
lemma mono_afilter: "r \<sqsubseteq> r' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> afilter e r S \<sqsubseteq> afilter e r' S'"
nipkow@45963
   177
apply(induction e arbitrary: r r' S S')
nipkow@45963
   178
apply(auto simp: Let_def split: up.splits prod.splits)
nipkow@45963
   179
apply (metis le_rep subsetD)
nipkow@45963
   180
apply(drule_tac x = "list" in mono_lookup)
nipkow@45963
   181
apply (metis mono_meet le_trans)
nipkow@45963
   182
apply (metis mono_meet mono_lookup mono_update le_trans)
nipkow@45963
   183
apply(metis mono_aval' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv)
nipkow@45963
   184
done
nipkow@45963
   185
nipkow@45963
   186
lemma mono_bfilter: "S \<sqsubseteq> S' \<Longrightarrow> bfilter b r S \<sqsubseteq> bfilter b r S'"
nipkow@45963
   187
apply(induction b arbitrary: r S S')
nipkow@45963
   188
apply(auto simp: le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] split: prod.splits)
nipkow@45963
   189
apply(metis mono_aval' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv)
nipkow@45963
   190
done
nipkow@45963
   191
nipkow@45963
   192
nipkow@45963
   193
lemma post_le_post: "c \<sqsubseteq> c' \<Longrightarrow> post c \<sqsubseteq> post c'"
nipkow@45963
   194
by (induction c c' rule: le_acom.induct) simp_all
nipkow@45963
   195
nipkow@45963
   196
lemma mono_step: "S \<sqsubseteq> S' \<Longrightarrow> c \<sqsubseteq> c' \<Longrightarrow> step S c \<sqsubseteq> step S' c'"
nipkow@45963
   197
apply(induction c c' arbitrary: S S' rule: le_acom.induct)
nipkow@45963
   198
apply (auto simp: post_le_post Let_def mono_bfilter mono_update mono_aval' le_join_disj
nipkow@45963
   199
  split: up.split)
nipkow@45963
   200
done
nipkow@45963
   201
nipkow@45963
   202
nipkow@45963
   203
subsubsection "Soundness"
nipkow@45963
   204
nipkow@45963
   205
lemma in_rep_update: "\<lbrakk> s <:f S; i <: a \<rbrakk> \<Longrightarrow> s(x := i) <:f update S x a"
nipkow@45963
   206
by(simp add: rep_st_def lookup_update)
nipkow@45963
   207
nipkow@45963
   208
lemma While_final_False: "(WHILE b DO c, s) \<Rightarrow> t \<Longrightarrow> \<not> bval b t"
nipkow@45963
   209
by(induct "WHILE b DO c" s t rule: big_step_induct) simp_all
nipkow@45963
   210
nipkow@45963
   211
lemma step_sound:
nipkow@45963
   212
  "step S c \<sqsubseteq> c \<Longrightarrow> (strip c,s) \<Rightarrow> t \<Longrightarrow> s <:up S \<Longrightarrow> t <:up post c"
nipkow@45963
   213
proof(induction c arbitrary: S s t)
nipkow@45963
   214
  case SKIP thus ?case
nipkow@45963
   215
    by simp (metis skipE up_fun_in_rep_le)
nipkow@45963
   216
next
nipkow@45963
   217
  case Assign thus ?case
nipkow@45963
   218
    apply (auto simp del: fun_upd_apply split: up.splits)
nipkow@45963
   219
    by (metis aval'_sound fun_in_rep_le in_rep_update rep_up.simps(2))
nipkow@45963
   220
next
nipkow@45963
   221
  case Semi thus ?case by simp blast
nipkow@45963
   222
next
nipkow@45963
   223
  case (If b c1 c2 S0)
nipkow@45963
   224
  show ?case
nipkow@45963
   225
  proof cases
nipkow@45963
   226
    assume "bval b s"
nipkow@45963
   227
    with If.prems have 1: "step (bfilter b True S) c1 \<sqsubseteq> c1"
nipkow@45963
   228
      and 2: "(strip c1, s) \<Rightarrow> t" and 3: "post c1 \<sqsubseteq> S0" by(auto simp: Let_def)
nipkow@45963
   229
    from If.IH(1)[OF 1 2 bfilter_sound[OF `s <:up S`]] `bval b s` 3
nipkow@45963
   230
    show ?thesis by simp (metis up_fun_in_rep_le)
nipkow@45963
   231
  next
nipkow@45963
   232
    assume "\<not> bval b s"
nipkow@45963
   233
    with If.prems have 1: "step (bfilter b False S) c2 \<sqsubseteq> c2"
nipkow@45963
   234
      and 2: "(strip c2, s) \<Rightarrow> t" and 3: "post c2 \<sqsubseteq> S0" by(auto simp: Let_def)
nipkow@45963
   235
    from If.IH(2)[OF 1 2 bfilter_sound[OF `s <:up S`]] `\<not> bval b s` 3
nipkow@45963
   236
    show ?thesis by simp (metis up_fun_in_rep_le)
nipkow@45963
   237
  qed
nipkow@45963
   238
next
nipkow@45963
   239
  case (While Inv b c P)
nipkow@45963
   240
  from While.prems have inv: "step (bfilter b True Inv) c \<sqsubseteq> c"
nipkow@45963
   241
    and "post c \<sqsubseteq> Inv" and "S \<sqsubseteq> Inv" and "bfilter b False Inv \<sqsubseteq> P"
nipkow@45963
   242
    by(auto simp: Let_def)
nipkow@45963
   243
  { fix s t have "(WHILE b DO strip c,s) \<Rightarrow> t \<Longrightarrow> s <:up Inv \<Longrightarrow> t <:up Inv"
nipkow@45963
   244
    proof(induction "WHILE b DO strip c" s t rule: big_step_induct)
nipkow@45963
   245
      case WhileFalse thus ?case by simp
nipkow@45963
   246
    next
nipkow@45963
   247
      case (WhileTrue s1 s2 s3)
nipkow@45963
   248
      from WhileTrue.hyps(5)[OF up_fun_in_rep_le[OF While.IH[OF inv `(strip c, s1) \<Rightarrow> s2` bfilter_sound[OF `s1 <:up Inv`]] `post c \<sqsubseteq> Inv`]] `bval b s1`
nipkow@45963
   249
      show ?case by simp
nipkow@45963
   250
    qed
nipkow@45963
   251
  } note Inv = this
nipkow@45963
   252
  from  While.prems(2) have "(WHILE b DO strip c, s) \<Rightarrow> t" and "\<not> bval b t"
nipkow@45963
   253
    by(auto dest: While_final_False)
nipkow@45963
   254
  from Inv[OF this(1) up_fun_in_rep_le[OF `s <:up S` `S \<sqsubseteq> Inv`]]
nipkow@45963
   255
  have "t <:up Inv" .
nipkow@45963
   256
  from up_fun_in_rep_le[OF bfilter_sound[OF this]  `bfilter b False Inv \<sqsubseteq> P`]
nipkow@45963
   257
  show ?case using `\<not> bval b t` by simp
nipkow@45963
   258
qed
nipkow@45963
   259
nipkow@45963
   260
lemma AI_sound: "(c,s) \<Rightarrow> t \<Longrightarrow> t <:up post(AI c)"
nipkow@45963
   261
by(fastforce simp: AI_def strip_pfp mono_def in_rep_Top_up
nipkow@45963
   262
  intro: step_sound pfp  mono_step[OF le_refl])
nipkow@45963
   263
nipkow@45963
   264
end
nipkow@45963
   265
nipkow@45963
   266
end