| author | blanchet | 
| Tue, 28 Aug 2012 17:17:41 +0200 | |
| changeset 48978 | dcb486124b6a | 
| parent 48759 | ff570720ba1c | 
| child 49095 | 7df19036392e | 
| permissions | -rw-r--r-- | 
| 47613 | 1  | 
(* Author: Tobias Nipkow *)  | 
2  | 
||
3  | 
theory Abs_Int2  | 
|
4  | 
imports Abs_Int1  | 
|
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  | 
class L_top_bot = SL_top + Bot +  | 
|
25  | 
fixes meet :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 65)  | 
|
26  | 
assumes meet_le1 [simp]: "x \<sqinter> y \<sqsubseteq> x"  | 
|
27  | 
and meet_le2 [simp]: "x \<sqinter> y \<sqsubseteq> y"  | 
|
28  | 
and meet_greatest: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<sqinter> z"  | 
|
29  | 
begin  | 
|
30  | 
||
31  | 
lemma mono_meet: "x \<sqsubseteq> x' \<Longrightarrow> y \<sqsubseteq> y' \<Longrightarrow> x \<sqinter> y \<sqsubseteq> x' \<sqinter> y'"  | 
|
32  | 
by (metis meet_greatest meet_le1 meet_le2 le_trans)  | 
|
33  | 
||
34  | 
end  | 
|
35  | 
||
36  | 
locale Val_abs1_gamma =  | 
|
37  | 
Gamma where \<gamma> = \<gamma> for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set" +  | 
|
38  | 
assumes inter_gamma_subset_gamma_meet:  | 
|
39  | 
"\<gamma> a1 \<inter> \<gamma> a2 \<subseteq> \<gamma>(a1 \<sqinter> a2)"  | 
|
40  | 
and gamma_Bot[simp]: "\<gamma> \<bottom> = {}"
 | 
