| author | wenzelm | 
| Fri, 08 Apr 2016 20:52:40 +0200 | |
| changeset 62914 | 930a30c1a9af | 
| parent 58252 | 7e353ced895e | 
| permissions | -rw-r--r-- | 
| 45111 | 1  | 
(* Author: Tobias Nipkow *)  | 
2  | 
||
| 47602 | 3  | 
theory Abs_Int2_ITP  | 
| 48480 | 4  | 
imports Abs_Int1_ITP "../Vars"  | 
| 45111 | 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"  | 
|
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
33  | 
begin  | 
| 45111 | 34  | 
|
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
35  | 
lemma mono_meet: "x \<sqsubseteq> x' \<Longrightarrow> y \<sqsubseteq> y' \<Longrightarrow> x \<sqinter> y \<sqsubseteq> x' \<sqinter> y'"  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
36  | 
by (metis meet_greatest meet_le1 meet_le2 le_trans)  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
37  | 
|
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
38  | 
end  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
39  | 
|
| 46039 | 40  | 
locale Val_abs1_gamma =  | 
| 
46346
 
10c18630612a
removed duplicate definitions that made locale inconsistent
 
nipkow 
parents: 
46251 
diff
changeset
 | 
41  | 
Gamma where \<gamma> = \<gamma> for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set" +  | 
| 46039 | 42  | 
assumes inter_gamma_subset_gamma_meet:  | 
43  | 
"\<gamma> a1 \<inter> \<gamma> a2 \<subseteq> \<gamma>(a1 \<sqinter> a2)"  | 
|
44  | 
and gamma_Bot[simp]: "\<gamma> \<bottom> = {}"
 | 
