author | nipkow |
Mon, 02 Jan 2012 11:54:21 +0100 | |
changeset 46070 | 8392c28d7868 |
parent 46068 | b9d4ec0f79ac |
child 46153 | 7e4a18db7384 |
permissions | -rw-r--r-- |
45111 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
theory Abs_Int0 |
|
4 |
imports Abs_State |
|
5 |
begin |
|
6 |
||
7 |
subsection "Computable Abstract Interpretation" |
|
8 |
||
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
9 |
text{* Abstract interpretation over type @{text st} instead of |
45111 | 10 |
functions. *} |
11 |
||
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
12 |
context Val_abs |
45111 | 13 |
begin |
14 |
||
46063 | 15 |
fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where |
46039 | 16 |
"aval' (N n) S = num' n" | |
45111 | 17 |
"aval' (V x) S = lookup S x" | |
18 |
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" |
|
19 |
||
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
20 |
lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)" |
46039 | 21 |
by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_st_def lookup_def) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
22 |
|
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
23 |
end |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
24 |
|
46063 | 25 |
text{* The for-clause (here and elsewhere) only serves the purpose of fixing |
26 |
the name of the type parameter @{typ 'av} which would otherwise be renamed to |
|
27 |
@{typ 'a}. *} |
|
28 |
||
29 |
locale Abs_Int = Val_abs \<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set" |
|
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
30 |
begin |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
31 |
|
46063 | 32 |
fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom" where |
45655
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
33 |
"step' S (SKIP {P}) = (SKIP {S})" | |
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
34 |
"step' S (x ::= e {P}) = |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
35 |
x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S))}" | |
45655
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
36 |
"step' S (c1; c2) = step' S c1; step' (post c1) c2" | |
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
37 |
"step' S (IF b THEN c1 ELSE c2 {P}) = |
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
38 |
(let c1' = step' S c1; c2' = step' S c2 |
45111 | 39 |
in IF b THEN c1' ELSE c2' {post c1 \<squnion> post c2})" | |
45655
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
40 |
"step' S ({Inv} WHILE b DO c {P}) = |
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
41 |
{S \<squnion> post c} WHILE b DO step' Inv c {Inv}" |
45111 | 42 |
|
46063 | 43 |
definition AI :: "com \<Rightarrow> 'av st option acom option" where |
45655
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
44 |
"AI = lpfp\<^isub>c (step' \<top>)" |
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
45 |
|
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
46 |
|
45655
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
47 |
lemma strip_step'[simp]: "strip(step' S c) = strip c" |
45111 | 48 |
by(induct c arbitrary: S) (simp_all add: Let_def) |
49 |
||
50 |
||
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
51 |
text{* Soundness: *} |
45111 | 52 |
|
46039 | 53 |
lemma in_gamma_update: |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
54 |
"\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(update S x a)" |
46039 | 55 |
by(simp add: \<gamma>_st_def lookup_update) |
45111 | 56 |
|
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
57 |
text{* The soundness proofs are textually identical to the ones for the step |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
58 |
function operating on states as functions. *} |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
59 |
|
46068 | 60 |
lemma step_preserves_le: |
61 |
"\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; cs \<le> \<gamma>\<^isub>c ca \<rbrakk> \<Longrightarrow> step S cs \<le> \<gamma>\<^isub>c (step' S' ca)" |
|
62 |
proof(induction cs arbitrary: ca S S') |
|
63 |
case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) |
|
45111 | 64 |
next |
65 |
case Assign thus ?case |
|
46068 | 66 |
by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
67 |
split: option.splits del:subsetD) |
45111 | 68 |
next |
46068 | 69 |
case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
70 |
by (metis le_post post_map_acom) |
45111 | 71 |
next |
46068 | 72 |
case (If b cs1 cs2 P) |
73 |
then obtain ca1 ca2 Pa where |
|
74 |
"ca= IF b THEN ca1 ELSE ca2 {Pa}" |
|
46039 | 75 |
"P \<subseteq> \<gamma>\<^isub>o Pa" "cs1 \<le> \<gamma>\<^isub>c ca1" "cs2 \<le> \<gamma>\<^isub>c ca2" |
46068 | 76 |
by (fastforce simp: If_le map_acom_If) |
46039 | 77 |
moreover have "post cs1 \<subseteq> \<gamma>\<^isub>o(post ca1 \<squnion> post ca2)" |
78 |
by (metis (no_types) `cs1 \<le> \<gamma>\<^isub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom) |
|
79 |
moreover have "post cs2 \<subseteq> \<gamma>\<^isub>o(post ca1 \<squnion> post ca2)" |
|
80 |
by (metis (no_types) `cs2 \<le> \<gamma>\<^isub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom) |
|
46068 | 81 |
ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` by (simp add: If.IH subset_iff) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
82 |
next |
46068 | 83 |
case (While I b cs1 P) |
84 |
then obtain ca1 Ia Pa where |
|
85 |
"ca = {Ia} WHILE b DO ca1 {Pa}" |
|
46039 | 86 |
"I \<subseteq> \<gamma>\<^isub>o Ia" "P \<subseteq> \<gamma>\<^isub>o Pa" "cs1 \<le> \<gamma>\<^isub>c ca1" |
46068 | 87 |
by (fastforce simp: map_acom_While While_le) |
46067 | 88 |
moreover have "S \<union> post cs1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post ca1)" |
89 |
using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `cs1 \<le> \<gamma>\<^isub>c ca1`, simplified] |
|
46039 | 90 |
by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
91 |
ultimately show ?case by (simp add: While.IH subset_iff) |
45111 | 92 |
qed |
93 |
||
46070 | 94 |
lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'" |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
95 |
proof(simp add: CS_def AI_def) |
45655
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
96 |
assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'" |
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
97 |
have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1]) |
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
98 |
have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c" |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
99 |
by(simp add: strip_lpfpc[OF _ 1]) |
46066 | 100 |
have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')" |
45903 | 101 |
proof(rule lfp_lowerbound[simplified,OF 3]) |
45655
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
102 |
show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')" |
46068 | 103 |
proof(rule step_preserves_le[OF _ _]) |
46039 | 104 |
show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp |
105 |
show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2]) |
|
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
106 |
qed |
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
107 |
qed |
46066 | 108 |
from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'" |
46039 | 109 |
by (blast intro: mono_gamma_c order_trans) |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
110 |
qed |
45111 | 111 |
|
112 |
end |
|
113 |
||
114 |
||
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
115 |
subsubsection "Monotonicity" |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
116 |
|
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
117 |
locale Abs_Int_mono = Abs_Int + |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
118 |
assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2" |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
119 |
begin |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
120 |
|
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
121 |
lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'" |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
122 |
by(induction e) (auto simp: le_st_def lookup_def mono_plus') |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
123 |
|
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
124 |
lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> update S x a \<sqsubseteq> update S' x a'" |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
125 |
by(auto simp add: le_st_def lookup_def update_def) |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
126 |
|
45655
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents:
45623
diff
changeset
|
127 |
lemma step'_mono: "S \<sqsubseteq> S' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c" |
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
128 |
apply(induction c arbitrary: S S') |
45623
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
45127
diff
changeset
|
129 |
apply (auto simp: Let_def mono_update mono_aval' le_join_disj split: option.split) |
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
130 |
done |
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
131 |
|
45111 | 132 |
end |
45127
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
133 |
|
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents:
45111
diff
changeset
|
134 |
end |