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