|
41  | 
begin  | 
|
42  | 
||
43  | 
lemma in_gamma_meet: "x : \<gamma> a1 \<Longrightarrow> x : \<gamma> a2 \<Longrightarrow> x : \<gamma>(a1 \<sqinter> a2)"  | 
|
44  | 
by (metis IntI inter_gamma_subset_gamma_meet set_mp)  | 
|
45  | 
||
46  | 
lemma gamma_meet[simp]: "\<gamma>(a1 \<sqinter> a2) = \<gamma> a1 \<inter> \<gamma> a2"  | 
|
47  | 
by (metis equalityI inter_gamma_subset_gamma_meet le_inf_iff mono_gamma meet_le1 meet_le2)  | 
|
48  | 
||
49  | 
end  | 
|
50  | 
||
51  | 
||
52  | 
locale Val_abs1 =  | 
|
53  | 
Val_abs1_gamma where \<gamma> = \<gamma>  | 
|
54  | 
for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set" +  | 
|
55  | 
fixes test_num' :: "val \<Rightarrow> 'av \<Rightarrow> bool"  | 
|
56  | 
and filter_plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av \<Rightarrow> 'av * 'av"  | 
|
57  | 
and filter_less' :: "bool \<Rightarrow> 'av \<Rightarrow> 'av \<Rightarrow> 'av * 'av"  | 
|
58  | 
assumes test_num': "test_num' n a = (n : \<gamma> a)"  | 
|
59  | 
and filter_plus': "filter_plus' a a1 a2 = (b1,b2) \<Longrightarrow>  | 
|
60  | 
n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1+n2 : \<gamma> a \<Longrightarrow> n1 : \<gamma> b1 \<and> n2 : \<gamma> b2"  | 
|
61  | 
and filter_less': "filter_less' (n1<n2) a1 a2 = (b1,b2) \<Longrightarrow>  | 
|
62  | 
n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1 : \<gamma> b1 \<and> n2 : \<gamma> b2"  | 
|
63  | 
||
64  | 
||
65  | 
locale Abs_Int1 =  | 
|
66  | 
Val_abs1 where \<gamma> = \<gamma> for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set"  | 
|
67  | 
begin  | 
|
68  | 
||
69  | 
lemma in_gamma_join_UpI:  | 
|
70  | 
"wt S1 X \<Longrightarrow> wt S2 X \<Longrightarrow> s : \<gamma>\<^isub>o S1 \<or> s : \<gamma>\<^isub>o S2 \<Longrightarrow> s : \<gamma>\<^isub>o(S1 \<squnion> S2)"  | 
|
71  | 
by (metis (hide_lams, no_types) SL_top_wt_class.join_ge1 SL_top_wt_class.join_ge2 mono_gamma_o subsetD)  | 
|
72  | 
||
73  | 
fun aval'' :: "aexp \<Rightarrow> 'av st option \<Rightarrow> 'av" where  | 
|
74  | 
"aval'' e None = \<bottom>" |  | 
|
75  | 
"aval'' e (Some sa) = aval' e sa"  | 
|
76  | 
||
77  | 
lemma aval''_sound: "s : \<gamma>\<^isub>o S \<Longrightarrow> wt S X \<Longrightarrow> vars a \<subseteq> X \<Longrightarrow> aval a s : \<gamma>(aval'' a S)"  | 
|
78  | 
by(simp add: wt_option_def wt_st_def aval'_sound split: option.splits)  | 
|
79  | 
||
80  | 
subsubsection "Backward analysis"  | 
|
81  | 
||
82  | 
fun afilter :: "aexp \<Rightarrow> 'av \<Rightarrow> 'av st option \<Rightarrow> 'av st option" where  | 
|
83  | 
"afilter (N n) a S = (if test_num' n a then S else None)" |  | 
|
84  | 
"afilter (V x) a S = (case S of None \<Rightarrow> None | Some S \<Rightarrow>  | 
|
85  | 
let a' = fun S x \<sqinter> a in  | 
|
86  | 
if a' \<sqsubseteq> \<bottom> then None else Some(update S x a'))" |  | 
|
87  | 
"afilter (Plus e1 e2) a S =  | 
|
88  | 
(let (a1,a2) = filter_plus' a (aval'' e1 S) (aval'' e2 S)  | 
|
89  | 
in afilter e1 a1 (afilter e2 a2 S))"  | 
|
90  | 
||
91  | 
text{* The test for @{const bot} in the @{const V}-case is important: @{const
 | 
|
92  | 
bot} indicates that a variable has no possible values, i.e.\ that the current  | 
|
93  | 
program point is unreachable. But then the abstract state should collapse to  | 
|
94  | 
@{const None}. Put differently, we maintain the invariant that in an abstract
 | 
|
95  | 
state of the form @{term"Some s"}, all variables are mapped to non-@{const
 | 
|
96  | 
bot} values. Otherwise the (pointwise) join of two abstract states, one of  | 
|
97  | 
which contains @{const bot} values, may produce too large a result, thus
 | 
|
98  | 
making the analysis less precise. *}  | 
|
99  | 
||
100  | 
||
101  | 
fun bfilter :: "bexp \<Rightarrow> bool \<Rightarrow> 'av st option \<Rightarrow> 'av st option" where  | 
|
102  | 
"bfilter (Bc v) res S = (if v=res then S else None)" |  | 
|
103  | 
"bfilter (Not b) res S = bfilter b (\<not> res) S" |  | 
|
104  | 
"bfilter (And b1 b2) res S =  | 
|
105  | 
(if res then bfilter b1 True (bfilter b2 True S)  | 
|
106  | 
else bfilter b1 False S \<squnion> bfilter b2 False S)" |  | 
|
107  | 
"bfilter (Less e1 e2) res S =  | 
|
108  | 
(let (res1,res2) = filter_less' res (aval'' e1 S) (aval'' e2 S)  | 
|
109  | 
in afilter e1 res1 (afilter e2 res2 S))"  | 
|
110  | 
||
111  | 
lemma wt_afilter: "wt S X \<Longrightarrow> vars e \<subseteq> X ==> wt (afilter e a S) X"  | 
|
112  | 
by(induction e arbitrary: a S)  | 
|
113  | 
(auto simp: Let_def update_def wt_st_def  | 
|
114  | 
split: option.splits prod.split)  | 
|
115  | 
||
116  | 
lemma afilter_sound: "wt S X \<Longrightarrow> vars e \<subseteq> X \<Longrightarrow>  | 
|
117  | 
s : \<gamma>\<^isub>o S \<Longrightarrow> aval e s : \<gamma> a \<Longrightarrow> s : \<gamma>\<^isub>o (afilter e a S)"  | 
|
118  | 
proof(induction e arbitrary: a S)  | 
|
119  | 
case N thus ?case by simp (metis test_num')  | 
|
120  | 
next  | 
|
121  | 
case (V x)  | 
|
122  | 
obtain S' where "S = Some S'" and "s : \<gamma>\<^isub>f S'" using `s : \<gamma>\<^isub>o S`  | 
|
123  | 
by(auto simp: in_gamma_option_iff)  | 
|
124  | 
moreover hence "s x : \<gamma> (fun S' x)"  | 
|
125  | 
using V(1,2) by(simp add: \<gamma>_st_def wt_st_def)  | 
|
126  | 
moreover have "s x : \<gamma> a" using V by simp  | 
|
127  | 
ultimately show ?case using V(3)  | 
|
128  | 
by(simp add: Let_def \<gamma>_st_def)  | 
|
129  | 
(metis mono_gamma emptyE in_gamma_meet gamma_Bot subset_empty)  | 
|
130  | 
next  | 
|
131  | 
case (Plus e1 e2) thus ?case  | 
|
132  | 
using filter_plus'[OF _ aval''_sound[OF Plus.prems(3)] aval''_sound[OF Plus.prems(3)]]  | 
|
133  | 
by (auto simp: wt_afilter split: prod.split)  | 
|
134  | 
qed  | 
|
135  | 
||
136  | 
lemma wt_bfilter: "wt S X \<Longrightarrow> vars b \<subseteq> X \<Longrightarrow> wt (bfilter b bv S) X"  | 
|
137  | 
by(induction b arbitrary: bv S)  | 
|
138  | 
(auto simp: wt_afilter split: prod.split)  | 
|
139  | 
||
140  | 
lemma bfilter_sound: "wt S X \<Longrightarrow> vars b \<subseteq> X \<Longrightarrow>  | 
|
141  | 
s : \<gamma>\<^isub>o S \<Longrightarrow> bv = bval b s \<Longrightarrow> s : \<gamma>\<^isub>o(bfilter b bv S)"  | 
|
142  | 
proof(induction b arbitrary: S bv)  | 
|
143  | 
case Bc thus ?case by simp  | 
|
144  | 
next  | 
|
145  | 
case (Not b) thus ?case by simp  | 
|
146  | 
next  | 
|
147  | 
case (And b1 b2) thus ?case  | 
|
148  | 
by simp (metis And(1) And(2) wt_bfilter in_gamma_join_UpI)  | 
|
149  | 
next  | 
|
150  | 
case (Less e1 e2) thus ?case  | 
|
151  | 
by(auto split: prod.split)  | 
|
152  | 
(metis (lifting) wt_afilter afilter_sound aval''_sound filter_less')  | 
|
153  | 
qed  | 
|
154  | 
||
155  | 
||
156  | 
fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom"  | 
|
157  | 
where  | 
|
158  | 
"step' S (SKIP {P}) = (SKIP {S})" |
 | 
|
159  | 
"step' S (x ::= e {P}) =
 | 
|
160  | 
  x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S))}" |
 | 
