| author | blanchet | 
| Tue, 09 Sep 2014 20:51:36 +0200 | |
| changeset 58255 | 9dfe8506c04d | 
| parent 57512 | cc97b347b301 | 
| child 58947 | 79684150ed6a | 
| permissions | -rw-r--r-- | 
| 47613 | 1  | 
(* Author: Tobias Nipkow *)  | 
2  | 
||
3  | 
theory Abs_Int0  | 
|
4  | 
imports Abs_Int_init  | 
|
5  | 
begin  | 
|
6  | 
||
7  | 
subsection "Orderings"  | 
|
8  | 
||
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
52504 
diff
changeset
 | 
9  | 
text{* The basic type classes @{class order}, @{class semilattice_sup} and @{class order_top} are
 | 
| 51625 | 10  | 
defined in @{theory Main}, more precisely in theories @{theory Orderings} and @{theory Lattices}.
 | 
11  | 
If you view this theory with jedit, just click on the names to get there. *}  | 
|
| 47613 | 12  | 
|
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
52504 
diff
changeset
 | 
13  | 
class semilattice_sup_top = semilattice_sup + order_top  | 
| 47613 | 14  | 
|
15  | 
||
| 51826 | 16  | 
instance "fun" :: (type, semilattice_sup_top) semilattice_sup_top ..  | 
| 47613 | 17  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
18  | 
instantiation option :: (order)order  | 
| 47613 | 19  | 
begin  | 
20  | 
||
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
21  | 
fun less_eq_option where  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
22  | 
"Some x \<le> Some y = (x \<le> y)" |  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
23  | 
"None \<le> y = True" |  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
24  | 
"Some _ \<le> None = False"  | 
| 47613 | 25  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
26  | 
definition less_option where "x < (y::'a option) = (x \<le> y \<and> \<not> y \<le> x)"  | 
| 47613 | 27  | 
|
| 51628 | 28  | 
lemma le_None[simp]: "(x \<le> None) = (x = None)"  | 
| 47613 | 29  | 
by (cases x) simp_all  | 
30  | 
||
| 51628 | 31  | 
lemma Some_le[simp]: "(Some x \<le> u) = (\<exists>y. u = Some y \<and> x \<le> y)"  | 
| 47613 | 32  | 
by (cases u) auto  | 
33  | 
||
34  | 
instance proof  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
35  | 
case goal1 show ?case by(rule less_option_def)  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
36  | 
next  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
37  | 
case goal2 show ?case by(cases x, simp_all)  | 
| 47613 | 38  | 
next  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
39  | 
case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, auto)  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
40  | 
next  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
41  | 
case goal4 thus ?case by(cases y, simp, cases x, auto)  | 
| 47613 | 42  | 
qed  | 
43  | 
||
44  | 
end  | 
|
45  | 
||
| 51389 | 46  | 
instantiation option :: (sup)sup  | 
| 47613 | 47  | 
begin  | 
48  | 
||
| 51389 | 49  | 
fun sup_option where  | 
| 47613 | 50  | 
"Some x \<squnion> Some y = Some(x \<squnion> y)" |  | 
51  | 
"None \<squnion> y = y" |  | 
|
52  | 
"x \<squnion> None = x"  | 
|
53  | 
||
| 51389 | 54  | 
lemma sup_None2[simp]: "x \<squnion> None = x"  | 
| 47613 | 55  | 
by (cases x) simp_all  | 
56  | 
||
57  | 
instance ..  | 
|
58  | 
||
59  | 
end  | 
|
60  | 
||
| 51826 | 61  | 
instantiation option :: (semilattice_sup_top)semilattice_sup_top  | 
| 47613 | 62  | 
begin  | 
63  | 
||
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
64  | 
definition top_option where "\<top> = Some \<top>"  | 
| 47613 | 65  | 
|
66  | 
instance proof  | 
|
| 51389 | 67  | 
case goal4 show ?case by(cases a, simp_all add: top_option_def)  | 
| 47613 | 68  | 
next  | 
| 51389 | 69  | 
case goal1 thus ?case by(cases x, simp, cases y, simp_all)  | 
| 47613 | 70  | 
next  | 
| 51389 | 71  | 
case goal2 thus ?case by(cases y, simp, cases x, simp_all)  | 
| 47613 | 72  | 
next  | 
| 51389 | 73  | 
case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all)  | 
| 47613 | 74  | 
qed  | 
75  | 
||
76  | 
end  | 
|
77  | 
||
| 51390 | 78  | 
lemma [simp]: "(Some x < Some y) = (x < y)"  | 
79  | 
by(auto simp: less_le)  | 
|
80  | 
||
| 
52729
 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 
haftmann 
parents: 
52504 
diff
changeset
 | 