|
| 45111 | 45  | 
begin  | 
46  | 
||
| 46039 | 47  | 
lemma in_gamma_meet: "x : \<gamma> a1 \<Longrightarrow> x : \<gamma> a2 \<Longrightarrow> x : \<gamma>(a1 \<sqinter> a2)"  | 
48  | 
by (metis IntI inter_gamma_subset_gamma_meet set_mp)  | 
|
| 45111 | 49  | 
|
| 46039 | 50  | 
lemma gamma_meet[simp]: "\<gamma>(a1 \<sqinter> a2) = \<gamma> a1 \<inter> \<gamma> a2"  | 
51  | 
by (metis equalityI inter_gamma_subset_gamma_meet le_inf_iff mono_gamma meet_le1 meet_le2)  | 
|
| 45111 | 52  | 
|
53  | 
end  | 
|
54  | 
||
55  | 
||
| 46063 | 56  | 
locale Val_abs1 =  | 
| 
46355
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
57  | 
Val_abs1_gamma where \<gamma> = \<gamma>  | 
| 
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
58  | 
for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set" +  | 
| 
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
59  | 
fixes test_num' :: "val \<Rightarrow> 'av \<Rightarrow> bool"  | 
| 
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
60  | 
and filter_plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av \<Rightarrow> 'av * 'av"  | 
| 46063 | 61  | 
and filter_less' :: "bool \<Rightarrow> 'av \<Rightarrow> 'av \<Rightarrow> 'av * 'av"  | 
| 
46355
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
62  | 
assumes test_num': "test_num' n a = (n : \<gamma> a)"  | 
| 
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
63  | 
and filter_plus': "filter_plus' a a1 a2 = (b1,b2) \<Longrightarrow>  | 
| 
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
64  | 
n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1+n2 : \<gamma> a \<Longrightarrow> n1 : \<gamma> b1 \<and> n2 : \<gamma> b2"  | 
| 
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
65  | 
and filter_less': "filter_less' (n1<n2) a1 a2 = (b1,b2) \<Longrightarrow>  | 
| 
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
66  | 
n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1 : \<gamma> b1 \<and> n2 : \<gamma> b2"  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
67  | 
|
| 45111 | 68  | 
|
| 46063 | 69  | 
locale Abs_Int1 =  | 
70  | 
Val_abs1 where \<gamma> = \<gamma> for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set"  | 
|
| 45111 | 71  | 
begin  | 
72  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
73  | 
lemma in_gamma_join_UpI: "s : \<gamma>\<^sub>o S1 \<or> s : \<gamma>\<^sub>o S2 \<Longrightarrow> s : \<gamma>\<^sub>o(S1 \<squnion> S2)"  | 
| 46039 | 74  | 
by (metis (no_types) join_ge1 join_ge2 mono_gamma_o set_rev_mp)  | 
| 45111 | 75  | 
|
| 46063 | 76  | 
fun aval'' :: "aexp \<Rightarrow> 'av st option \<Rightarrow> 'av" where  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
77  | 
"aval'' e None = \<bottom>" |  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
78  | 
"aval'' e (Some sa) = aval' e sa"  | 
| 45111 | 79  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
80  | 
lemma aval''_sound: "s : \<gamma>\<^sub>o S \<Longrightarrow> aval a s : \<gamma>(aval'' a S)"  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
81  | 
by(cases S)(simp add: aval'_sound)+  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
82  | 
|
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
83  | 
subsubsection "Backward analysis"  | 
| 45111 | 84  | 
|
| 46063 | 85  | 
fun afilter :: "aexp \<Rightarrow> 'av \<Rightarrow> 'av st option \<Rightarrow> 'av st option" where  | 
| 
46355
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
86  | 
"afilter (N n) a S = (if test_num' n a then S else None)" |  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
87  | 
"afilter (V x) a S = (case S of None \<Rightarrow> None | Some S \<Rightarrow>  | 
| 45111 | 88  | 
let a' = lookup S x \<sqinter> a in  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
89  | 
if a' \<sqsubseteq> \<bottom> then None else Some(update S x a'))" |  | 
| 45111 | 90  | 
"afilter (Plus e1 e2) a S =  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
91  | 
(let (a1,a2) = filter_plus' a (aval'' e1 S) (aval'' e2 S)  | 
| 45111 | 92  | 
in afilter e1 a1 (afilter e2 a2 S))"  | 
93  | 
||
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
94  | 
text{* The test for @{const bot} in the @{const V}-case is important: @{const
 | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
95  | 
bot} indicates that a variable has no possible values, i.e.\ that the current  | 
| 45111 | 96  | 
program point is unreachable. But then the abstract state should collapse to  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
97  | 
@{const None}. Put differently, we maintain the invariant that in an abstract
 | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
98  | 
state of the form @{term"Some s"}, all variables are mapped to non-@{const
 | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
99  | 
bot} values. Otherwise the (pointwise) join of two abstract states, one of  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
100  | 
which contains @{const bot} values, may produce too large a result, thus
 | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
101  | 
making the analysis less precise. *}  | 
| 45111 | 102  | 
|
103  | 
||
| 46063 | 104  | 
fun bfilter :: "bexp \<Rightarrow> bool \<Rightarrow> 'av st option \<Rightarrow> 'av st option" where  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
105  | 
"bfilter (Bc v) res S = (if v=res then S else None)" |  | 
| 45111 | 106  | 
"bfilter (Not b) res S = bfilter b (\<not> res) S" |  | 
107  | 
"bfilter (And b1 b2) res S =  | 
|
108  | 
(if res then bfilter b1 True (bfilter b2 True S)  | 
|
109  | 
else bfilter b1 False S \<squnion> bfilter b2 False S)" |  | 
|
110  | 
"bfilter (Less e1 e2) res S =  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
111  | 
(let (res1,res2) = filter_less' res (aval'' e1 S) (aval'' e2 S)  | 
| 45111 | 112  | 
in afilter e1 res1 (afilter e2 res2 S))"  | 
113  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
114  | 
lemma afilter_sound: "s : \<gamma>\<^sub>o S \<Longrightarrow> aval e s : \<gamma> a \<Longrightarrow> s : \<gamma>\<^sub>o (afilter e a S)"  | 
| 45111 | 115  | 
proof(induction e arbitrary: a S)  | 
| 
46355
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
116  | 
case N thus ?case by simp (metis test_num')  | 
| 45111 | 117  | 
next  | 
118  | 
case (V x)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
119  | 
obtain S' where "S = Some S'" and "s : \<gamma>\<^sub>f S'" using `s : \<gamma>\<^sub>o S`  | 
| 46039 | 120  | 
by(auto simp: in_gamma_option_iff)  | 
121  | 
moreover hence "s x : \<gamma> (lookup S' x)" by(simp add: \<gamma>_st_def)  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
122  | 
moreover have "s x : \<gamma> a" using V by simp  | 
| 45111 | 123  | 
ultimately show ?case using V(1)  | 
| 46039 | 124  | 
by(simp add: lookup_update Let_def \<gamma>_st_def)  | 
125  | 
(metis mono_gamma emptyE in_gamma_meet gamma_Bot subset_empty)  | 
|
| 45111 | 126  | 
next  | 
127  | 
case (Plus e1 e2) thus ?case  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
128  | 
using filter_plus'[OF _ aval''_sound[OF Plus(3)] aval''_sound[OF Plus(3)]]  | 
| 45111 | 129  | 
by (auto split: prod.split)  | 
130  | 
qed  | 
|
131  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
132  | 
lemma bfilter_sound: "s : \<gamma>\<^sub>o S \<Longrightarrow> bv = bval b s \<Longrightarrow> s : \<gamma>\<^sub>o(bfilter b bv S)"  | 
| 45111 | 133  | 
proof(induction b arbitrary: S bv)  | 
| 45200 | 134  | 
case Bc thus ?case by simp  | 
| 45111 | 135  | 
next  | 
136  | 
case (Not b) thus ?case by simp  | 
|
137  | 
next  | 
|
| 
57492
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
53015 
diff
changeset
 | 
138  | 
case (And b1 b2) thus ?case  | 
| 
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
53015 
diff
changeset
 | 
139  | 
apply hypsubst_thin  | 
| 
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
53015 
diff
changeset
 | 
140  | 
apply (fastforce simp: in_gamma_join_UpI)  | 
| 
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
53015 
diff
changeset
 | 
141  | 
done  | 
| 45111 | 142  | 
next  | 
143  | 
case (Less e1 e2) thus ?case  | 
|
| 
57492
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
53015 
diff
changeset
 | 
144  | 
apply hypsubst_thin  | 
| 
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
53015 
diff
changeset
 | 
145  | 
apply (auto split: prod.split)  | 
| 
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
53015 
diff
changeset
 | 
146  | 
apply (metis afilter_sound filter_less' aval''_sound Less)  | 
| 
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
53015 
diff
changeset
 | 
147  | 
done  | 
| 45111 | 148  | 
qed  | 
149  | 
||
150  | 
||
| 46063 | 151  | 
fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom"  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
152  | 
where  | 
| 
45655
 
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
 
nipkow 
parents: 
45623 
diff
changeset
 | 
153  | 
"step' S (SKIP {P}) = (SKIP {S})" |
 | 
| 
 
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
 
nipkow 
parents: 
45623 
diff
changeset
 | 
154  | 
"step' S (x ::= e {P}) =
 | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
155  | 
  x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S))}" |
 | 
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
51698 
diff
changeset
 | 