|
161  | 
"step' S (C1; C2) = step' S C1; step' (post C1) C2" |  | 
|
162  | 
"step' S (IF b THEN C1 ELSE C2 {P}) =
 | 
|
163  | 
(let D1 = step' (bfilter b True S) C1; D2 = step' (bfilter b False S) C2  | 
|
164  | 
   in IF b THEN D1 ELSE D2 {post C1 \<squnion> post C2})" |
 | 
|
165  | 
"step' S ({Inv} WHILE b DO C {P}) =
 | 
|
166  | 
   {S \<squnion> post C}
 | 
|
167  | 
WHILE b DO step' (bfilter b True Inv) C  | 
|
168  | 
   {bfilter b False Inv}"
 | 
|
169  | 
||
170  | 
definition AI :: "com \<Rightarrow> 'av st option acom option" where  | 
|
171  | 
"AI c = lpfp (step' \<top>\<^bsub>c\<^esub>) c"  | 
|
172  | 
||
173  | 
lemma strip_step'[simp]: "strip(step' S c) = strip c"  | 
|
174  | 
by(induct c arbitrary: S) (simp_all add: Let_def)  | 
|
175  | 
||
176  | 
||
177  | 
subsubsection "Soundness"  | 
|
178  | 
||
179  | 
lemma in_gamma_update:  | 
|
180  | 
"\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(update S x a)"  | 
|
181  | 
by(simp add: \<gamma>_st_def)  | 
|
182  | 
||
183  | 
theorem step_preserves_le:  | 
|
184  | 
"\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; C \<le> \<gamma>\<^isub>c C'; wt C' X; wt S' X \<rbrakk> \<Longrightarrow> step S C \<le> \<gamma>\<^isub>c (step' S' C')"  | 
|
185  | 
proof(induction C arbitrary: C' S S')  | 
|
186  | 
case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP)  | 
|
187  | 
next  | 
|
188  | 
case Assign thus ?case  | 
|
189  | 
by (fastforce simp: Assign_le map_acom_Assign wt_option_def wt_st_def  | 
|
190  | 
intro: aval'_sound in_gamma_update split: option.splits del:subsetD)  | 
|
191  | 
next  | 
|
| 47818 | 192  | 
case Seq thus ?case apply (auto simp: Seq_le map_acom_Seq)  | 
| 47613 | 193  | 
by (metis le_post post_map_acom wt_post)  | 
194  | 
next  | 
|
195  | 
case (If b C1 C2 P)  | 
|
196  | 
then obtain C1' C2' P' where  | 
|
197  | 
      "C' = IF b THEN C1' ELSE C2' {P'}"
 | 
|
198  | 
"P \<subseteq> \<gamma>\<^isub>o P'" "C1 \<le> \<gamma>\<^isub>c C1'" "C2 \<le> \<gamma>\<^isub>c C2'"  | 
|
199  | 
by (fastforce simp: If_le map_acom_If)  | 
|
200  | 
moreover from this(1) `wt C' X` have wt: "wt C1' X" "wt C2' X"  | 
|
201  | 
and "vars b \<subseteq> X" by simp_all  | 
|
202  | 
moreover have "post C1 \<subseteq> \<gamma>\<^isub>o(post C1' \<squnion> post C2')"  | 
|
203  | 
by (metis (no_types) `C1 \<le> \<gamma>\<^isub>c C1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom wt_post wt)  | 
|
204  | 
moreover have "post C2 \<subseteq> \<gamma>\<^isub>o(post C1' \<squnion> post C2')"  | 
|
205  | 
by (metis (no_types) `C2 \<le> \<gamma>\<^isub>c C2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom wt_post wt)  | 
|
206  | 
moreover note vars = `wt S' X` `vars b \<subseteq> X`  | 
|
207  | 
ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'`  | 
|
208  | 
by (simp add: If.IH subset_iff bfilter_sound[OF vars] wt_bfilter[OF vars])  | 
|
209  | 
next  | 
|
210  | 
case (While I b C1 P)  | 
|
211  | 
then obtain C1' I' P' where  | 
|
212  | 
    "C' = {I'} WHILE b DO C1' {P'}"
 | 
|
213  | 
"I \<subseteq> \<gamma>\<^isub>o I'" "P \<subseteq> \<gamma>\<^isub>o P'" "C1 \<le> \<gamma>\<^isub>c C1'"  | 
|
214  | 
by (fastforce simp: map_acom_While While_le)  | 
|
215  | 
moreover from this(1) `wt C' X`  | 
|
216  | 
have wt: "wt C1' X" "wt I' X" and "vars b \<subseteq> X" by simp_all  | 
|
217  | 
moreover note compat = `wt S' X` wt_post[OF wt(1)]  | 
|
218  | 
moreover note vars = `wt I' X` `vars b \<subseteq> X`  | 
|
219  | 
moreover have "S \<union> post C1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post C1')"  | 
|
220  | 
using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `C1 \<le> \<gamma>\<^isub>c C1'`, simplified]  | 
|
221  | 
by (metis (no_types) join_ge1[OF compat] join_ge2[OF compat] le_sup_iff mono_gamma_o order_trans)  | 
|
222  | 
ultimately show ?case  | 
|
223  | 
by (simp add: While.IH subset_iff bfilter_sound[OF vars] wt_bfilter[OF vars])  | 
|
224  | 
qed  | 
|
225  | 
||
226  | 
lemma wt_step'[simp]:  | 
|
227  | 
"\<lbrakk> wt C X; wt S X \<rbrakk> \<Longrightarrow> wt (step' S C) X"  | 
|
228  | 
proof(induction C arbitrary: S)  | 
|
229  | 
case Assign thus ?case by(simp add: wt_option_def wt_st_def update_def split: option.splits)  | 
|
230  | 
qed (auto simp add: wt_bfilter)  | 
|
231  | 
||
232  | 
theorem AI_sound: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C"  | 
|
233  | 
proof(simp add: CS_def AI_def)  | 
|
234  | 
assume 1: "lpfp (step' (top c)) c = Some C"  | 
|
235  | 
have "wt C (vars c)"  | 
|
236  | 
by(rule lpfp_inv[where P = "%C. wt C (vars c)", OF 1 _ wt_bot])  | 
|
237  | 
(erule wt_step'[OF _ wt_top])  | 
|
238  | 
have 2: "step' (top c) C \<sqsubseteq> C" by(rule lpfpc_pfp[OF 1])  | 
|
239  | 
have 3: "strip (\<gamma>\<^isub>c (step' (top c) C)) = c"  | 
|
240  | 
by(simp add: strip_lpfp[OF _ 1])  | 
|
| 
48759
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
241  | 
have "lfp c (step UNIV) \<le> \<gamma>\<^isub>c (step' (top c) C)"  | 
| 47613 | 242  | 
proof(rule lfp_lowerbound[simplified,OF 3])  | 
243  | 
show "step UNIV (\<gamma>\<^isub>c (step' (top c) C)) \<le> \<gamma>\<^isub>c (step' (top c) C)"  | 
|
244  | 
proof(rule step_preserves_le[OF _ _ `wt C (vars c)` wt_top])  | 
|
245  | 
show "UNIV \<subseteq> \<gamma>\<^isub>o (top c)" by simp  | 
|
246  | 
show "\<gamma>\<^isub>c (step' (top c) C) \<le> \<gamma>\<^isub>c C" by(rule mono_gamma_c[OF 2])  | 
|
247  | 
qed  | 
|
248  | 
qed  | 
|
| 
48759
 
ff570720ba1c
Improved complete lattice formalisation - no more index set.
 
nipkow 
parents: 
47818 
diff
changeset
 | 
249  | 
from this 2 show "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C"  | 
| 47613 | 250  | 
by (blast intro: mono_gamma_c order_trans)  | 
251  | 
qed  | 
|
252  | 
||
253  | 
end  | 
|
254  | 
||
255  | 
||
256  | 
subsubsection "Monotonicity"  | 
|
257  | 
||
258  | 
locale Abs_Int1_mono = Abs_Int1 +  | 
|
259  | 
assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2"  | 
|
260  | 
and mono_filter_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> r \<sqsubseteq> r' \<Longrightarrow>  | 
|
261  | 
filter_plus' r a1 a2 \<sqsubseteq> filter_plus' r' b1 b2"  | 
|
262  | 
and mono_filter_less': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow>  | 
|
263  | 
filter_less' bv a1 a2 \<sqsubseteq> filter_less' bv b1 b2"  | 
|
264  | 
begin  | 
|
265  | 
||
266  | 
lemma mono_aval':  | 
|
267  | 
"S1 \<sqsubseteq> S2 \<Longrightarrow> wt S1 X \<Longrightarrow> vars e \<subseteq> X \<Longrightarrow> aval' e S1 \<sqsubseteq> aval' e S2"  | 
|
268  | 
by(induction e) (auto simp: le_st_def mono_plus' wt_st_def)  | 
|
269  | 
||
270  | 
lemma mono_aval'':  | 
|
271  | 
"S1 \<sqsubseteq> S2 \<Longrightarrow> wt S1 X \<Longrightarrow> vars e \<subseteq> X \<Longrightarrow> aval'' e S1 \<sqsubseteq> aval'' e S2"  | 
|
272  | 
apply(cases S1)  | 
|
273  | 
apply simp  | 
|
274  | 
apply(cases S2)  | 
|
275  | 
apply simp  | 
|
276  | 
by (simp add: mono_aval')  | 
|
277  | 
||
278  | 
lemma mono_afilter: "wt S1 X \<Longrightarrow> wt S2 X \<Longrightarrow> vars e \<subseteq> X \<Longrightarrow>  | 
|
279  | 
r1 \<sqsubseteq> r2 \<Longrightarrow> S1 \<sqsubseteq> S2 \<Longrightarrow> afilter e r1 S1 \<sqsubseteq> afilter e r2 S2"  | 
|
280  | 
apply(induction e arbitrary: r1 r2 S1 S2)  | 
|
281  | 
apply(auto simp: test_num' Let_def mono_meet split: option.splits prod.splits)  | 
|
282  | 
apply (metis mono_gamma subsetD)  | 
|
283  | 
apply(drule (2) mono_fun_wt)  | 
|
284  | 
apply (metis mono_meet le_trans)  | 
|
285  | 
apply(metis mono_aval'' mono_filter_plus'[simplified le_prod_def] fst_conv snd_conv wt_afilter)  | 
|
286  | 
done  | 
|
287  | 
||
288  | 
lemma mono_bfilter: "wt S1 X \<Longrightarrow> wt S2 X \<Longrightarrow> vars b \<subseteq> X \<Longrightarrow>  | 
|
289  | 
S1 \<sqsubseteq> S2 \<Longrightarrow> bfilter b bv S1 \<sqsubseteq> bfilter b bv S2"  | 
|
290  | 
apply(induction b arbitrary: bv S1 S2)  | 
|
291  | 
apply(simp)  | 
|
292  | 
apply(simp)  | 
|
293  | 
apply simp  | 
|
294  | 
apply(metis join_least le_trans[OF _ join_ge1] le_trans[OF _ join_ge2] wt_bfilter)  | 
|
295  | 
apply (simp split: prod.splits)  | 
|
296  | 
apply(metis mono_aval'' mono_afilter mono_filter_less'[simplified le_prod_def] fst_conv snd_conv wt_afilter)  | 
|
297  | 
done  | 
|
298  | 
||
299  | 
theorem mono_step': "wt S1 X \<Longrightarrow> wt S2 X \<Longrightarrow> wt C1 X \<Longrightarrow> wt C2 X \<Longrightarrow>  | 
|
300  | 
S1 \<sqsubseteq> S2 \<Longrightarrow> C1 \<sqsubseteq> C2 \<Longrightarrow> step' S1 C1 \<sqsubseteq> step' S2 C2"  | 
|
301  | 
apply(induction C1 C2 arbitrary: S1 S2 rule: le_acom.induct)  | 
|
302  | 
apply (auto simp: Let_def mono_bfilter mono_aval' mono_post  | 
|
303  | 
le_join_disj le_join_disj[OF wt_post wt_post] wt_bfilter  | 
|
304  | 
split: option.split)  | 
|
305  | 
done  | 
|
306  | 
||
307  | 
lemma mono_step'_top: "wt C1 (vars c) \<Longrightarrow> wt C2 (vars c) \<Longrightarrow> C1 \<sqsubseteq> C2 \<Longrightarrow> step' (top c) C1 \<sqsubseteq> step' (top c) C2"  | 
|
308  | 
by (metis wt_top mono_step' preord_class.le_refl)  | 
|
309  | 
||
310  | 
end  | 
|
311  | 
||
312  | 
end  |