|
44656
|
1 |
(* Author: Tobias Nipkow *)
|
|
|
2 |
|
|
|
3 |
theory AbsInt1
|
|
|
4 |
imports AbsInt0_const
|
|
|
5 |
begin
|
|
|
6 |
|
|
|
7 |
subsection "Backward Analysis of Expressions"
|
|
|
8 |
|
|
|
9 |
class L_top_bot = SL_top +
|
|
|
10 |
fixes meet :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 65)
|
|
|
11 |
and Bot :: "'a"
|
|
|
12 |
assumes meet_le1 [simp]: "x \<sqinter> y \<sqsubseteq> x"
|
|
|
13 |
and meet_le2 [simp]: "x \<sqinter> y \<sqsubseteq> y"
|
|
|
14 |
and meet_greatest: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<sqinter> z"
|
|
|
15 |
assumes bot[simp]: "Bot \<sqsubseteq> x"
|
|
|
16 |
|
|
|
17 |
locale Rep1 = Rep rep for rep :: "'a::L_top_bot \<Rightarrow> 'b set" +
|
|
|
18 |
assumes inter_rep_subset_rep_meet: "rep a1 \<inter> rep a2 \<subseteq> rep(a1 \<sqinter> a2)"
|
|
|
19 |
and rep_Bot: "rep Bot = {}"
|
|
|
20 |
begin
|
|
|
21 |
|
|
|
22 |
lemma in_rep_meet: "x <: a1 \<Longrightarrow> x <: a2 \<Longrightarrow> x <: a1 \<sqinter> a2"
|
|
|
23 |
by (metis IntI inter_rep_subset_rep_meet set_mp)
|
|
|
24 |
|
|
|
25 |
lemma rep_meet[simp]: "rep(a1 \<sqinter> a2) = rep a1 \<inter> rep a2"
|
|
|
26 |
by (metis equalityI inter_rep_subset_rep_meet le_inf_iff le_rep meet_le1 meet_le2)
|
|
|
27 |
|
|
|
28 |
end
|
|
|
29 |
|
|
|
30 |
|
|
|
31 |
locale Val_abs1 = Val_abs rep num' plus' + Rep1 rep
|
|
|
32 |
for rep :: "'a::L_top_bot \<Rightarrow> int set" and num' plus' +
|
|
|
33 |
fixes inv_plus' :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a * 'a"
|
|
|
34 |
and inv_less' :: "'a \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> 'a * 'a"
|
|
|
35 |
assumes inv_plus': "inv_plus' a1 a2 a = (a1',a2') \<Longrightarrow>
|
|
|
36 |
n1 <: a1 \<Longrightarrow> n2 <: a2 \<Longrightarrow> n1+n2 <: a \<Longrightarrow> n1 <: a1' \<and> n2 <: a2'"
|
|
|
37 |
and inv_less': "inv_less' a1 a2 (n1<n2) = (a1',a2') \<Longrightarrow>
|
|
|
38 |
n1 <: a1 \<Longrightarrow> n2 <: a2 \<Longrightarrow> n1 <: a1' \<and> n2 <: a2'"
|
|
|
39 |
|
|
|
40 |
datatype 'a up = bot | Up 'a
|
|
|
41 |
|
|
|
42 |
instantiation up :: (SL_top)SL_top
|
|
|
43 |
begin
|
|
|
44 |
|
|
|
45 |
fun le_up where
|
|
|
46 |
"Up x \<sqsubseteq> Up y = (x \<sqsubseteq> y)" |
|
|
|
47 |
"bot \<sqsubseteq> y = True" |
|
|
|
48 |
"Up _ \<sqsubseteq> bot = False"
|
|
|
49 |
|
|
|
50 |
lemma [simp]: "(x \<sqsubseteq> bot) = (x = bot)"
|
|
|
51 |
by (cases x) simp_all
|
|
|
52 |
|
|
|
53 |
lemma [simp]: "(Up x \<sqsubseteq> u) = (EX y. u = Up y & x \<sqsubseteq> y)"
|
|
|
54 |
by (cases u) auto
|
|
|
55 |
|
|
|
56 |
fun join_up where
|
|
|
57 |
"Up x \<squnion> Up y = Up(x \<squnion> y)" |
|
|
|
58 |
"bot \<squnion> y = y" |
|
|
|
59 |
"x \<squnion> bot = x"
|
|
|
60 |
|
|
|
61 |
lemma [simp]: "x \<squnion> bot = x"
|
|
|
62 |
by (cases x) simp_all
|
|
|
63 |
|
|
|
64 |
|
|
|
65 |
definition "Top = Up Top"
|
|
|
66 |
|
|
|
67 |
(* register <= as transitive - see orderings *)
|
|
|
68 |
|
|
|
69 |
instance proof
|
|
|
70 |
case goal1 show ?case by(cases x, simp_all)
|
|
|
71 |
next
|
|
|
72 |
case goal2 thus ?case
|
|
|
73 |
by(cases z, simp, cases y, simp, cases x, auto intro: le_trans)
|
|
|
74 |
next
|
|
|
75 |
case goal3 thus ?case by(cases x, simp, cases y, simp_all)
|
|
|
76 |
next
|
|
|
77 |
case goal4 thus ?case by(cases y, simp, cases x, simp_all)
|
|
|
78 |
next
|
|
|
79 |
case goal5 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all)
|
|
|
80 |
next
|
|
|
81 |
case goal6 thus ?case by(cases x, simp_all add: Top_up_def)
|
|
|
82 |
qed
|
|
|
83 |
|
|
|
84 |
end
|
|
|
85 |
|
|
|
86 |
|
|
|
87 |
locale Abs_Int1 = Val_abs1
|
|
|
88 |
begin
|
|
|
89 |
|
|
|
90 |
(* FIXME avoid duplicating this defn *)
|
|
|
91 |
abbreviation astate_in_rep (infix "<:" 50) where
|
|
|
92 |
"s <: S == ALL x. s x <: lookup S x"
|
|
|
93 |
|
|
|
94 |
abbreviation in_rep_up :: "state \<Rightarrow> 'a astate up \<Rightarrow> bool" (infix "<::" 50) where
|
|
|
95 |
"s <:: S == EX S0. S = Up S0 \<and> s <: S0"
|
|
|
96 |
|
|
|
97 |
lemma in_rep_up_trans: "(s::state) <:: S \<Longrightarrow> S \<sqsubseteq> T \<Longrightarrow> s <:: T"
|
|
|
98 |
apply auto
|
|
|
99 |
by (metis in_mono le_astate_def le_rep lookup_def top)
|
|
|
100 |
|
|
|
101 |
lemma in_rep_join_UpI: "s <:: S1 | s <:: S2 \<Longrightarrow> s <:: S1 \<squnion> S2"
|
|
|
102 |
by (metis in_rep_up_trans SL_top_class.join_ge1 SL_top_class.join_ge2)
|
|
|
103 |
|
|
|
104 |
fun aval' :: "aexp \<Rightarrow> 'a astate up \<Rightarrow> 'a" ("aval\<^isup>#") where
|
|
|
105 |
"aval' _ bot = Bot" |
|
|
|
106 |
"aval' (N n) _ = num' n" |
|
|
|
107 |
"aval' (V x) (Up S) = lookup S x" |
|
|
|
108 |
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
|
|
|
109 |
|
|
|
110 |
lemma aval'_sound: "s <:: S \<Longrightarrow> aval a s <: aval' a S"
|
|
|
111 |
by (induct a) (auto simp: rep_num' rep_plus')
|
|
|
112 |
|
|
|
113 |
fun afilter :: "aexp \<Rightarrow> 'a \<Rightarrow> 'a astate up \<Rightarrow> 'a astate up" where
|
|
|
114 |
"afilter (N n) a S = (if n <: a then S else bot)" |
|
|
|
115 |
"afilter (V x) a S = (case S of bot \<Rightarrow> bot | Up S \<Rightarrow>
|
|
|
116 |
let a' = lookup S x \<sqinter> a in
|
|
|
117 |
if a'\<sqsubseteq> Bot then bot else Up(update S x a'))" |
|
|
|
118 |
"afilter (Plus e1 e2) a S =
|
|
|
119 |
(let (a1,a2) = inv_plus' (aval' e1 S) (aval' e2 S) a
|
|
|
120 |
in afilter e1 a1 (afilter e2 a2 S))"
|
|
|
121 |
|
|
|
122 |
fun bfilter :: "bexp \<Rightarrow> bool \<Rightarrow> 'a astate up \<Rightarrow> 'a astate up" where
|
|
|
123 |
"bfilter (B bv) res S = (if bv=res then S else bot)" |
|
|
|
124 |
"bfilter (Not b) res S = bfilter b (\<not> res) S" |
|
|
|
125 |
"bfilter (And b1 b2) res S =
|
|
|
126 |
(if res then bfilter b1 True (bfilter b2 True S)
|
|
|
127 |
else bfilter b1 False S \<squnion> bfilter b2 False S)" |
|
|
|
128 |
"bfilter (Less e1 e2) res S =
|
|
|
129 |
(let (res1,res2) = inv_less' (aval' e1 S) (aval' e2 S) res
|
|
|
130 |
in afilter e1 res1 (afilter e2 res2 S))"
|
|
|
131 |
|
|
|
132 |
lemma afilter_sound: "s <:: S \<Longrightarrow> aval e s <: a \<Longrightarrow> s <:: afilter e a S"
|
|
|
133 |
proof(induct e arbitrary: a S)
|
|
|
134 |
case N thus ?case by simp
|
|
|
135 |
next
|
|
|
136 |
case (V x)
|
|
|
137 |
obtain S' where "S = Up S'" and "s <: S'" using `s <:: S` by auto
|
|
|
138 |
moreover hence "s x <: lookup S' x" by(simp)
|
|
|
139 |
moreover have "s x <: a" using V by simp
|
|
|
140 |
ultimately show ?case using V(1)
|
|
|
141 |
by(simp add: lookup_update Let_def)
|
|
|
142 |
(metis le_rep emptyE in_rep_meet rep_Bot subset_empty)
|
|
|
143 |
next
|
|
|
144 |
case (Plus e1 e2) thus ?case
|
|
|
145 |
using inv_plus'[OF _ aval'_sound[OF Plus(3)] aval'_sound[OF Plus(3)]]
|
|
|
146 |
by (auto split: prod.split)
|
|
|
147 |
qed
|
|
|
148 |
|
|
|
149 |
lemma bfilter_sound: "s <:: S \<Longrightarrow> bv = bval b s \<Longrightarrow> s <:: bfilter b bv S"
|
|
|
150 |
proof(induct b arbitrary: S bv)
|
|
|
151 |
case B thus ?case by simp
|
|
|
152 |
next
|
|
|
153 |
case (Not b) thus ?case by simp
|
|
|
154 |
next
|
|
|
155 |
case (And b1 b2) thus ?case by (auto simp: in_rep_join_UpI)
|
|
|
156 |
next
|
|
|
157 |
case (Less e1 e2) thus ?case
|
|
|
158 |
by (auto split: prod.split)
|
|
|
159 |
(metis afilter_sound inv_less' aval'_sound Less)
|
|
|
160 |
qed
|
|
|
161 |
|
|
|
162 |
fun AI :: "com \<Rightarrow> 'a astate up \<Rightarrow> 'a astate up" where
|
|
|
163 |
"AI SKIP S = S" |
|
|
|
164 |
"AI (x ::= a) S =
|
|
|
165 |
(case S of bot \<Rightarrow> bot | Up S \<Rightarrow> Up(update S x (aval' a (Up S))))" |
|
|
|
166 |
"AI (c1;c2) S = AI c2 (AI c1 S)" |
|
|
|
167 |
"AI (IF b THEN c1 ELSE c2) S =
|
|
|
168 |
AI c1 (bfilter b True S) \<squnion> AI c2 (bfilter b False S)" |
|
|
|
169 |
"AI (WHILE b DO c) S =
|
|
|
170 |
bfilter b False (pfp_above (%S. AI c (bfilter b True S)) S)"
|
|
|
171 |
|
|
|
172 |
lemma AI_sound: "(c,s) \<Rightarrow> t \<Longrightarrow> s <:: S \<Longrightarrow> t <:: AI c S"
|
|
|
173 |
proof(induct c arbitrary: s t S)
|
|
|
174 |
case SKIP thus ?case by fastsimp
|
|
|
175 |
next
|
|
|
176 |
case Assign thus ?case
|
|
|
177 |
by (auto simp: lookup_update aval'_sound)
|
|
|
178 |
next
|
|
|
179 |
case Semi thus ?case by fastsimp
|
|
|
180 |
next
|
|
|
181 |
case If thus ?case by (auto simp: in_rep_join_UpI bfilter_sound)
|
|
|
182 |
next
|
|
|
183 |
case (While b c)
|
|
|
184 |
let ?P = "pfp_above (%S. AI c (bfilter b True S)) S"
|
|
|
185 |
have pfp: "AI c (bfilter b True ?P) \<sqsubseteq> ?P" and "S \<sqsubseteq> ?P"
|
|
|
186 |
by (rule pfp_above_pfp(1), rule pfp_above_pfp(2))
|
|
|
187 |
{ fix s t
|
|
|
188 |
have "(WHILE b DO c,s) \<Rightarrow> t \<Longrightarrow> s <:: ?P \<Longrightarrow>
|
|
|
189 |
t <:: bfilter b False ?P"
|
|
|
190 |
proof(induct "WHILE b DO c" s t rule: big_step_induct)
|
|
|
191 |
case WhileFalse thus ?case by(metis bfilter_sound)
|
|
|
192 |
next
|
|
|
193 |
case WhileTrue show ?case
|
|
|
194 |
by(rule WhileTrue, rule in_rep_up_trans[OF _ pfp],
|
|
|
195 |
rule While.hyps[OF WhileTrue(2)],
|
|
|
196 |
rule bfilter_sound[OF WhileTrue.prems], simp add: WhileTrue(1))
|
|
|
197 |
qed
|
|
|
198 |
}
|
|
|
199 |
with in_rep_up_trans[OF `s <:: S` `S \<sqsubseteq> ?P`] While(2,3) AI.simps(5)
|
|
|
200 |
show ?case by simp
|
|
|
201 |
qed
|
|
|
202 |
|
|
|
203 |
text{* Exercise: *}
|
|
|
204 |
|
|
|
205 |
lemma afilter_strict: "afilter e res bot = bot"
|
|
|
206 |
by (induct e arbitrary: res) simp_all
|
|
|
207 |
|
|
|
208 |
lemma bfilter_strict: "bfilter b res bot = bot"
|
|
|
209 |
by (induct b arbitrary: res) (simp_all add: afilter_strict)
|
|
|
210 |
|
|
|
211 |
lemma pfp_strict: "f bot = bot \<Longrightarrow> pfp_above f bot = bot"
|
|
|
212 |
by(simp add: pfp_above_def pfp_def eval_nat_numeral)
|
|
|
213 |
|
|
|
214 |
lemma AI_strict: "AI c bot = bot"
|
|
|
215 |
by(induct c) (simp_all add: bfilter_strict pfp_strict)
|
|
|
216 |
|
|
|
217 |
end
|
|
|
218 |
|
|
|
219 |
end
|