author | nipkow |
Tue, 15 Jan 2013 13:46:19 +0100 | |
changeset 50896 | fb0fcd278ac5 |
parent 49547 | 78be750222cf |
child 50986 | c54ea7f5418f |
permissions | -rw-r--r-- |
47613 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
theory Abs_Int1 |
|
4 |
imports Abs_State |
|
5 |
begin |
|
6 |
||
7 |
lemma le_iff_le_annos_zip: "C1 \<sqsubseteq> C2 \<longleftrightarrow> |
|
8 |
(\<forall> (a1,a2) \<in> set(zip (annos C1) (annos C2)). a1 \<sqsubseteq> a2) \<and> strip C1 = strip C2" |
|
9 |
by(induct C1 C2 rule: le_acom.induct) (auto simp: size_annos_same2) |
|
10 |
||
11 |
lemma le_iff_le_annos: "C1 \<sqsubseteq> C2 \<longleftrightarrow> |
|
12 |
strip C1 = strip C2 \<and> (\<forall> i<size(annos C1). annos C1 ! i \<sqsubseteq> annos C2 ! i)" |
|
13 |
by(auto simp add: le_iff_le_annos_zip set_zip) (metis size_annos_same2) |
|
14 |
||
15 |
||
49396 | 16 |
lemma mono_fun_L[simp]: "F \<in> L X \<Longrightarrow> F \<sqsubseteq> G \<Longrightarrow> x : X \<Longrightarrow> fun F x \<sqsubseteq> fun G x" |
17 |
by(simp add: mono_fun L_st_def) |
|
47613 | 18 |
|
49396 | 19 |
lemma bot_in_L[simp]: "bot c \<in> L(vars c)" |
20 |
by(simp add: L_acom_def bot_def) |
|
47613 | 21 |
|
49396 | 22 |
lemma L_acom_simps[simp]: "SKIP {P} \<in> L X \<longleftrightarrow> P \<in> L X" |
23 |
"(x ::= e {P}) \<in> L X \<longleftrightarrow> x : X \<and> vars e \<subseteq> X \<and> P \<in> L X" |
|
24 |
"(C1;C2) \<in> L X \<longleftrightarrow> C1 \<in> L X \<and> C2 \<in> L X" |
|
25 |
"(IF b THEN {P1} C1 ELSE {P2} C2 {Q}) \<in> L X \<longleftrightarrow> |
|
26 |
vars b \<subseteq> X \<and> C1 \<in> L X \<and> C2 \<in> L X \<and> P1 \<in> L X \<and> P2 \<in> L X \<and> Q \<in> L X" |
|
27 |
"({I} WHILE b DO {P} C {Q}) \<in> L X \<longleftrightarrow> |
|
28 |
I \<in> L X \<and> vars b \<subseteq> X \<and> P \<in> L X \<and> C \<in> L X \<and> Q \<in> L X" |
|
29 |
by(auto simp add: L_acom_def) |
|
47613 | 30 |
|
49353 | 31 |
lemma post_in_annos: "post C : set(annos C)" |
32 |
by(induction C) auto |
|
33 |
||
49396 | 34 |
lemma post_in_L[simp]: "C \<in> L X \<Longrightarrow> post C \<in> L X" |
35 |
by(simp add: L_acom_def post_in_annos) |
|
47613 | 36 |
|
37 |
||
38 |
subsection "Computable Abstract Interpretation" |
|
39 |
||
40 |
text{* Abstract interpretation over type @{text st} instead of |
|
41 |
functions. *} |
|
42 |
||
43 |
context Gamma |
|
44 |
begin |
|
45 |
||
46 |
fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where |
|
50896 | 47 |
"aval' (N i) S = num' i" | |
47613 | 48 |
"aval' (V x) S = fun S x" | |
49 |
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" |
|
50 |
||
49497 | 51 |
lemma aval'_sound: "s : \<gamma>\<^isub>s S \<Longrightarrow> vars a \<subseteq> dom S \<Longrightarrow> aval a s : \<gamma>(aval' a S)" |
47613 | 52 |
by (induction a) (auto simp: gamma_num' gamma_plus' \<gamma>_st_def) |
53 |
||
54 |
end |
|
55 |
||
56 |
text{* The for-clause (here and elsewhere) only serves the purpose of fixing |
|
57 |
the name of the type parameter @{typ 'av} which would otherwise be renamed to |
|
58 |
@{typ 'a}. *} |
|
59 |
||
49396 | 60 |
locale Abs_Int = Gamma where \<gamma>=\<gamma> for \<gamma> :: "'av::semilattice \<Rightarrow> val set" |
47613 | 61 |
begin |
62 |
||
63 |
fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom" where |
|
64 |
"step' S (SKIP {P}) = (SKIP {S})" | |
|
65 |
"step' S (x ::= e {P}) = |
|
66 |
x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S))}" | |
|
67 |
"step' S (C1; C2) = step' S C1; step' (post C1) C2" | |
|
49344 | 68 |
"step' S (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) = |
69 |
(IF b THEN {S} step' P1 C1 ELSE {S} step' P2 C2 {post C1 \<squnion> post C2})" | |
|
70 |
"step' S ({I} WHILE b DO {P} C {Q}) = |
|
71 |
{S \<squnion> post C} WHILE b DO {I} step' P C {I}" |
|
47613 | 72 |
|
73 |
definition AI :: "com \<Rightarrow> 'av st option acom option" where |
|
49464 | 74 |
"AI c = pfp (step' (top c)) (bot c)" |
47613 | 75 |
|
76 |
||
77 |
lemma strip_step'[simp]: "strip(step' S C) = strip C" |
|
78 |
by(induct C arbitrary: S) (simp_all add: Let_def) |
|
79 |
||
80 |
||
81 |
text{* Soundness: *} |
|
82 |
||
83 |
lemma in_gamma_update: |
|
49497 | 84 |
"\<lbrakk> s : \<gamma>\<^isub>s S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>s(update S x a)" |
47613 | 85 |
by(simp add: \<gamma>_st_def) |
86 |
||
87 |
theorem step_preserves_le: |
|
49396 | 88 |
"\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; C \<le> \<gamma>\<^isub>c C'; C' \<in> L X; S' \<in> L X \<rbrakk> \<Longrightarrow> step S C \<le> \<gamma>\<^isub>c (step' S' C')" |
47613 | 89 |
proof(induction C arbitrary: C' S S') |
90 |
case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) |
|
91 |
next |
|
92 |
case Assign thus ?case |
|
49396 | 93 |
by(fastforce simp: Assign_le map_acom_Assign L_st_def |
47613 | 94 |
intro: aval'_sound in_gamma_update split: option.splits) |
95 |
next |
|
47818 | 96 |
case Seq thus ?case apply (auto simp: Seq_le map_acom_Seq) |
49396 | 97 |
by (metis le_post post_map_acom post_in_L) |
47613 | 98 |
next |
49344 | 99 |
case (If b P1 C1 P2 C2 Q) |
100 |
then obtain P1' P2' C1' C2' Q' where |
|
101 |
"C' = IF b THEN {P1'} C1' ELSE {P2'} C2' {Q'}" |
|
102 |
"P1 \<subseteq> \<gamma>\<^isub>o P1'" "P2 \<subseteq> \<gamma>\<^isub>o P2'" "Q \<subseteq> \<gamma>\<^isub>o Q'" "C1 \<le> \<gamma>\<^isub>c C1'" "C2 \<le> \<gamma>\<^isub>c C2'" |
|
47613 | 103 |
by (fastforce simp: If_le map_acom_If) |
49396 | 104 |
moreover from this(1) `C' \<in> L X` |
105 |
have L: "C1' \<in> L X" "C2' \<in> L X" "P1' \<in> L X" "P2' \<in> L X" by simp_all |
|
47613 | 106 |
moreover have "post C1 \<subseteq> \<gamma>\<^isub>o(post C1' \<squnion> post C2')" |
49396 | 107 |
by (metis (no_types) `C1 \<le> \<gamma>\<^isub>c C1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom L post_in_L) |
47613 | 108 |
moreover have "post C2 \<subseteq> \<gamma>\<^isub>o(post C1' \<squnion> post C2')" |
49396 | 109 |
by (metis (no_types) `C2 \<le> \<gamma>\<^isub>c C2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom L post_in_L) |
110 |
ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` `S' \<in> L X` |
|
47613 | 111 |
by (simp add: If.IH subset_iff) |
112 |
next |
|
49344 | 113 |
case (While I b P1 C1 Q) |
114 |
then obtain C1' I' P1' Q' where |
|
115 |
"C' = {I'} WHILE b DO {P1'} C1' {Q'}" |
|
116 |
"I \<subseteq> \<gamma>\<^isub>o I'" "P1 \<subseteq> \<gamma>\<^isub>o P1'" "C1 \<le> \<gamma>\<^isub>c C1'" "Q \<subseteq> \<gamma>\<^isub>o Q'" |
|
47613 | 117 |
by (fastforce simp: map_acom_While While_le) |
49396 | 118 |
moreover from this(1) `C' \<in> L X` |
119 |
have L: "C1' \<in> L X" "I' \<in> L X" "P1' \<in> L X" by simp_all |
|
120 |
moreover note compat = `S' \<in> L X` post_in_L[OF L(1)] |
|
47613 | 121 |
moreover have "S \<union> post C1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post C1')" |
122 |
using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `C1 \<le> \<gamma>\<^isub>c C1'`, simplified] |
|
123 |
by (metis (no_types) join_ge1[OF compat] join_ge2[OF compat] le_sup_iff mono_gamma_o order_trans) |
|
124 |
ultimately show ?case by (simp add: While.IH subset_iff) |
|
125 |
qed |
|
126 |
||
49396 | 127 |
lemma step'_in_L[simp]: |
128 |
"\<lbrakk> C \<in> L X; S \<in> L X \<rbrakk> \<Longrightarrow> (step' S C) \<in> L X" |
|
47613 | 129 |
proof(induction C arbitrary: S) |
130 |
case Assign thus ?case |
|
49396 | 131 |
by(auto simp: L_st_def update_def split: option.splits) |
47613 | 132 |
qed auto |
133 |
||
134 |
theorem AI_sound: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C" |
|
135 |
proof(simp add: CS_def AI_def) |
|
49464 | 136 |
assume 1: "pfp (step' (top c)) (bot c) = Some C" |
49396 | 137 |
have "C \<in> L(vars c)" |
49464 | 138 |
by(rule pfp_inv[where P = "%C. C \<in> L(vars c)", OF 1 _ bot_in_L]) |
49396 | 139 |
(erule step'_in_L[OF _ top_in_L]) |
49464 | 140 |
have 2: "step' (top c) C \<sqsubseteq> C" by(rule pfp_pfp[OF 1]) |
47613 | 141 |
have 3: "strip (\<gamma>\<^isub>c (step' (top c) C)) = c" |
49464 | 142 |
by(simp add: strip_pfp[OF _ 1]) |
48759
ff570720ba1c
Improved complete lattice formalisation - no more index set.
nipkow
parents:
47818
diff
changeset
|
143 |
have "lfp c (step UNIV) \<le> \<gamma>\<^isub>c (step' (top c) C)" |
47613 | 144 |
proof(rule lfp_lowerbound[simplified,OF 3]) |
145 |
show "step UNIV (\<gamma>\<^isub>c (step' (top c) C)) \<le> \<gamma>\<^isub>c (step' (top c) C)" |
|
49396 | 146 |
proof(rule step_preserves_le[OF _ _ `C \<in> L(vars c)` top_in_L]) |
47613 | 147 |
show "UNIV \<subseteq> \<gamma>\<^isub>o (top c)" by simp |
148 |
show "\<gamma>\<^isub>c (step' (top c) C) \<le> \<gamma>\<^isub>c C" by(rule mono_gamma_c[OF 2]) |
|
149 |
qed |
|
150 |
qed |
|
48759
ff570720ba1c
Improved complete lattice formalisation - no more index set.
nipkow
parents:
47818
diff
changeset
|
151 |
from this 2 show "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" |
47613 | 152 |
by (blast intro: mono_gamma_c order_trans) |
153 |
qed |
|
154 |
||
155 |
end |
|
156 |
||
157 |
||
158 |
subsubsection "Monotonicity" |
|
159 |
||
49396 | 160 |
lemma le_join_disj: "y \<in> L X \<Longrightarrow> (z::_::semilatticeL) \<in> L X \<Longrightarrow> |
161 |
x \<sqsubseteq> y \<or> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<squnion> z" |
|
47613 | 162 |
by (metis join_ge1 join_ge2 preord_class.le_trans) |
163 |
||
164 |
locale Abs_Int_mono = Abs_Int + |
|
165 |
assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2" |
|
166 |
begin |
|
167 |
||
49396 | 168 |
lemma mono_aval': |
49402 | 169 |
"S1 \<sqsubseteq> S2 \<Longrightarrow> S1 \<in> L X \<Longrightarrow> S2 \<in> L X \<Longrightarrow> vars e \<subseteq> X \<Longrightarrow> aval' e S1 \<sqsubseteq> aval' e S2" |
49396 | 170 |
by(induction e) (auto simp: le_st_def mono_plus' L_st_def) |
47613 | 171 |
|
49396 | 172 |
theorem mono_step': "S1 \<in> L X \<Longrightarrow> S2 \<in> L X \<Longrightarrow> C1 \<in> L X \<Longrightarrow> C2 \<in> L X \<Longrightarrow> |
47613 | 173 |
S1 \<sqsubseteq> S2 \<Longrightarrow> C1 \<sqsubseteq> C2 \<Longrightarrow> step' S1 C1 \<sqsubseteq> step' S2 C2" |
174 |
apply(induction C1 C2 arbitrary: S1 S2 rule: le_acom.induct) |
|
175 |
apply (auto simp: Let_def mono_aval' mono_post |
|
49396 | 176 |
le_join_disj le_join_disj[OF post_in_L post_in_L] |
47613 | 177 |
split: option.split) |
178 |
done |
|
179 |
||
49396 | 180 |
lemma mono_step'_top: "C \<in> L(vars c) \<Longrightarrow> C' \<in> L(vars c) \<Longrightarrow> |
181 |
C \<sqsubseteq> C' \<Longrightarrow> step' (top c) C \<sqsubseteq> step' (top c) C'" |
|
182 |
by (metis top_in_L mono_step' preord_class.le_refl) |
|
47613 | 183 |
|
49464 | 184 |
lemma pfp_bot_least: |
185 |
assumes "\<forall>x\<in>L(vars c)\<inter>{C. strip C = c}.\<forall>y\<in>L(vars c)\<inter>{C. strip C = c}. |
|
186 |
x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y" |
|
187 |
and "\<forall>C. C \<in> L(vars c)\<inter>{C. strip C = c} \<longrightarrow> f C \<in> L(vars c)\<inter>{C. strip C = c}" |
|
188 |
and "f C' \<sqsubseteq> C'" "strip C' = c" "C' \<in> L(vars c)" "pfp f (bot c) = Some C" |
|
189 |
shows "C \<sqsubseteq> C'" |
|
190 |
apply(rule while_least[OF assms(1,2) _ _ assms(3) _ assms(6)[unfolded pfp_def]]) |
|
191 |
by (simp_all add: assms(4,5) bot_least) |
|
192 |
||
193 |
lemma AI_least_pfp: assumes "AI c = Some C" |
|
194 |
and "step' (top c) C' \<sqsubseteq> C'" "strip C' = c" "C' \<in> L(vars c)" |
|
195 |
shows "C \<sqsubseteq> C'" |
|
196 |
apply(rule pfp_bot_least[OF _ _ assms(2-4) assms(1)[unfolded AI_def]]) |
|
197 |
by (simp_all add: mono_step'_top) |
|
198 |
||
47613 | 199 |
end |
200 |
||
201 |
||
202 |
subsubsection "Termination" |
|
203 |
||
204 |
abbreviation sqless (infix "\<sqsubset>" 50) where |
|
205 |
"x \<sqsubset> y == x \<sqsubseteq> y \<and> \<not> y \<sqsubseteq> x" |
|
206 |
||
207 |
lemma pfp_termination: |
|
208 |
fixes x0 :: "'a::preord" and m :: "'a \<Rightarrow> nat" |
|
209 |
assumes mono: "\<And>x y. I x \<Longrightarrow> I y \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" |
|
210 |
and m: "\<And>x y. I x \<Longrightarrow> I y \<Longrightarrow> x \<sqsubset> y \<Longrightarrow> m x > m y" |
|
211 |
and I: "\<And>x y. I x \<Longrightarrow> I(f x)" and "I x0" and "x0 \<sqsubseteq> f x0" |
|
49464 | 212 |
shows "\<exists>x. pfp f x0 = Some x" |
47613 | 213 |
proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. I x & x \<sqsubseteq> f x"]) |
214 |
show "wf {(y,x). ((I x \<and> x \<sqsubseteq> f x) \<and> \<not> f x \<sqsubseteq> x) \<and> y = f x}" |
|
215 |
by(rule wf_subset[OF wf_measure[of m]]) (auto simp: m I) |
|
216 |
next |
|
217 |
show "I x0 \<and> x0 \<sqsubseteq> f x0" using `I x0` `x0 \<sqsubseteq> f x0` by blast |
|
218 |
next |
|
219 |
fix x assume "I x \<and> x \<sqsubseteq> f x" thus "I(f x) \<and> f x \<sqsubseteq> f(f x)" |
|
220 |
by (blast intro: I mono) |
|
221 |
qed |
|
222 |
||
223 |
||
49547 | 224 |
locale Measure1 = |
225 |
fixes m :: "'av::preord \<Rightarrow> nat" |
|
47613 | 226 |
fixes h :: "nat" |
227 |
assumes m1: "x \<sqsubseteq> y \<Longrightarrow> m x \<ge> m y" |
|
228 |
assumes h: "m x \<le> h" |
|
229 |
begin |
|
230 |
||
49547 | 231 |
definition m_s :: "'av st \<Rightarrow> nat" ("m\<^isub>s") where |
232 |
"m_s S = (\<Sum> x \<in> dom S. m(fun S x))" |
|
47613 | 233 |
|
49547 | 234 |
lemma m_s_h: "x \<in> L X \<Longrightarrow> finite X \<Longrightarrow> m_s x \<le> h * card X" |
235 |
by(simp add: L_st_def m_s_def) |
|
49431 | 236 |
(metis nat_mult_commute of_nat_id setsum_bounded[OF h]) |
237 |
||
49547 | 238 |
lemma m_s1: "S1 \<sqsubseteq> S2 \<Longrightarrow> m_s S1 \<ge> m_s S2" |
239 |
proof(auto simp add: le_st_def m_s_def) |
|
47613 | 240 |
assume "\<forall>x\<in>dom S2. fun S1 x \<sqsubseteq> fun S2 x" |
241 |
hence "\<forall>x\<in>dom S2. m(fun S1 x) \<ge> m(fun S2 x)" by (metis m1) |
|
242 |
thus "(\<Sum>x\<in>dom S2. m (fun S2 x)) \<le> (\<Sum>x\<in>dom S2. m (fun S1 x))" |
|
243 |
by (metis setsum_mono) |
|
244 |
qed |
|
245 |
||
49432 | 246 |
definition m_o :: "nat \<Rightarrow> 'av st option \<Rightarrow> nat" ("m\<^isub>o") where |
49547 | 247 |
"m_o d opt = (case opt of None \<Rightarrow> h*d+1 | Some S \<Rightarrow> m_s S)" |
47613 | 248 |
|
49431 | 249 |
lemma m_o_h: "ost \<in> L X \<Longrightarrow> finite X \<Longrightarrow> m_o (card X) ost \<le> (h*card X + 1)" |
49547 | 250 |
by(auto simp add: m_o_def m_s_h split: option.split dest!:m_s_h) |
47613 | 251 |
|
49396 | 252 |
lemma m_o1: "finite X \<Longrightarrow> o1 \<in> L X \<Longrightarrow> o2 \<in> L X \<Longrightarrow> |
47613 | 253 |
o1 \<sqsubseteq> o2 \<Longrightarrow> m_o (card X) o1 \<ge> m_o (card X) o2" |
254 |
proof(induction o1 o2 rule: le_option.induct) |
|
49547 | 255 |
case 1 thus ?case by (simp add: m_o_def)(metis m_s1) |
47613 | 256 |
next |
257 |
case 2 thus ?case |
|
49547 | 258 |
by(simp add: L_option_def m_o_def le_SucI m_s_h split: option.splits) |
47613 | 259 |
next |
260 |
case 3 thus ?case by simp |
|
261 |
qed |
|
262 |
||
49432 | 263 |
definition m_c :: "'av st option acom \<Rightarrow> nat" ("m\<^isub>c") where |
49431 | 264 |
"m_c C = (\<Sum>i<size(annos C). m_o (card(vars(strip C))) (annos C ! i))" |
265 |
||
266 |
lemma m_c_h: assumes "C \<in> L(vars(strip C))" |
|
267 |
shows "m_c C \<le> size(annos C) * (h * card(vars(strip C)) + 1)" |
|
268 |
proof- |
|
269 |
let ?X = "vars(strip C)" let ?n = "card ?X" let ?a = "size(annos C)" |
|
270 |
{ fix i assume "i < ?a" |
|
271 |
hence "annos C ! i \<in> L ?X" using assms by(simp add: L_acom_def) |
|
272 |
note m_o_h[OF this finite_cvars] |
|
273 |
} note 1 = this |
|
274 |
have "m_c C = (\<Sum>i<?a. m_o ?n (annos C ! i))" by(simp add: m_c_def) |
|
275 |
also have "\<dots> \<le> (\<Sum>i<?a. h * ?n + 1)" |
|
276 |
apply(rule setsum_mono) using 1 by simp |
|
277 |
also have "\<dots> = ?a * (h * ?n + 1)" by simp |
|
278 |
finally show ?thesis . |
|
279 |
qed |
|
280 |
||
49547 | 281 |
end |
282 |
||
283 |
locale Measure = Measure1 + |
|
284 |
assumes m2: "x \<sqsubset> y \<Longrightarrow> m x > m y" |
|
285 |
begin |
|
286 |
||
287 |
lemma m_s2: "finite(dom S1) \<Longrightarrow> S1 \<sqsubset> S2 \<Longrightarrow> m_s S1 > m_s S2" |
|
288 |
proof(auto simp add: le_st_def m_s_def) |
|
289 |
assume "finite(dom S2)" and 0: "\<forall>x\<in>dom S2. fun S1 x \<sqsubseteq> fun S2 x" |
|
290 |
hence 1: "\<forall>x\<in>dom S2. m(fun S1 x) \<ge> m(fun S2 x)" by (metis m1) |
|
291 |
fix x assume "x \<in> dom S2" "\<not> fun S2 x \<sqsubseteq> fun S1 x" |
|
292 |
hence 2: "\<exists>x\<in>dom S2. m(fun S1 x) > m(fun S2 x)" using 0 m2 by blast |
|
293 |
from setsum_strict_mono_ex1[OF `finite(dom S2)` 1 2] |
|
294 |
show "(\<Sum>x\<in>dom S2. m (fun S2 x)) < (\<Sum>x\<in>dom S2. m (fun S1 x))" . |
|
295 |
qed |
|
296 |
||
297 |
lemma m_o2: "finite X \<Longrightarrow> o1 \<in> L X \<Longrightarrow> o2 \<in> L X \<Longrightarrow> |
|
298 |
o1 \<sqsubset> o2 \<Longrightarrow> m_o (card X) o1 > m_o (card X) o2" |
|
299 |
proof(induction o1 o2 rule: le_option.induct) |
|
300 |
case 1 thus ?case by (simp add: m_o_def L_st_def m_s2) |
|
301 |
next |
|
302 |
case 2 thus ?case |
|
303 |
by(auto simp add: m_o_def le_imp_less_Suc m_s_h) |
|
304 |
next |
|
305 |
case 3 thus ?case by simp |
|
306 |
qed |
|
307 |
||
308 |
||
49396 | 309 |
lemma m_c2: "C1 \<in> L(vars(strip C1)) \<Longrightarrow> C2 \<in> L(vars(strip C2)) \<Longrightarrow> |
310 |
C1 \<sqsubset> C2 \<Longrightarrow> m_c C1 > m_c C2" |
|
311 |
proof(auto simp add: le_iff_le_annos m_c_def size_annos_same[of C1 C2] L_acom_def) |
|
312 |
let ?X = "vars(strip C2)" |
|
47613 | 313 |
let ?n = "card ?X" |
49396 | 314 |
assume V1: "\<forall>a\<in>set(annos C1). a \<in> L ?X" |
315 |
and V2: "\<forall>a\<in>set(annos C2). a \<in> L ?X" |
|
316 |
and strip_eq: "strip C1 = strip C2" |
|
317 |
and 0: "\<forall>i<size(annos C2). annos C1 ! i \<sqsubseteq> annos C2 ! i" |
|
318 |
hence 1: "\<forall>i<size(annos C2). m_o ?n (annos C1 ! i) \<ge> m_o ?n (annos C2 ! i)" |
|
47613 | 319 |
by (auto simp: all_set_conv_all_nth) |
320 |
(metis finite_cvars m_o1 size_annos_same2) |
|
49396 | 321 |
fix i assume "i < size(annos C2)" "\<not> annos C2 ! i \<sqsubseteq> annos C1 ! i" |
322 |
hence "m_o ?n (annos C1 ! i) > m_o ?n (annos C2 ! i)" (is "?P i") |
|
323 |
by(metis m_o2[OF finite_cvars] V1 V2 nth_mem size_annos_same[OF strip_eq] 0) |
|
324 |
hence 2: "\<exists>i < size(annos C2). ?P i" using `i < size(annos C2)` by blast |
|
325 |
show "(\<Sum>i<size(annos C2). m_o ?n (annos C2 ! i)) |
|
326 |
< (\<Sum>i<size(annos C2). m_o ?n (annos C1 ! i))" |
|
47613 | 327 |
apply(rule setsum_strict_mono_ex1) using 1 2 by (auto) |
328 |
qed |
|
329 |
||
49547 | 330 |
end |
331 |
||
332 |
locale Abs_Int_measure = |
|
333 |
Abs_Int_mono where \<gamma>=\<gamma> + Measure where m=m |
|
334 |
for \<gamma> :: "'av::semilattice \<Rightarrow> val set" and m :: "'av \<Rightarrow> nat" |
|
335 |
begin |
|
336 |
||
47613 | 337 |
lemma AI_Some_measure: "\<exists>C. AI c = Some C" |
338 |
unfolding AI_def |
|
49464 | 339 |
apply(rule pfp_termination[where I = "%C. strip C = c \<and> C \<in> L(vars c)" |
47613 | 340 |
and m="m_c"]) |
49464 | 341 |
apply(simp_all add: m_c2 mono_step'_top bot_least) |
47613 | 342 |
done |
343 |
||
344 |
end |
|
345 |
||
346 |
end |