81  | 
instantiation option :: (order)order_bot  | 
| 47613 | 82  | 
begin  | 
83  | 
||
| 49396 | 84  | 
definition bot_option :: "'a option" where  | 
| 47613 | 85  | 
"\<bottom> = None"  | 
86  | 
||
87  | 
instance  | 
|
88  | 
proof  | 
|
| 49396 | 89  | 
case goal1 thus ?case by(auto simp: bot_option_def)  | 
| 47613 | 90  | 
qed  | 
91  | 
||
92  | 
end  | 
|
93  | 
||
94  | 
||
95  | 
definition bot :: "com \<Rightarrow> 'a option acom" where  | 
|
| 
52019
 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 
nipkow 
parents: 
51974 
diff
changeset
 | 
96  | 
"bot c = annotate (\<lambda>p. None) c"  | 
| 47613 | 97  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
98  | 
lemma bot_least: "strip C = c \<Longrightarrow> bot c \<le> C"  | 
| 
52019
 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 
nipkow 
parents: 
51974 
diff
changeset
 | 
99  | 
by(auto simp: bot_def less_eq_acom_def)  | 
| 47613 | 100  | 
|
101  | 
lemma strip_bot[simp]: "strip(bot c) = c"  | 
|
102  | 
by(simp add: bot_def)  | 
|
103  | 
||
104  | 
||
| 51722 | 105  | 
subsubsection "Pre-fixpoint iteration"  | 
| 47613 | 106  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
107  | 
definition pfp :: "(('a::order) \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where
 | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
108  | 
"pfp f = while_option (\<lambda>x. \<not> f x \<le> x) f"  | 
| 47613 | 109  | 
|
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
110  | 
lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x \<le> x"  | 
| 47613 | 111  | 
using while_option_stop[OF assms[simplified pfp_def]] by simp  | 
112  | 
||
| 49464 | 113  | 
lemma while_least:  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
114  | 
fixes q :: "'a::order"  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
115  | 
assumes "\<forall>x\<in>L.\<forall>y\<in>L. x \<le> y \<longrightarrow> f x \<le> f y" and "\<forall>x. x \<in> L \<longrightarrow> f x \<in> L"  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
116  | 
and "\<forall>x \<in> L. b \<le> x" and "b \<in> L" and "f q \<le> q" and "q \<in> L"  | 
| 49464 | 117  | 
and "while_option P f b = Some p"  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
118  | 
shows "p \<le> q"  | 
| 49464 | 119  | 
using while_option_rule[OF _ assms(7)[unfolded pfp_def],  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
120  | 
where P = "%x. x \<in> L \<and> x \<le> q"]  | 
| 
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
121  | 
by (metis assms(1-6) order_trans)  | 
| 47613 | 122  | 
|
| 51710 | 123  | 
lemma pfp_bot_least:  | 
124  | 
assumes "\<forall>x\<in>{C. strip C = c}.\<forall>y\<in>{C. strip C = c}. x \<le> y \<longrightarrow> f x \<le> f y"
 | 
|
125  | 
and "\<forall>C. C \<in> {C. strip C = c} \<longrightarrow> f C \<in> {C. strip C = c}"
 | 
|
126  | 
and "f C' \<le> C'" "strip C' = c" "pfp f (bot c) = Some C"  | 
|
127  | 
shows "C \<le> C'"  | 
|
128  | 
by(rule while_least[OF assms(1,2) _ _ assms(3) _ assms(5)[unfolded pfp_def]])  | 
|
129  | 
(simp_all add: assms(4) bot_least)  | 
|
130  | 
||
| 49464 | 131  | 
lemma pfp_inv:  | 
132  | 
"pfp f x = Some y \<Longrightarrow> (\<And>x. P x \<Longrightarrow> P(f x)) \<Longrightarrow> P x \<Longrightarrow> P y"  | 
|
133  | 
unfolding pfp_def by (metis (lifting) while_option_rule)  | 
|
| 47613 | 134  | 
|
135  | 
lemma strip_pfp:  | 
|
136  | 
assumes "\<And>x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0"  | 
|
| 49464 | 137  | 
using pfp_inv[OF assms(2), where P = "%x. g x = g x0"] assms(1) by simp  | 
| 47613 | 138  | 
|
139  | 
||
140  | 
subsection "Abstract Interpretation"  | 
|
141  | 
||
142  | 
definition \<gamma>_fun :: "('a \<Rightarrow> 'b set) \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> ('c \<Rightarrow> 'b)set" where
 | 
|
143  | 
"\<gamma>_fun \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(F x)}"
 | 
|
144  | 
||
145  | 
fun \<gamma>_option :: "('a \<Rightarrow> 'b set) \<Rightarrow> 'a option \<Rightarrow> 'b set" where
 | 
|
146  | 
"\<gamma>_option \<gamma> None = {}" |
 | 
|
147  | 
"\<gamma>_option \<gamma> (Some a) = \<gamma> a"  | 
|
148  | 
||
149  | 
text{* The interface for abstract values: *}
 | 