156  | 
"step' S (c1;; c2) = step' S c1;; step' (post c1) c2" |  | 
| 
45655
 
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
 
nipkow 
parents: 
45623 
diff
changeset
 | 
157  | 
"step' S (IF b THEN c1 ELSE c2 {P}) =
 | 
| 
 
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
 
nipkow 
parents: 
45623 
diff
changeset
 | 
158  | 
(let c1' = step' (bfilter b True S) c1; c2' = step' (bfilter b False S) c2  | 
| 45111 | 159  | 
   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
 | 
160  | 
"step' S ({Inv} WHILE b DO c {P}) =
 | 
| 45111 | 161  | 
   {S \<squnion> post c}
 | 
| 
45655
 
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
 
nipkow 
parents: 
45623 
diff
changeset
 | 
162  | 
WHILE b DO step' (bfilter b True Inv) c  | 
| 45111 | 163  | 
   {bfilter b False Inv}"
 | 
164  | 
||
| 46063 | 165  | 
definition AI :: "com \<Rightarrow> 'av st option acom option" where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
166  | 
"AI = lpfp\<^sub>c (step' \<top>)"  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
167  | 
|
| 
45655
 
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
 
nipkow 
parents: 
45623 
diff
changeset
 | 
168  | 
lemma strip_step'[simp]: "strip(step' S c) = strip c"  | 
| 45111 | 169  | 
by(induct c arbitrary: S) (simp_all add: Let_def)  | 
170  | 
||
171  | 
||
172  | 
subsubsection "Soundness"  | 
|
173  | 
||
| 46039 | 174  | 
lemma in_gamma_update:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
175  | 
"\<lbrakk> s : \<gamma>\<^sub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^sub>f(update S x a)"  | 
| 46039 | 176  | 
by(simp add: \<gamma>_st_def lookup_update)  | 
| 45111 | 177  | 
|
| 46068 | 178  | 
lemma step_preserves_le:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
179  | 
"\<lbrakk> S \<subseteq> \<gamma>\<^sub>o S'; cs \<le> \<gamma>\<^sub>c ca \<rbrakk> \<Longrightarrow> step S cs \<le> \<gamma>\<^sub>c (step' S' ca)"  | 
| 46068 | 180  | 
proof(induction cs arbitrary: ca S S')  | 
181  | 
case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP)  | 
|
| 45111 | 182  | 
next  | 
183  | 
case Assign thus ?case  | 
|
| 46068 | 184  | 
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: 
45200 
diff
changeset
 | 
