author | nipkow |
Sat, 31 Dec 2011 17:53:50 +0100 | |
changeset 46063 | 81ebd0cdb300 |
parent 46039 | 698de142f6f9 |
child 46158 | 8b5f1f91ef38 |
permissions | -rw-r--r-- |
45111 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
theory Abs_Int0_const |
|
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45200
diff
changeset
|
4 |
imports Abs_Int0 Abs_Int_Tests |
45111 | 5 |
begin |
6 |
||
7 |
subsection "Constant Propagation" |
|
8 |
||
9 |
datatype cval = Const val | Any |
|
10 |
||
46039 | 11 |
fun \<gamma>_cval where |
12 |
"\<gamma>_cval (Const n) = {n}" | |
|
13 |
"\<gamma>_cval (Any) = UNIV" |
|
45111 | 14 |
|
15 |
fun plus_cval where |
|
16 |
"plus_cval (Const m) (Const n) = Const(m+n)" | |
|
17 |
"plus_cval _ _ = Any" |
|
18 |
||
19 |
lemma plus_cval_cases: "plus_cval a1 a2 = |
|
20 |
(case (a1,a2) of (Const m, Const n) \<Rightarrow> Const(m+n) | _ \<Rightarrow> Any)" |
|
21 |
by(auto split: prod.split cval.split) |
|
22 |
||
23 |
instantiation cval :: SL_top |
|
24 |
begin |
|
25 |
||
26 |
fun le_cval where |
|
27 |
"_ \<sqsubseteq> Any = True" | |
|
28 |
"Const n \<sqsubseteq> Const m = (n=m)" | |
|
29 |
"Any \<sqsubseteq> Const _ = False" |
|
30 |
||
31 |
fun join_cval where |
|
32 |
"Const m \<squnion> Const n = (if n=m then Const m else Any)" | |
|
33 |
"_ \<squnion> _ = Any" |
|
34 |
||
35 |
definition "\<top> = Any" |
|
36 |
||
37 |
instance |
|
38 |
proof |
|
39 |
case goal1 thus ?case by (cases x) simp_all |
|
40 |
next |
|
41 |
case goal2 thus ?case by(cases z, cases y, cases x, simp_all) |
|
42 |
next |
|
43 |
case goal3 thus ?case by(cases x, cases y, simp_all) |
|
44 |
next |
|
45 |
case goal4 thus ?case by(cases y, cases x, simp_all) |
|
46 |
next |
|
47 |
case goal5 thus ?case by(cases z, cases y, cases x, simp_all) |
|
48 |
next |
|
49 |
case goal6 thus ?case by(simp add: Top_cval_def) |
|
50 |
qed |
|
51 |
||
52 |
end |
|
53 |
||
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
54 |
|
46063 | 55 |
interpretation Val_abs |
56 |
where \<gamma> = \<gamma>_cval and num' = Const and plus' = plus_cval |
|
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45200
diff
changeset
|
57 |
defines aval'_const is aval' |
45111 | 58 |
proof |
59 |
case goal1 thus ?case |
|
60 |
by(cases a, cases b, simp, simp, cases b, simp, simp) |
|
61 |
next |
|
62 |
case goal2 show ?case by(simp add: Top_cval_def) |
|
63 |
next |
|
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
64 |
case goal3 show ?case by simp |
45111 | 65 |
next |
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
66 |
case goal4 thus ?case |
45111 | 67 |
by(auto simp: plus_cval_cases split: cval.split) |
68 |
qed |
|
69 |
||
46063 | 70 |
interpretation Abs_Int |
71 |
where \<gamma> = \<gamma>_cval and num' = Const and plus' = plus_cval |
|
45111 | 72 |
defines AI_const is AI |
45655
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
73 |
and step_const is step' |
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
74 |
proof qed |
45111 | 75 |
|
76 |
||
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45200
diff
changeset
|
77 |
text{* Monotonicity: *} |
45111 | 78 |
|
46063 | 79 |
interpretation Abs_Int_mono |
80 |
where \<gamma> = \<gamma>_cval and num' = Const and plus' = plus_cval |
|
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45200
diff
changeset
|
81 |
proof |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45200
diff
changeset
|
82 |
case goal1 thus ?case |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45200
diff
changeset
|
83 |
by(auto simp: plus_cval_cases split: cval.split) |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45200
diff
changeset
|
84 |
qed |
45111 | 85 |
|
86 |
||
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45200
diff
changeset
|
87 |
subsubsection "Tests" |
45111 | 88 |
|
89 |
value [code] "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test1_const))" |
|
90 |
value [code] "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test1_const))" |
|
91 |
value [code] "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test1_const))" |
|
92 |
value [code] "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test1_const))" |
|
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
93 |
value [code] "show_acom_opt (AI_const test1_const)" |
45111 | 94 |
|
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
95 |
value [code] "show_acom_opt (AI_const test2_const)" |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
96 |
value [code] "show_acom_opt (AI_const test3_const)" |
45111 | 97 |
|
98 |
value [code] "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test4_const))" |
|
99 |
value [code] "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test4_const))" |
|
100 |
value [code] "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test4_const))" |
|
101 |
value [code] "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test4_const))" |
|
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
102 |
value [code] "show_acom_opt (AI_const test4_const)" |
45111 | 103 |
|
104 |
value [code] "show_acom (((step_const \<top>)^^0) (\<bottom>\<^sub>c test5_const))" |
|
105 |
value [code] "show_acom (((step_const \<top>)^^1) (\<bottom>\<^sub>c test5_const))" |
|
106 |
value [code] "show_acom (((step_const \<top>)^^2) (\<bottom>\<^sub>c test5_const))" |
|
107 |
value [code] "show_acom (((step_const \<top>)^^3) (\<bottom>\<^sub>c test5_const))" |
|
108 |
value [code] "show_acom (((step_const \<top>)^^4) (\<bottom>\<^sub>c test5_const))" |
|
109 |
value [code] "show_acom (((step_const \<top>)^^5) (\<bottom>\<^sub>c test5_const))" |
|
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
110 |
value [code] "show_acom_opt (AI_const test5_const)" |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
111 |
|
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
112 |
value [code] "show_acom_opt (AI_const test6_const)" |
45111 | 113 |
|
114 |
end |