|
150  | 
||
| 52504 | 151  | 
locale Val_semilattice =  | 
| 51826 | 152  | 
fixes \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set"  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
153  | 
assumes mono_gamma: "a \<le> b \<Longrightarrow> \<gamma> a \<le> \<gamma> b"  | 
| 47613 | 154  | 
and gamma_Top[simp]: "\<gamma> \<top> = UNIV"  | 
155  | 
fixes num' :: "val \<Rightarrow> 'av"  | 
|
156  | 
and plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av"  | 
|
| 51036 | 157  | 
assumes gamma_num': "i \<in> \<gamma>(num' i)"  | 
158  | 
and gamma_plus': "i1 \<in> \<gamma> a1 \<Longrightarrow> i2 \<in> \<gamma> a2 \<Longrightarrow> i1+i2 \<in> \<gamma>(plus' a1 a2)"  | 
|
| 47613 | 159  | 
|
160  | 
type_synonym 'av st = "(vname \<Rightarrow> 'av)"  | 
|
161  | 
||
| 51826 | 162  | 
text{* The for-clause (here and elsewhere) only serves the purpose of fixing
 | 
163  | 
the name of the type parameter @{typ 'av} which would otherwise be renamed to
 | 
|
164  | 
@{typ 'a}. *}
 | 
|
165  | 
||
| 52504 | 166  | 
locale Abs_Int_fun = Val_semilattice where \<gamma>=\<gamma>  | 
| 51826 | 167  | 
for \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set"  | 
| 47613 | 168  | 
begin  | 
169  | 
||
170  | 
fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where  | 
|
| 50896 | 171  | 
"aval' (N i) S = num' i" |  | 
| 47613 | 172  | 
"aval' (V x) S = S x" |  | 
173  | 
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"  | 
|
174  | 
||
| 51807 | 175  | 
definition "asem x e S = (case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(S(x := aval' e S)))"  | 
| 51694 | 176  | 
|
| 51807 | 177  | 
definition "step' = Step asem (\<lambda>b S. S)"  | 
| 51389 | 178  | 
|
| 51722 | 179  | 
lemma strip_step'[simp]: "strip(step' S C) = strip C"  | 
180  | 
by(simp add: step'_def)  | 
|
181  | 
||
| 47613 | 182  | 
definition AI :: "com \<Rightarrow> 'av st option acom option" where  | 
| 49464 | 183  | 
"AI c = pfp (step' \<top>) (bot c)"  | 
| 47613 | 184  | 
|
185  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
186  | 
abbreviation \<gamma>\<^sub>s :: "'av st \<Rightarrow> state set"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
187  | 
where "\<gamma>\<^sub>s == \<gamma>_fun \<gamma>"  | 
| 47613 | 188  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
189  | 
abbreviation \<gamma>\<^sub>o :: "'av st option \<Rightarrow> state set"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
190  | 
where "\<gamma>\<^sub>o == \<gamma>_option \<gamma>\<^sub>s"  | 
| 47613 | 191  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
192  | 
abbreviation \<gamma>\<^sub>c :: "'av st option acom \<Rightarrow> state set acom"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
193  | 
where "\<gamma>\<^sub>c == map_acom \<gamma>\<^sub>o"  | 
| 47613 | 194  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
195  | 
lemma gamma_s_Top[simp]: "\<gamma>\<^sub>s \<top> = UNIV"  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
196  | 
by(simp add: top_fun_def \<gamma>_fun_def)  | 
| 47613 | 197  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
198  | 
lemma gamma_o_Top[simp]: "\<gamma>\<^sub>o \<top> = UNIV"  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
199  | 
by (simp add: top_option_def)  | 
| 47613 | 200  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
201  | 
lemma mono_gamma_s: "f1 \<le> f2 \<Longrightarrow> \<gamma>\<^sub>s f1 \<subseteq> \<gamma>\<^sub>s f2"  | 
| 47613 | 202  | 
by(auto simp: le_fun_def \<gamma>_fun_def dest: mono_gamma)  | 
203  | 
||
204  | 
lemma mono_gamma_o:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
205  | 
"S1 \<le> S2 \<Longrightarrow> \<gamma>\<^sub>o S1 \<subseteq> \<gamma>\<^sub>o S2"  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
206  | 
by(induction S1 S2 rule: less_eq_option.induct)(simp_all add: mono_gamma_s)  | 
| 47613 | 207  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
208  | 
lemma mono_gamma_c: "C1 \<le> C2 \<Longrightarrow> \<gamma>\<^sub>c C1 \<le> \<gamma>\<^sub>c C2"  | 
| 
52019
 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 
nipkow 
parents: 
51974 
diff
changeset
 | 
209  | 
by (simp add: less_eq_acom_def mono_gamma_o size_annos anno_map_acom size_annos_same[of C1 C2])  | 
| 47613 | 210  | 
|
| 51974 | 211  | 
text{* Correctness: *}
 | 
| 47613 | 212  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
213  | 
lemma aval'_correct: "s : \<gamma>\<^sub>s S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"  | 
| 47613 | 214  | 
by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_fun_def)  | 
215  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
216  | 
lemma in_gamma_update: "\<lbrakk> s : \<gamma>\<^sub>s S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^sub>s(S(x := a))"  | 
| 47613 | 217  | 
by(simp add: \<gamma>_fun_def)  | 
218  | 
||
| 51390 | 219  | 
lemma gamma_Step_subcomm:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
220  | 
assumes "!!x e S. f1 x e (\<gamma>\<^sub>o S) \<subseteq> \<gamma>\<^sub>o (f2 x e S)" "!!b S. g1 b (\<gamma>\<^sub>o S) \<subseteq> \<gamma>\<^sub>o (g2 b S)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
221  | 
shows "Step f1 g1 (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (Step f2 g2 S C)"  | 
| 54944 | 222  | 
by (induction C arbitrary: S) (auto simp: mono_gamma_o assms)  | 
| 51390 | 223  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
224  | 
lemma step_step': "step (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' S C)"  | 
| 51390 | 225  | 
unfolding step_def step'_def  | 
| 51694 | 226  | 
by(rule gamma_Step_subcomm)  | 
| 51974 | 227  | 
(auto simp: aval'_correct in_gamma_update asem_def split: option.splits)  | 
| 47613 | 228  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
229  | 
lemma AI_correct: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c C"  | 
| 47613 | 230  | 
proof(simp add: CS_def AI_def)  | 
| 49464 | 231  | 
assume 1: "pfp (step' \<top>) (bot c) = Some C"  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
232  | 
have pfp': "step' \<top> C \<le> C" by(rule pfp_pfp[OF 1])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
233  | 
have 2: "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c C" --"transfer the pfp'"  | 
| 50986 | 234  | 
proof(rule order_trans)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
235  | 
show "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' \<top> C)" by(rule step_step')  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
236  | 
show "... \<le> \<gamma>\<^sub>c C" by (metis mono_gamma_c[OF pfp'])  | 
| 47613 | 237  | 
qed  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
238  | 
have 3: "strip (\<gamma>\<^sub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
239  | 
have "lfp c (step (\<gamma>\<^sub>o \<top>)) \<le> \<gamma>\<^sub>c C"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
240  | 
by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^sub>o \<top>)", OF 3 2])  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
241  | 
thus "lfp c (step UNIV) \<le> \<gamma>\<^sub>c C" by simp  | 
| 47613 | 242  | 
qed  | 
243  | 
||
244  | 
end  | 
|
245  | 
||
246  | 
||
247  | 
subsubsection "Monotonicity"  | 
|
248  | 
||
| 51721 | 249  | 
locale Abs_Int_fun_mono = Abs_Int_fun +  | 
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
250  | 
assumes mono_plus': "a1 \<le> b1 \<Longrightarrow> a2 \<le> b2 \<Longrightarrow> plus' a1 a2 \<le> plus' b1 b2"  | 
| 47613 | 251  | 
begin  | 
252  | 
||
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
253  | 
lemma mono_aval': "S \<le> S' \<Longrightarrow> aval' e S \<le> aval' e S'"  | 
| 47613 | 254  | 
by(induction e)(auto simp: le_fun_def mono_plus')  | 
255  | 
||
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
256  | 
lemma mono_update: "a \<le> a' \<Longrightarrow> S \<le> S' \<Longrightarrow> S(x := a) \<le> S'(x := a')"  | 
| 47613 | 257  | 
by(simp add: le_fun_def)  | 
258  | 
||
| 
51359
 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 
nipkow 
parents: 
51036 
diff
changeset
 | 
259  | 
lemma mono_step': "S1 \<le> S2 \<Longrightarrow> C1 \<le> C2 \<Longrightarrow> step' S1 C1 \<le> step' S2 C2"  | 
| 51390 | 260  | 
unfolding step'_def  | 
| 51694 | 261  | 
by(rule mono2_Step)  | 
| 51807 | 262  | 
(auto simp: mono_update mono_aval' asem_def split: option.split)  | 
| 47613 | 263  | 
|
| 51722 | 264  | 
lemma mono_step'_top: "C \<le> C' \<Longrightarrow> step' \<top> C \<le> step' \<top> C'"  | 
265  | 
by (metis mono_step' order_refl)  | 
|
266  | 
||
267  | 
lemma AI_least_pfp: assumes "AI c = Some C" "step' \<top> C' \<le> C'" "strip C' = c"  | 
|
268  | 
shows "C \<le> C'"  | 
|
269  | 
by(rule pfp_bot_least[OF _ _ assms(2,3) assms(1)[unfolded AI_def]])  | 
|
270  | 
(simp_all add: mono_step'_top)  | 
|
271  | 
||
272  | 
end  | 
|
273  | 
||
274  | 
||
275  | 
instantiation acom :: (type) vars  | 
|
276  | 
begin  | 
|
277  | 
||
278  | 
definition "vars_acom = vars o strip"  | 
|
279  | 
||
280  | 
instance ..  | 
|
281  | 
||
282  | 
end  | 
|
283  | 
||
284  | 
lemma finite_Cvars: "finite(vars(C::'a acom))"  | 
|
285  | 
by(simp add: vars_acom_def)  | 
|
286  | 
||
287  | 
||
288  | 
subsubsection "Termination"  | 
|
289  | 
||
290  | 
lemma pfp_termination:  | 
|
291  | 
fixes x0 :: "'a::order" and m :: "'a \<Rightarrow> nat"  | 
|
292  | 
assumes mono: "\<And>x y. I x \<Longrightarrow> I y \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"  | 
|
293  | 
and m: "\<And>x y. I x \<Longrightarrow> I y \<Longrightarrow> x < y \<Longrightarrow> m x > m y"  | 
|
294  | 
and I: "\<And>x y. I x \<Longrightarrow> I(f x)" and "I x0" and "x0 \<le> f x0"  | 
|
295  | 
shows "\<exists>x. pfp f x0 = Some x"  | 
|
296  | 
proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. I x & x \<le> f x"])  | 
|
297  | 
  show "wf {(y,x). ((I x \<and> x \<le> f x) \<and> \<not> f x \<le> x) \<and> y = f x}"
 | 
|
298  | 
by(rule wf_subset[OF wf_measure[of m]]) (auto simp: m I)  | 
|
299  | 
next  | 
|
300  | 
show "I x0 \<and> x0 \<le> f x0" using `I x0` `x0 \<le> f x0` by blast  | 
|
301  | 
next  | 
|
302  | 
fix x assume "I x \<and> x \<le> f x" thus "I(f x) \<and> f x \<le> f(f x)"  | 
|
303  | 
by (blast intro: I mono)  | 
|
304  | 
qed  | 
|
305  | 
||
| 
52019
 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 
nipkow 
parents: 
51974 
diff
changeset
 | 
306  | 
lemma le_iff_le_annos: "C1 \<le> C2 \<longleftrightarrow>  | 
| 
 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 
nipkow 
parents: 
51974 
diff
changeset
 | 
307  | 
strip C1 = strip C2 \<and> (\<forall> i<size(annos C1). annos C1 ! i \<le> annos C2 ! i)"  | 
| 
 
a4cbca8f7342
finally: acom with pointwise access and update of annotations
 
nipkow 
parents: 
51974 
diff
changeset
 | 
308  | 
by(simp add: less_eq_acom_def anno_def)  | 
| 51722 | 309  | 
|
310  | 
locale Measure1_fun =  | 
|
| 
51749
 
c27bb7994bd3
moved defs into locale to reduce unnecessary polymorphism; tuned
 
nipkow 
parents: 
51722 
diff
changeset
 | 
311  | 
fixes m :: "'av::top \<Rightarrow> nat"  | 
| 51722 | 312  | 
fixes h :: "nat"  | 
313  | 
assumes h: "m x \<le> h"  | 
|
314  | 
begin  | 
|
315  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
316  | 
definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>s") where
 | 
| 51791 | 317  | 
"m_s S X = (\<Sum> x \<in> X. m(S x))"  | 
| 51722 | 318  | 
|
| 51791 | 319  | 
lemma m_s_h: "finite X \<Longrightarrow> m_s S X \<le> h * card X"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
54944 
diff
changeset
 | 
320  | 
by(simp add: m_s_def) (metis mult.commute of_nat_id setsum_bounded[OF h])  | 
| 51722 | 321  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
322  | 
fun m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>o") where
 | 
| 51791 | 323  | 
"m_o (Some S) X = m_s S X" |  | 
324  | 
"m_o None X = h * card X + 1"  | 
|
| 51722 | 325  | 
|
| 51791 | 326  | 
lemma m_o_h: "finite X \<Longrightarrow> m_o opt X \<le> (h*card X + 1)"  | 
| 
51749
 
c27bb7994bd3
moved defs into locale to reduce unnecessary polymorphism; tuned
 
nipkow 
parents: 
51722 
diff
changeset
 | 
327  | 
by(cases opt)(auto simp add: m_s_h le_SucI dest: m_s_h)  | 
| 51722 | 328  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
329  | 
definition m_c :: "'av st option acom \<Rightarrow> nat" ("m\<^sub>c") where
 | 
| 51791 | 330  | 
"m_c C = listsum (map (\<lambda>a. m_o a (vars C)) (annos C))"  | 
| 51722 | 331  | 
|
332  | 
text{* Upper complexity bound: *}
 | 
|
333  | 
lemma m_c_h: "m_c C \<le> size(annos C) * (h * card(vars C) + 1)"  | 
|
334  | 
proof-  | 
|
335  | 
let ?X = "vars C" let ?n = "card ?X" let ?a = "size(annos C)"  | 
|
| 51791 | 336  | 
have "m_c C = (\<Sum>i<?a. m_o (annos C ! i) ?X)"  | 
| 51783 | 337  | 
by(simp add: m_c_def listsum_setsum_nth atLeast0LessThan)  | 
| 51722 | 338  | 
also have "\<dots> \<le> (\<Sum>i<?a. h * ?n + 1)"  | 
339  | 
apply(rule setsum_mono) using m_o_h[OF finite_Cvars] by simp  | 
|
340  | 
also have "\<dots> = ?a * (h * ?n + 1)" by simp  | 
|
341  | 
finally show ?thesis .  | 
|
342  | 
qed  | 
|
343  | 
||
344  | 
end  | 
|
345  | 
||
346  | 
||
| 51826 | 347  | 
locale Measure_fun = Measure1_fun where m=m  | 
348  | 
for m :: "'av::semilattice_sup_top \<Rightarrow> nat" +  | 
|
| 
51749
 
c27bb7994bd3
moved defs into locale to reduce unnecessary polymorphism; tuned
 
nipkow 
parents: 
51722 
diff
changeset
 | 
349  | 
assumes m2: "x < y \<Longrightarrow> m x > m y"  | 
| 
 
c27bb7994bd3
moved defs into locale to reduce unnecessary polymorphism; tuned
 
nipkow 
parents: 
51722 
diff
changeset
 | 
350  | 
begin  | 
| 
 
c27bb7994bd3
moved defs into locale to reduce unnecessary polymorphism; tuned
 
nipkow 
parents: 
51722 
diff
changeset
 | 
351  | 
|
| 51785 | 352  | 
text{* The predicates @{text "top_on_ty a X"} that follow describe that any abstract
 | 
| 
51749
 
c27bb7994bd3
moved defs into locale to reduce unnecessary polymorphism; tuned
 
nipkow 
parents: 
51722 
diff
changeset
 | 
353  | 
state in @{text a} maps all variables in @{text X} to @{term \<top>}.
 | 
| 51722 | 354  | 
This is an important invariant for the termination proof where we argue that only  | 
355  | 
the finitely many variables in the program change. That the others do not change  | 
|
356  | 
follows because they remain @{term \<top>}. *}
 | 
|
357  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
358  | 
fun top_on_st :: "'av st \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>s") where
 | 
| 51785 | 359  | 
"top_on_st S X = (\<forall>x\<in>X. S x = \<top>)"  | 
| 51722 | 360  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
361  | 
fun top_on_opt :: "'av st option \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>o") where
 | 
| 51785 | 362  | 
"top_on_opt (Some S) X = top_on_st S X" |  | 
363  | 
"top_on_opt None X = True"  | 
|
| 51722 | 364  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52729 
diff
changeset
 | 
365  | 
definition top_on_acom :: "'av st option acom \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>c") where
 | 
| 51785 | 366  | 
"top_on_acom C X = (\<forall>a \<in> set(annos C). top_on_opt a X)"  | 
| 51722 | 367  | 
|
| 51785 | 368  | 
lemma top_on_top: "top_on_opt \<top> X"  | 
| 51722 | 369  | 
by(auto simp: top_option_def)  | 
370  | 
||
| 51785 | 371  | 
lemma top_on_bot: "top_on_acom (bot c) X"  | 
| 51722 | 372  | 
by(auto simp add: top_on_acom_def bot_def)  | 
373  | 
||
| 51785 | 374  | 
lemma top_on_post: "top_on_acom C X \<Longrightarrow> top_on_opt (post C) X"  | 
| 51722 | 375  | 
by(simp add: top_on_acom_def post_in_annos)  | 
376  | 
||
377  | 
lemma top_on_acom_simps:  | 
|
| 51785 | 378  | 
  "top_on_acom (SKIP {Q}) X = top_on_opt Q X"
 | 
379  | 
  "top_on_acom (x ::= e {Q}) X = top_on_opt Q X"
 | 
|
| 
52046
 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 
nipkow 
parents: 
52022 
diff
changeset
 | 
380  | 
"top_on_acom (C1;;C2) X = (top_on_acom C1 X \<and> top_on_acom C2 X)"  | 
| 51785 | 381  | 
  "top_on_acom (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) X =
 | 
382  | 
(top_on_opt P1 X \<and> top_on_acom C1 X \<and> top_on_opt P2 X \<and> top_on_acom C2 X \<and> top_on_opt Q X)"  | 
|
383  | 
  "top_on_acom ({I} WHILE b DO {P} C {Q}) X =
 | 
|
384  | 
(top_on_opt I X \<and> top_on_acom C X \<and> top_on_opt P X \<and> top_on_opt Q X)"  | 
|
| 51722 | 385  | 
by(auto simp add: top_on_acom_def)  | 
386  | 
||
387  | 
lemma top_on_sup:  | 
|
| 51785 | 388  | 
"top_on_opt o1 X \<Longrightarrow> top_on_opt o2 X \<Longrightarrow> top_on_opt (o1 \<squnion> o2) X"  | 
| 51722 | 389  | 
apply(induction o1 o2 rule: sup_option.induct)  | 
390  | 
apply(auto)  | 
|
391  | 
done  | 
|
392  | 
||
| 
51749
 
c27bb7994bd3
moved defs into locale to reduce unnecessary polymorphism; tuned
 
nipkow 
parents: 
51722 
diff
changeset
 | 
393  | 
lemma top_on_Step: fixes C :: "'av st option acom"  | 
| 51785 | 394  | 
assumes "!!x e S. \<lbrakk>top_on_opt S X; x \<notin> X; vars e \<subseteq> -X\<rbrakk> \<Longrightarrow> top_on_opt (f x e S) X"  | 
395  | 
"!!b S. top_on_opt S X \<Longrightarrow> vars b \<subseteq> -X \<Longrightarrow> top_on_opt (g b S) X"  | 
|
396  | 
shows "\<lbrakk> vars C \<subseteq> -X; top_on_opt S X; top_on_acom C X \<rbrakk> \<Longrightarrow> top_on_acom (Step f g S C) X"  | 
|
| 51722 | 397  | 
proof(induction C arbitrary: S)  | 
398  | 
qed (auto simp: top_on_acom_simps vars_acom_def top_on_post top_on_sup assms)  | 
|
399  | 
||
400  | 
lemma m1: "x \<le> y \<Longrightarrow> m x \<ge> m y"  | 
|
401  | 
by(auto simp: le_less m2)  | 
|
402  | 
||
403  | 
lemma m_s2_rep: assumes "finite(X)" and "S1 = S2 on -X" and "\<forall>x. S1 x \<le> S2 x" and "S1 \<noteq> S2"  | 
|
404  | 
shows "(\<Sum>x\<in>X. m (S2 x)) < (\<Sum>x\<in>X. m (S1 x))"  | 
|
405  | 
proof-  | 
|
406  | 
from assms(3) have 1: "\<forall>x\<in>X. m(S1 x) \<ge> m(S2 x)" by (simp add: m1)  | 
|
407  | 
from assms(2,3,4) have "EX x:X. S1 x < S2 x"  | 
|
408  | 
by(simp add: fun_eq_iff) (metis Compl_iff le_neq_trans)  | 
|
409  | 
hence 2: "\<exists>x\<in>X. m(S1 x) > m(S2 x)" by (metis m2)  | 
|
410  | 
from setsum_strict_mono_ex1[OF `finite X` 1 2]  | 
|
411  | 
show "(\<Sum>x\<in>X. m (S2 x)) < (\<Sum>x\<in>X. m (S1 x))" .  | 
|
412  | 
qed  | 
|
413  | 
||
| 51791 | 414  | 
lemma m_s2: "finite(X) \<Longrightarrow> S1 = S2 on -X \<Longrightarrow> S1 < S2 \<Longrightarrow> m_s S1 X > m_s S2 X"  | 
| 51722 | 415  | 
apply(auto simp add: less_fun_def m_s_def)  | 
416  | 
apply(simp add: m_s2_rep le_fun_def)  | 
|
417  | 
done  | 
|
418  | 
||
| 51785 | 419  | 
lemma m_o2: "finite X \<Longrightarrow> top_on_opt o1 (-X) \<Longrightarrow> top_on_opt o2 (-X) \<Longrightarrow>  | 
| 51791 | 420  | 
o1 < o2 \<Longrightarrow> m_o o1 X > m_o o2 X"  | 
| 51722 | 421  | 
proof(induction o1 o2 rule: less_eq_option.induct)  | 
| 
51749
 
c27bb7994bd3
moved defs into locale to reduce unnecessary polymorphism; tuned
 
nipkow 
parents: 
51722 
diff
changeset
 | 
422  | 
case 1 thus ?case by (auto simp: m_s2 less_option_def)  | 
| 51722 | 423  | 
next  | 
| 
51749
 
c27bb7994bd3
moved defs into locale to reduce unnecessary polymorphism; tuned
 
nipkow 
parents: 
51722 
diff
changeset
 | 
424  | 
case 2 thus ?case by(auto simp: less_option_def le_imp_less_Suc m_s_h)  | 
| 51722 | 425  | 
next  | 
426  | 
case 3 thus ?case by (auto simp: less_option_def)  | 
|
427  | 
qed  | 
|
428  | 
||
| 51785 | 429  | 
lemma m_o1: "finite X \<Longrightarrow> top_on_opt o1 (-X) \<Longrightarrow> top_on_opt o2 (-X) \<Longrightarrow>  | 
| 51791 | 430  | 
o1 \<le> o2 \<Longrightarrow> m_o o1 X \<ge> m_o o2 X"  | 
| 51722 | 431  | 
by(auto simp: le_less m_o2)  | 
432  | 
||
433  | 
||
| 51785 | 434  | 
lemma m_c2: "top_on_acom C1 (-vars C1) \<Longrightarrow> top_on_acom C2 (-vars C2) \<Longrightarrow>  | 
| 51722 | 435  | 
C1 < C2 \<Longrightarrow> m_c C1 > m_c C2"  | 
| 51783 | 436  | 
proof(auto simp add: le_iff_le_annos size_annos_same[of C1 C2] vars_acom_def less_acom_def)  | 
| 51722 | 437  | 
let ?X = "vars(strip C2)"  | 
| 51785 | 438  | 
assume top: "top_on_acom C1 (- vars(strip C2))" "top_on_acom C2 (- vars(strip C2))"  | 
| 51722 | 439  | 
and strip_eq: "strip C1 = strip C2"  | 
440  | 
and 0: "\<forall>i<size(annos C2). annos C1 ! i \<le> annos C2 ! i"  | 
|
| 51791 | 441  | 
hence 1: "\<forall>i<size(annos C2). m_o (annos C1 ! i) ?X \<ge> m_o (annos C2 ! i) ?X"  | 
| 51722 | 442  | 
apply (auto simp: all_set_conv_all_nth vars_acom_def top_on_acom_def)  | 
443  | 
by (metis (lifting, no_types) finite_cvars m_o1 size_annos_same2)  | 
|
444  | 
fix i assume i: "i < size(annos C2)" "\<not> annos C2 ! i \<le> annos C1 ! i"  | 
|
| 51785 | 445  | 
have topo1: "top_on_opt (annos C1 ! i) (- ?X)"  | 
| 51722 | 446  | 
using i(1) top(1) by(simp add: top_on_acom_def size_annos_same[OF strip_eq])  | 
| 51785 | 447  | 
have topo2: "top_on_opt (annos C2 ! i) (- ?X)"  | 
| 51722 | 448  | 
using i(1) top(2) by(simp add: top_on_acom_def size_annos_same[OF strip_eq])  | 
| 51791 | 449  | 
from i have "m_o (annos C1 ! i) ?X > m_o (annos C2 ! i) ?X" (is "?P i")  | 
| 51722 | 450  | 
by (metis 0 less_option_def m_o2[OF finite_cvars topo1] topo2)  | 
451  | 
hence 2: "\<exists>i < size(annos C2). ?P i" using `i < size(annos C2)` by blast  | 
|
| 51791 | 452  | 
have "(\<Sum>i<size(annos C2). m_o (annos C2 ! i) ?X)  | 
453  | 
< (\<Sum>i<size(annos C2). m_o (annos C1 ! i) ?X)"  | 
|
| 51722 | 454  | 
apply(rule setsum_strict_mono_ex1) using 1 2 by (auto)  | 
| 51783 | 455  | 
thus ?thesis  | 
456  | 
by(simp add: m_c_def vars_acom_def strip_eq listsum_setsum_nth atLeast0LessThan size_annos_same[OF strip_eq])  | 
|
| 51722 | 457  | 
qed  | 
458  | 
||
459  | 
end  | 
|
460  | 
||
461  | 
||
462  | 
locale Abs_Int_fun_measure =  | 
|
463  | 
Abs_Int_fun_mono where \<gamma>=\<gamma> + Measure_fun where m=m  | 
|
| 51826 | 464  | 
for \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set" and m :: "'av \<Rightarrow> nat"  | 
| 51722 | 465  | 
begin  | 
466  | 
||
| 51785 | 467  | 
lemma top_on_step': "top_on_acom C (-vars C) \<Longrightarrow> top_on_acom (step' \<top> C) (-vars C)"  | 
| 51722 | 468  | 
unfolding step'_def  | 
469  | 
by(rule top_on_Step)  | 
|
| 51807 | 470  | 
(auto simp add: top_option_def asem_def split: option.splits)  | 
| 51722 | 471  | 
|
472  | 
lemma AI_Some_measure: "\<exists>C. AI c = Some C"  | 
|
473  | 
unfolding AI_def  | 
|
| 51785 | 474  | 
apply(rule pfp_termination[where I = "\<lambda>C. top_on_acom C (- vars C)" and m="m_c"])  | 
| 51722 | 475  | 
apply(simp_all add: m_c2 mono_step'_top bot_least top_on_bot)  | 
| 51754 | 476  | 
using top_on_step' apply(auto simp add: vars_acom_def)  | 
| 51722 | 477  | 
done  | 
478  | 
||
| 47613 | 479  | 
end  | 
480  | 
||
481  | 
text{* Problem: not executable because of the comparison of abstract states,
 | 
|
| 52022 | 482  | 
i.e. functions, in the pre-fixpoint computation. *}  | 
| 47613 | 483  | 
|
484  | 
end  |