185  | 
split: option.splits del:subsetD)  | 
| 45111 | 186  | 
next  | 
| 47818 | 187  | 
case Seq thus ?case apply (auto simp: Seq_le map_acom_Seq)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
188  | 
by (metis le_post post_map_acom)  | 
| 45111 | 189  | 
next  | 
| 46068 | 190  | 
case (If b cs1 cs2 P)  | 
191  | 
then obtain ca1 ca2 Pa where  | 
|
192  | 
      "ca= IF b THEN ca1 ELSE ca2 {Pa}"
 | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
193  | 
"P \<subseteq> \<gamma>\<^sub>o Pa" "cs1 \<le> \<gamma>\<^sub>c ca1" "cs2 \<le> \<gamma>\<^sub>c ca2"  | 
| 46068 | 194  | 
by (fastforce simp: If_le map_acom_If)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
195  | 
moreover have "post cs1 \<subseteq> \<gamma>\<^sub>o(post ca1 \<squnion> post ca2)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
196  | 
by (metis (no_types) `cs1 \<le> \<gamma>\<^sub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
197  | 
moreover have "post cs2 \<subseteq> \<gamma>\<^sub>o(post ca1 \<squnion> post ca2)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
198  | 
by (metis (no_types) `cs2 \<le> \<gamma>\<^sub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
199  | 
ultimately show ?case using `S \<subseteq> \<gamma>\<^sub>o S'`  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
200  | 
by (simp add: If.IH subset_iff bfilter_sound)  | 
| 45111 | 201  | 
next  | 
| 46068 | 202  | 
case (While I b cs1 P)  | 
203  | 
then obtain ca1 Ia Pa where  | 
|
204  | 
    "ca = {Ia} WHILE b DO ca1 {Pa}"
 | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
205  | 
"I \<subseteq> \<gamma>\<^sub>o Ia" "P \<subseteq> \<gamma>\<^sub>o Pa" "cs1 \<le> \<gamma>\<^sub>c ca1"  | 
| 46068 | 206  | 
by (fastforce simp: map_acom_While While_le)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
207  | 
moreover have "S \<union> post cs1 \<subseteq> \<gamma>\<^sub>o (S' \<squnion> post ca1)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
208  | 
using `S \<subseteq> \<gamma>\<^sub>o S'` le_post[OF `cs1 \<le> \<gamma>\<^sub>c ca1`, simplified]  | 
| 46039 | 209  | 
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: 
45200 
diff
changeset
 | 
210  | 
ultimately show ?case by (simp add: While.IH subset_iff bfilter_sound)  | 
| 45111 | 211  | 
qed  | 
212  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
213  | 
lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c c'"  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
214  | 
proof(simp add: CS_def AI_def)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
215  | 
assume 1: "lpfp\<^sub>c (step' \<top>) c = Some c'"  | 
| 
45655
 
a49f9428aba4
simplified Collecting1 and renamed: step -> step', step_cs -> step
 
nipkow 
parents: 
45623 
diff
changeset
 | 
216  | 
have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
217  | 
have 3: "strip (\<gamma>\<^sub>c (step' \<top> c')) = c"  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
218  | 
by(simp add: strip_lpfpc[OF _ 1])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
219  | 
have "lfp (step UNIV) c \<le> \<gamma>\<^sub>c (step' \<top> c')"  | 
| 45903 | 220  | 
proof(rule lfp_lowerbound[simplified,OF 3])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
221  | 
show "step UNIV (\<gamma>\<^sub>c (step' \<top> c')) \<le> \<gamma>\<^sub>c (step' \<top> c')"  | 
| 46068 | 222  | 
proof(rule step_preserves_le[OF _ _])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
223  | 
show "UNIV \<subseteq> \<gamma>\<^sub>o \<top>" by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
224  | 
show "\<gamma>\<^sub>c (step' \<top> c') \<le> \<gamma>\<^sub>c c'" by(rule mono_gamma_c[OF 2])  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
225  | 
qed  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
226  | 
qed  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52046 
diff
changeset
 | 
227  | 
from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^sub>c c'"  | 
| 46039 | 228  | 
by (blast intro: mono_gamma_c order_trans)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
229  | 
qed  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
230  | 
|
| 46246 | 231  | 
|
232  | 
subsubsection "Commands over a set of variables"  | 
|
233  | 
||
234  | 
text{* Key invariant: the domains of all abstract states are subsets of the
 | 
|
235  | 
set of variables of the program. *}  | 
|
236  | 
||
237  | 
definition "domo S = (case S of None \<Rightarrow> {} | Some S' \<Rightarrow> set(dom S'))"
 | 
|
238  | 
||
239  | 
definition Com :: "vname set \<Rightarrow> 'a st option acom set" where  | 
|
240  | 
"Com X = {c. \<forall>S \<in> set(annos c). domo S \<subseteq> X}"
 | 
|
241  | 
||
242  | 
lemma domo_Top[simp]: "domo \<top> = {}"
 | 
|
243  | 
by(simp add: domo_def Top_st_def Top_option_def)  | 
|
244  | 
||
| 46251 | 245  | 
lemma bot_acom_Com[simp]: "\<bottom>\<^sub>c c \<in> Com X"  | 
| 46246 | 246  | 
by(simp add: bot_acom_def Com_def domo_def set_annos_anno)  | 
247  | 
||
248  | 
lemma post_in_annos: "post c : set(annos c)"  | 
|
249  | 
by(induction c) simp_all  | 
|
250  | 
||
251  | 
lemma domo_join: "domo (S \<squnion> T) \<subseteq> domo S \<union> domo T"  | 
|
252  | 
by(auto simp: domo_def join_st_def split: option.split)  | 
|
253  | 
||
254  | 
lemma domo_afilter: "vars a \<subseteq> X \<Longrightarrow> domo S \<subseteq> X \<Longrightarrow> domo(afilter a i S) \<subseteq> X"  | 
|
255  | 
apply(induction a arbitrary: i S)  | 
|
256  | 
apply(simp add: domo_def)  | 
|
257  | 
apply(simp add: domo_def Let_def update_def lookup_def split: option.splits)  | 
|
258  | 
apply blast  | 
|
259  | 
apply(simp split: prod.split)  | 
|
260  | 
done  | 
|
261  | 
||
262  | 
lemma domo_bfilter: "vars b \<subseteq> X \<Longrightarrow> domo S \<subseteq> X \<Longrightarrow> domo(bfilter b bv S) \<subseteq> X"  | 
|
263  | 
apply(induction b arbitrary: bv S)  | 
|
264  | 
apply(simp add: domo_def)  | 
|
265  | 
apply(simp)  | 
|
266  | 
apply(simp)  | 
|
267  | 
apply rule  | 
|
268  | 
apply (metis le_sup_iff subset_trans[OF domo_join])  | 
|
269  | 
apply(simp split: prod.split)  | 
|
270  | 
by (metis domo_afilter)  | 
|
271  | 
||
272  | 
lemma step'_Com:  | 
|
273  | 
"domo S \<subseteq> X \<Longrightarrow> vars(strip c) \<subseteq> X \<Longrightarrow> c : Com X \<Longrightarrow> step' S c : Com X"  | 
|
274  | 
apply(induction c arbitrary: S)  | 
|
275  | 
apply(simp add: Com_def)  | 
|
276  | 
apply(simp add: Com_def domo_def update_def split: option.splits)  | 
|
277  | 
apply(simp (no_asm_use) add: Com_def ball_Un)  | 
|
278  | 
apply (metis post_in_annos)  | 
|
279  | 
apply(simp (no_asm_use) add: Com_def ball_Un)  | 
|
280  | 
apply rule  | 
|
281  | 
apply (metis Un_assoc domo_join order_trans post_in_annos subset_Un_eq)  | 
|
282  | 
apply (metis domo_bfilter)  | 
|
283  | 
apply(simp (no_asm_use) add: Com_def)  | 
|
284  | 
apply rule  | 
|
285  | 
apply (metis domo_join le_sup_iff post_in_annos subset_trans)  | 
|
286  | 
apply rule  | 
|
287  | 
apply (metis domo_bfilter)  | 
|
288  | 
by (metis domo_bfilter)  | 
|
289  | 
||
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
290  | 
end  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
291  | 
|
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
292  | 
|
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
293  | 
subsubsection "Monotonicity"  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
294  | 
|
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
295  | 
locale Abs_Int1_mono = Abs_Int1 +  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
296  | 
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
 | 
297  | 
and mono_filter_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> r \<sqsubseteq> r' \<Longrightarrow>  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
298  | 
filter_plus' r a1 a2 \<sqsubseteq> filter_plus' r' b1 b2"  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
299  | 
and mono_filter_less': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow>  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
300  | 
filter_less' bv a1 a2 \<sqsubseteq> filter_less' bv b1 b2"  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
301  | 
begin  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
302  | 
|
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
303  | 
lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'"  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
304  | 
by(induction e) (auto simp: le_st_def lookup_def mono_plus')  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
305  | 
|
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
306  | 
lemma mono_aval'': "S \<sqsubseteq> S' \<Longrightarrow> aval'' e S \<sqsubseteq> aval'' e S'"  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
307  | 
apply(cases S)  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
308  | 
apply simp  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
309  | 
apply(cases S')  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
310  | 
apply simp  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
311  | 
by (simp add: mono_aval')  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
312  | 
|
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
313  | 
lemma mono_afilter: "r \<sqsubseteq> r' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> afilter e r S \<sqsubseteq> afilter e r' S'"  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
314  | 
apply(induction e arbitrary: r r' S S')  | 
| 
46355
 
42a01315d998
removed accidental dependance of abstract interpreter on gamma
 
nipkow 
parents: 
46346 
diff
changeset
 | 
315  | 
apply(auto simp: test_num' Let_def split: option.splits prod.splits)  | 
| 46039 | 316  | 
apply (metis mono_gamma subsetD)  | 
| 58252 | 317  | 
apply(rename_tac list a b c d, drule_tac x = "list" in mono_lookup)  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
318  | 
apply (metis mono_meet le_trans)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
319  | 
apply (metis mono_meet mono_lookup mono_update)  | 
| 
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
320  | 
apply(metis mono_aval'' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv)  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
321  | 
done  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
322  | 
|
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
323  | 
lemma mono_bfilter: "S \<sqsubseteq> S' \<Longrightarrow> bfilter b r S \<sqsubseteq> bfilter b r S'"  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
324  | 
apply(induction b arbitrary: r S S')  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
325  | 
apply(auto simp: le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] split: prod.splits)  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
326  | 
apply(metis mono_aval'' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv)  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
327  | 
done  | 
| 
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
328  | 
|
| 46153 | 329  | 
lemma mono_step': "S \<sqsubseteq> S' \<Longrightarrow> c \<sqsubseteq> c' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c'"  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
330  | 
apply(induction c c' arbitrary: S S' rule: le_acom.induct)  | 
| 46153 | 331  | 
apply (auto simp: mono_post mono_bfilter mono_update mono_aval' Let_def le_join_disj  | 
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
332  | 
split: option.split)  | 
| 
45127
 
d2eb07a1e01b
separated monotonicity reasoning and defined narrowing with while_option
 
nipkow 
parents: 
45111 
diff
changeset
 | 
333  | 
done  | 
| 45111 | 334  | 
|
| 46153 | 335  | 
lemma mono_step'2: "mono (step' S)"  | 
336  | 
by(simp add: mono_def mono_step'[OF le_refl])  | 
|
| 
45623
 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 
nipkow 
parents: 
45200 
diff
changeset
 | 
337  | 
|
| 45111 | 338  | 
end  | 
339  | 
||
340  | 
end  |