| author | nipkow |
| Wed, 17 Apr 2013 21:23:35 +0200 | |
| changeset 51712 | 30624dab6054 |
| parent 51711 | df3426139651 |
| child 51722 | 3da99469cc1b |
| permissions | -rw-r--r-- |
| 47613 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
theory Abs_Int3 |
|
4 |
imports Abs_Int2_ivl |
|
5 |
begin |
|
6 |
||
7 |
||
8 |
subsection "Widening and Narrowing" |
|
9 |
||
10 |
class widen = |
|
11 |
fixes widen :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<nabla>" 65) |
|
12 |
||
13 |
class narrow = |
|
14 |
fixes narrow :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<triangle>" 65) |
|
15 |
||
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
16 |
class WN = widen + narrow + order + |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
17 |
assumes widen1: "x \<le> x \<nabla> y" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
18 |
assumes widen2: "y \<le> x \<nabla> y" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
19 |
assumes narrow1: "y \<le> x \<Longrightarrow> y \<le> x \<triangle> y" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
20 |
assumes narrow2: "y \<le> x \<Longrightarrow> x \<triangle> y \<le> x" |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
21 |
begin |
| 47613 | 22 |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
23 |
lemma narrowid[simp]: "x \<triangle> x = x" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
24 |
by (metis eq_iff narrow1 narrow2) |
| 47613 | 25 |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
26 |
end |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
27 |
|
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
28 |
lemma top_widen_top[simp]: "\<top> \<nabla> \<top> = (\<top>::_::{WN,top})"
|
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
29 |
by (metis eq_iff top_greatest widen2) |
| 47613 | 30 |
|
31 |
instantiation ivl :: WN |
|
32 |
begin |
|
33 |
||
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
34 |
definition "widen_rep p1 p2 = |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
35 |
(if is_empty_rep p1 then p2 else if is_empty_rep p2 then p1 else |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
36 |
let (l1,h1) = p1; (l2,h2) = p2 |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
37 |
in (if l2 < l1 then Minf else l1, if h1 < h2 then Pinf else h1))" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
38 |
|
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
39 |
lift_definition widen_ivl :: "ivl \<Rightarrow> ivl \<Rightarrow> ivl" is widen_rep |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
40 |
by(auto simp: widen_rep_def eq_ivl_iff) |
| 47613 | 41 |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
42 |
definition "narrow_rep p1 p2 = |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
43 |
(if is_empty_rep p1 \<or> is_empty_rep p2 then empty_rep else |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
44 |
let (l1,h1) = p1; (l2,h2) = p2 |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
45 |
in (if l1 = Minf then l2 else l1, if h1 = Pinf then h2 else h1))" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
46 |
|
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
47 |
lift_definition narrow_ivl :: "ivl \<Rightarrow> ivl \<Rightarrow> ivl" is narrow_rep |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
48 |
by(auto simp: narrow_rep_def eq_ivl_iff) |
| 47613 | 49 |
|
50 |
instance |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
51 |
proof |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
52 |
qed (transfer, auto simp: widen_rep_def narrow_rep_def le_iff_subset \<gamma>_rep_def subset_eq is_empty_rep_def empty_rep_def eq_ivl_def split: if_splits extended.splits)+ |
| 47613 | 53 |
|
54 |
end |
|
55 |
||
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
56 |
instantiation st :: ("{top,WN}")WN
|
| 47613 | 57 |
begin |
58 |
||
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
59 |
lift_definition widen_st :: "'a st \<Rightarrow> 'a st \<Rightarrow> 'a st" is "map2_st_rep (op \<nabla>)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
60 |
by(auto simp: eq_st_def) |
| 47613 | 61 |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
62 |
lift_definition narrow_st :: "'a st \<Rightarrow> 'a st \<Rightarrow> 'a st" is "map2_st_rep (op \<triangle>)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
63 |
by(auto simp: eq_st_def) |
| 47613 | 64 |
|
65 |
instance |
|
66 |
proof |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
67 |
case goal1 thus ?case |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
68 |
by transfer (simp add: less_eq_st_rep_iff widen1) |
| 47613 | 69 |
next |
70 |
case goal2 thus ?case |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
71 |
by transfer (simp add: less_eq_st_rep_iff widen2) |
| 47613 | 72 |
next |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
73 |
case goal3 thus ?case |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
74 |
by transfer (simp add: less_eq_st_rep_iff narrow1) |
| 47613 | 75 |
next |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
76 |
case goal4 thus ?case |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
77 |
by transfer (simp add: less_eq_st_rep_iff narrow2) |
| 47613 | 78 |
qed |
79 |
||
80 |
end |
|
81 |
||
82 |
||
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
83 |
instantiation option :: (WN)WN |
| 47613 | 84 |
begin |
85 |
||
86 |
fun widen_option where |
|
87 |
"None \<nabla> x = x" | |
|
88 |
"x \<nabla> None = x" | |
|
89 |
"(Some x) \<nabla> (Some y) = Some(x \<nabla> y)" |
|
90 |
||
91 |
fun narrow_option where |
|
92 |
"None \<triangle> x = None" | |
|
93 |
"x \<triangle> None = None" | |
|
94 |
"(Some x) \<triangle> (Some y) = Some(x \<triangle> y)" |
|
95 |
||
96 |
instance |
|
97 |
proof |
|
98 |
case goal1 thus ?case |
|
99 |
by(induct x y rule: widen_option.induct)(simp_all add: widen1) |
|
100 |
next |
|
101 |
case goal2 thus ?case |
|
102 |
by(induct x y rule: widen_option.induct)(simp_all add: widen2) |
|
103 |
next |
|
104 |
case goal3 thus ?case |
|
105 |
by(induct x y rule: narrow_option.induct) (simp_all add: narrow1) |
|
106 |
next |
|
107 |
case goal4 thus ?case |
|
108 |
by(induct x y rule: narrow_option.induct) (simp_all add: narrow2) |
|
109 |
qed |
|
110 |
||
111 |
end |
|
112 |
||
113 |
||
114 |
fun map2_acom :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a acom \<Rightarrow> 'a acom \<Rightarrow> 'a acom" where
|
|
115 |
"map2_acom f (SKIP {a1}) (SKIP {a2}) = (SKIP {f a1 a2})" |
|
|
116 |
"map2_acom f (x ::= e {a1}) (x' ::= e' {a2}) = (x ::= e {f a1 a2})" |
|
|
117 |
"map2_acom f (C1;C2) (D1;D2) = (map2_acom f C1 D1; map2_acom f C2 D2)" | |
|
| 49095 | 118 |
"map2_acom f (IF b THEN {p1} C1 ELSE {p2} C2 {a1}) (IF b' THEN {q1} D1 ELSE {q2} D2 {a2}) =
|
119 |
(IF b THEN {f p1 q1} map2_acom f C1 D1 ELSE {f p2 q2} map2_acom f C2 D2 {f a1 a2})" |
|
|
120 |
"map2_acom f ({a1} WHILE b DO {p} C {a2}) ({a3} WHILE b' DO {p'} C' {a4}) =
|
|
121 |
({f a1 a3} WHILE b DO {f p p'} map2_acom f C C' {f a2 a4})"
|
|
| 47613 | 122 |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
123 |
lemma annos_map2_acom[simp]: "strip C2 = strip C1 \<Longrightarrow> |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
124 |
annos(map2_acom f C1 C2) = map (%(x,y).f x y) (zip (annos C1) (annos C2))" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
125 |
by(induction f C1 C2 rule: map2_acom.induct)(simp_all add: size_annos_same2) |
| 49548 | 126 |
|
127 |
instantiation acom :: (widen)widen |
|
128 |
begin |
|
129 |
definition "widen_acom = map2_acom (op \<nabla>)" |
|
130 |
instance .. |
|
131 |
end |
|
132 |
||
133 |
instantiation acom :: (narrow)narrow |
|
134 |
begin |
|
135 |
definition "narrow_acom = map2_acom (op \<triangle>)" |
|
136 |
instance .. |
|
137 |
end |
|
138 |
||
| 47613 | 139 |
lemma strip_map2_acom[simp]: |
140 |
"strip C1 = strip C2 \<Longrightarrow> strip(map2_acom f C1 C2) = strip C1" |
|
141 |
by(induct f C1 C2 rule: map2_acom.induct) simp_all |
|
142 |
||
143 |
lemma strip_widen_acom[simp]: |
|
144 |
"strip C1 = strip C2 \<Longrightarrow> strip(C1 \<nabla> C2) = strip C1" |
|
| 49548 | 145 |
by(simp add: widen_acom_def) |
| 47613 | 146 |
|
147 |
lemma strip_narrow_acom[simp]: |
|
148 |
"strip C1 = strip C2 \<Longrightarrow> strip(C1 \<triangle> C2) = strip C1" |
|
| 49548 | 149 |
by(simp add: narrow_acom_def) |
| 47613 | 150 |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
151 |
lemma narrow1_acom: "C2 \<le> C1 \<Longrightarrow> C2 \<le> C1 \<triangle> (C2::'a::WN acom)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
152 |
by(induct C1 C2 rule: less_eq_acom.induct)(simp_all add: narrow_acom_def narrow1) |
| 47613 | 153 |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
154 |
lemma narrow2_acom: "C2 \<le> C1 \<Longrightarrow> C1 \<triangle> (C2::'a::WN acom) \<le> C1" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
155 |
by(induct C1 C2 rule: less_eq_acom.induct)(simp_all add: narrow_acom_def narrow2) |
| 47613 | 156 |
|
157 |
||
158 |
subsubsection "Post-fixed point computation" |
|
159 |
||
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
160 |
definition iter_widen :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> ('a::{order,widen})option"
|
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
161 |
where "iter_widen f = while_option (\<lambda>x. \<not> f x \<le> x) (\<lambda>x. x \<nabla> f x)" |
| 47613 | 162 |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
163 |
definition iter_narrow :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> ('a::{order,narrow})option"
|
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
164 |
where "iter_narrow f = while_option (\<lambda>x. x \<triangle> f x < x) (\<lambda>x. x \<triangle> f x)" |
| 47613 | 165 |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
166 |
definition pfp_wn :: "('a::{order,widen,narrow} \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option"
|
| 49548 | 167 |
where "pfp_wn f x = |
| 49576 | 168 |
(case iter_widen f x of None \<Rightarrow> None | Some p \<Rightarrow> iter_narrow f p)" |
| 47613 | 169 |
|
170 |
||
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
171 |
lemma iter_widen_pfp: "iter_widen f x = Some p \<Longrightarrow> f p \<le> p" |
| 47613 | 172 |
by(auto simp add: iter_widen_def dest: while_option_stop) |
173 |
||
174 |
lemma iter_widen_inv: |
|
175 |
assumes "!!x. P x \<Longrightarrow> P(f x)" "!!x1 x2. P x1 \<Longrightarrow> P x2 \<Longrightarrow> P(x1 \<nabla> x2)" and "P x" |
|
176 |
and "iter_widen f x = Some y" shows "P y" |
|
177 |
using while_option_rule[where P = "P", OF _ assms(4)[unfolded iter_widen_def]] |
|
178 |
by (blast intro: assms(1-3)) |
|
179 |
||
180 |
lemma strip_while: fixes f :: "'a acom \<Rightarrow> 'a acom" |
|
181 |
assumes "\<forall>C. strip (f C) = strip C" and "while_option P f C = Some C'" |
|
182 |
shows "strip C' = strip C" |
|
183 |
using while_option_rule[where P = "\<lambda>C'. strip C' = strip C", OF _ assms(2)] |
|
184 |
by (metis assms(1)) |
|
185 |
||
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
186 |
lemma strip_iter_widen: fixes f :: "'a::{order,widen} acom \<Rightarrow> 'a acom"
|
| 47613 | 187 |
assumes "\<forall>C. strip (f C) = strip C" and "iter_widen f C = Some C'" |
188 |
shows "strip C' = strip C" |
|
189 |
proof- |
|
190 |
have "\<forall>C. strip(C \<nabla> f C) = strip C" |
|
191 |
by (metis assms(1) strip_map2_acom widen_acom_def) |
|
192 |
from strip_while[OF this] assms(2) show ?thesis by(simp add: iter_widen_def) |
|
193 |
qed |
|
194 |
||
195 |
lemma iter_narrow_pfp: |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
196 |
assumes mono: "!!x1 x2::_::WN acom. P x1 \<Longrightarrow> P x2 \<Longrightarrow> x1 \<le> x2 \<Longrightarrow> f x1 \<le> f x2" |
| 49576 | 197 |
and Pinv: "!!x. P x \<Longrightarrow> P(f x)" "!!x1 x2. P x1 \<Longrightarrow> P x2 \<Longrightarrow> P(x1 \<triangle> x2)" |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
198 |
and "P p0" and "f p0 \<le> p0" and "iter_narrow f p0 = Some p" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
199 |
shows "P p \<and> f p \<le> p" |
| 47613 | 200 |
proof- |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
201 |
let ?Q = "%p. P p \<and> f p \<le> p \<and> p \<le> p0" |
| 49576 | 202 |
{ fix p assume "?Q p"
|
| 47613 | 203 |
note P = conjunct1[OF this] and 12 = conjunct2[OF this] |
204 |
note 1 = conjunct1[OF 12] and 2 = conjunct2[OF 12] |
|
| 49576 | 205 |
let ?p' = "p \<triangle> f p" |
206 |
have "?Q ?p'" |
|
| 47613 | 207 |
proof auto |
| 49576 | 208 |
show "P ?p'" by (blast intro: P Pinv) |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
209 |
have "f ?p' \<le> f p" by(rule mono[OF `P (p \<triangle> f p)` P narrow2_acom[OF 1]]) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
210 |
also have "\<dots> \<le> ?p'" by(rule narrow1_acom[OF 1]) |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
211 |
finally show "f ?p' \<le> ?p'" . |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
212 |
have "?p' \<le> p" by (rule narrow2_acom[OF 1]) |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
213 |
also have "p \<le> p0" by(rule 2) |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
214 |
finally show "?p' \<le> p0" . |
| 47613 | 215 |
qed |
216 |
} |
|
217 |
thus ?thesis |
|
218 |
using while_option_rule[where P = ?Q, OF _ assms(6)[simplified iter_narrow_def]] |
|
219 |
by (blast intro: assms(4,5) le_refl) |
|
220 |
qed |
|
221 |
||
222 |
lemma pfp_wn_pfp: |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
223 |
assumes mono: "!!x1 x2::_::WN acom. P x1 \<Longrightarrow> P x2 \<Longrightarrow> x1 \<le> x2 \<Longrightarrow> f x1 \<le> f x2" |
| 49548 | 224 |
and Pinv: "P x" "!!x. P x \<Longrightarrow> P(f x)" |
225 |
"!!x1 x2. P x1 \<Longrightarrow> P x2 \<Longrightarrow> P(x1 \<nabla> x2)" |
|
226 |
"!!x1 x2. P x1 \<Longrightarrow> P x2 \<Longrightarrow> P(x1 \<triangle> x2)" |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
227 |
and pfp_wn: "pfp_wn f x = Some p" shows "P p \<and> f p \<le> p" |
| 47613 | 228 |
proof- |
| 49576 | 229 |
from pfp_wn obtain p0 |
230 |
where its: "iter_widen f x = Some p0" "iter_narrow f p0 = Some p" |
|
| 47613 | 231 |
by(auto simp: pfp_wn_def split: option.splits) |
| 49576 | 232 |
have "P p0" by (blast intro: iter_widen_inv[where P="P"] its(1) Pinv(1-3)) |
| 47613 | 233 |
thus ?thesis |
234 |
by - (assumption | |
|
235 |
rule iter_narrow_pfp[where P=P] mono Pinv(2,4) iter_widen_pfp its)+ |
|
236 |
qed |
|
237 |
||
238 |
lemma strip_pfp_wn: |
|
| 49548 | 239 |
"\<lbrakk> \<forall>C. strip(f C) = strip C; pfp_wn f C = Some C' \<rbrakk> \<Longrightarrow> strip C' = strip C" |
| 47613 | 240 |
by(auto simp add: pfp_wn_def iter_narrow_def split: option.splits) |
| 51390 | 241 |
(metis (mono_tags) strip_iter_widen strip_narrow_acom strip_while) |
| 47613 | 242 |
|
243 |
||
244 |
locale Abs_Int2 = Abs_Int1_mono |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
245 |
where \<gamma>=\<gamma> for \<gamma> :: "'av::{WN,bounded_lattice} \<Rightarrow> val set"
|
| 47613 | 246 |
begin |
247 |
||
248 |
definition AI_wn :: "com \<Rightarrow> 'av st option acom option" where |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
249 |
"AI_wn c = pfp_wn (step' \<top>) (bot c)" |
| 47613 | 250 |
|
251 |
lemma AI_wn_sound: "AI_wn c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C" |
|
252 |
proof(simp add: CS_def AI_wn_def) |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
253 |
assume 1: "pfp_wn (step' \<top>) (bot c) = Some C" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
254 |
have 2: "strip C = c \<and> step' \<top> C \<le> C" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
255 |
by(rule pfp_wn_pfp[where x="bot c"]) (simp_all add: 1 mono_step'_top) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
256 |
have pfp: "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C" |
| 50986 | 257 |
proof(rule order_trans) |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
258 |
show "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' \<top> C)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
259 |
by(rule step_step') |
| 50986 | 260 |
show "... \<le> \<gamma>\<^isub>c C" |
261 |
by(rule mono_gamma_c[OF conjunct2[OF 2]]) |
|
| 47613 | 262 |
qed |
| 50986 | 263 |
have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp_wn[OF _ 1]) |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
264 |
have "lfp c (step (\<gamma>\<^isub>o \<top>)) \<le> \<gamma>\<^isub>c C" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
265 |
by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o \<top>)", OF 3 pfp]) |
| 50986 | 266 |
thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp |
| 47613 | 267 |
qed |
268 |
||
269 |
end |
|
270 |
||
271 |
interpretation Abs_Int2 |
|
| 51245 | 272 |
where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = "op +" |
| 47613 | 273 |
and test_num' = in_ivl |
274 |
and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl |
|
275 |
defines AI_ivl' is AI_wn |
|
276 |
.. |
|
277 |
||
278 |
||
279 |
subsubsection "Tests" |
|
280 |
||
281 |
definition "step_up_ivl n = |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
282 |
((\<lambda>C. C \<nabla> step_ivl \<top> C)^^n)" |
| 47613 | 283 |
definition "step_down_ivl n = |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
284 |
((\<lambda>C. C \<triangle> step_ivl \<top> C)^^n)" |
| 47613 | 285 |
|
286 |
text{* For @{const test3_ivl}, @{const AI_ivl} needed as many iterations as
|
|
287 |
the loop took to execute. In contrast, @{const AI_ivl'} converges in a
|
|
288 |
constant number of steps: *} |
|
289 |
||
290 |
value "show_acom (step_up_ivl 1 (bot test3_ivl))" |
|
291 |
value "show_acom (step_up_ivl 2 (bot test3_ivl))" |
|
292 |
value "show_acom (step_up_ivl 3 (bot test3_ivl))" |
|
293 |
value "show_acom (step_up_ivl 4 (bot test3_ivl))" |
|
294 |
value "show_acom (step_up_ivl 5 (bot test3_ivl))" |
|
| 49188 | 295 |
value "show_acom (step_up_ivl 6 (bot test3_ivl))" |
296 |
value "show_acom (step_up_ivl 7 (bot test3_ivl))" |
|
297 |
value "show_acom (step_up_ivl 8 (bot test3_ivl))" |
|
298 |
value "show_acom (step_down_ivl 1 (step_up_ivl 8 (bot test3_ivl)))" |
|
299 |
value "show_acom (step_down_ivl 2 (step_up_ivl 8 (bot test3_ivl)))" |
|
300 |
value "show_acom (step_down_ivl 3 (step_up_ivl 8 (bot test3_ivl)))" |
|
301 |
value "show_acom (step_down_ivl 4 (step_up_ivl 8 (bot test3_ivl)))" |
|
| 51036 | 302 |
value "show_acom_opt (AI_ivl' test3_ivl)" |
| 47613 | 303 |
|
304 |
||
305 |
text{* Now all the analyses terminate: *}
|
|
306 |
||
| 51036 | 307 |
value "show_acom_opt (AI_ivl' test4_ivl)" |
308 |
value "show_acom_opt (AI_ivl' test5_ivl)" |
|
309 |
value "show_acom_opt (AI_ivl' test6_ivl)" |
|
| 47613 | 310 |
|
311 |
||
312 |
subsubsection "Generic Termination Proof" |
|
313 |
||
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
314 |
lemma top_on_opt_widen: "top_on X o1 \<Longrightarrow> top_on X o2 \<Longrightarrow> top_on X (o1 \<nabla> o2 :: _ st option)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
315 |
apply(induct o1 o2 rule: widen_option.induct) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
316 |
apply (auto) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
317 |
by transfer simp |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
318 |
|
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
319 |
lemma top_on_opt_narrow: "top_on X o1 \<Longrightarrow> top_on X o2 \<Longrightarrow> top_on X (o1 \<triangle> o2 :: _ st option)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
320 |
apply(induct o1 o2 rule: narrow_option.induct) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
321 |
apply (auto) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
322 |
by transfer simp |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
323 |
|
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
324 |
lemma top_on_acom_widen: |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
325 |
"\<lbrakk>top_on X C1; strip C1 = strip C2; top_on X C2\<rbrakk> \<Longrightarrow> top_on X (C1 \<nabla> C2 :: _ st option acom)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
326 |
by(auto simp add: widen_acom_def top_on_acom_def)(metis top_on_opt_widen in_set_zipE) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
327 |
|
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
328 |
lemma top_on_acom_narrow: |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
329 |
"\<lbrakk>top_on X C1; strip C1 = strip C2; top_on X C2\<rbrakk> \<Longrightarrow> top_on X (C1 \<triangle> C2 :: _ st option acom)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
330 |
by(auto simp add: narrow_acom_def top_on_acom_def)(metis top_on_opt_narrow in_set_zipE) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
331 |
|
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
332 |
text{* The assumptions for widening and narrowing differ because during
|
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
333 |
narrowing we have the invariant @{prop"y \<le> x"} (where @{text y} is the next
|
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
334 |
iterate), but during widening there is no such invariant, there we only have |
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
335 |
that not yet @{prop"y \<le> x"}. This complicates the termination proof for
|
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
336 |
widening. *} |
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
337 |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
338 |
locale Measure_WN = Measure1 where m=m for m :: "'av::{top,WN} \<Rightarrow> nat" +
|
| 47613 | 339 |
fixes n :: "'av \<Rightarrow> nat" |
| 51372 | 340 |
assumes m_anti_mono: "x \<le> y \<Longrightarrow> m x \<ge> m y" |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
341 |
assumes m_widen: "~ y \<le> x \<Longrightarrow> m(x \<nabla> y) < m x" |
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
342 |
assumes n_narrow: "y \<le> x \<Longrightarrow> x \<triangle> y < x \<Longrightarrow> n(x \<triangle> y) < n x" |
| 47613 | 343 |
|
344 |
begin |
|
345 |
||
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
346 |
lemma m_s_anti_mono_rep: assumes "\<forall>x. S1 x \<le> S2 x" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
347 |
shows "(\<Sum>x\<in>X. m (S2 x)) \<le> (\<Sum>x\<in>X. m (S1 x))" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
348 |
proof- |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
349 |
from assms have "\<forall>x. m(S1 x) \<ge> m(S2 x)" by (metis m_anti_mono) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
350 |
thus "(\<Sum>x\<in>X. m (S2 x)) \<le> (\<Sum>x\<in>X. m (S1 x))" by (metis setsum_mono) |
| 51372 | 351 |
qed |
352 |
||
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
353 |
lemma m_s_anti_mono: "S1 \<le> S2 \<Longrightarrow> m_s X S1 \<ge> m_s X S2" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
354 |
unfolding m_s_def |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
355 |
apply (transfer fixing: m) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
356 |
apply(simp add: less_eq_st_rep_iff eq_st_def m_s_anti_mono_rep) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
357 |
done |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
358 |
|
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
359 |
lemma m_s_widen_rep: assumes "finite X" "S1 = S2 on -X" "\<not> S2 x \<le> S1 x" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
360 |
shows "(\<Sum>x\<in>X. m (S1 x \<nabla> S2 x)) < (\<Sum>x\<in>X. m (S1 x))" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
361 |
proof- |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
362 |
have 1: "\<forall>x\<in>X. m(S1 x) \<ge> m(S1 x \<nabla> S2 x)" |
| 51372 | 363 |
by (metis m_anti_mono WN_class.widen1) |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
364 |
have "x \<in> X" using assms(2,3) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
365 |
by(auto simp add: Ball_def) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
366 |
hence 2: "\<exists>x\<in>X. m(S1 x) > m(S1 x \<nabla> S2 x)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
367 |
using assms(3) m_widen by blast |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
368 |
from setsum_strict_mono_ex1[OF `finite X` 1 2] |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
369 |
show ?thesis . |
| 47613 | 370 |
qed |
371 |
||
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
372 |
lemma m_s_widen: "finite X \<Longrightarrow> fun S1 = fun S2 on -X ==> |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
373 |
~ S2 \<le> S1 \<Longrightarrow> m_s X (S1 \<nabla> S2) < m_s X S1" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
374 |
apply(auto simp add: less_st_def m_s_def) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
375 |
apply (transfer fixing: m) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
376 |
apply(auto simp add: less_eq_st_rep_iff m_s_widen_rep) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
377 |
done |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
378 |
|
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
379 |
lemma m_o_anti_mono: "finite X \<Longrightarrow> top_on (-X) o1 \<Longrightarrow> top_on (-X) o2 \<Longrightarrow> |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
380 |
o1 \<le> o2 \<Longrightarrow> m_o X o1 \<ge> m_o X o2" |
| 51372 | 381 |
proof(induction o1 o2 rule: less_eq_option.induct) |
382 |
case 1 thus ?case by (simp add: m_o_def)(metis m_s_anti_mono) |
|
383 |
next |
|
384 |
case 2 thus ?case |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
385 |
by(simp add: m_o_def le_SucI m_s_h split: option.splits) |
| 51372 | 386 |
next |
387 |
case 3 thus ?case by simp |
|
388 |
qed |
|
389 |
||
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
390 |
lemma m_o_widen: "\<lbrakk> finite X; top_on (-X) S1; top_on (-X) S2; \<not> S2 \<le> S1 \<rbrakk> \<Longrightarrow> |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
391 |
m_o X (S1 \<nabla> S2) < m_o X S1" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
392 |
by(auto simp: m_o_def m_s_h less_Suc_eq_le m_s_widen split: option.split) |
| 47613 | 393 |
|
| 49547 | 394 |
lemma m_c_widen: |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
395 |
"strip C1 = strip C2 \<Longrightarrow> top_on (-vars C1) C1 \<Longrightarrow> top_on (-vars C2) C2 |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
396 |
\<Longrightarrow> \<not> C2 \<le> C1 \<Longrightarrow> m_c (C1 \<nabla> C2) < m_c C1" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
397 |
apply(auto simp: m_c_def widen_acom_def) |
| 49547 | 398 |
apply(subgoal_tac "length(annos C2) = length(annos C1)") |
| 51390 | 399 |
prefer 2 apply (simp add: size_annos_same2) |
| 49547 | 400 |
apply (auto) |
401 |
apply(rule setsum_strict_mono_ex1) |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
402 |
apply(auto simp add: m_o_anti_mono vars_acom_def top_on_acom_def top_on_opt_widen widen1 le_iff_le_annos listrel_iff_nth) |
| 49547 | 403 |
apply(rule_tac x=i in bexI) |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
404 |
apply (auto simp: vars_acom_def m_o_widen top_on_acom_def) |
| 49547 | 405 |
done |
406 |
||
407 |
||
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
408 |
definition n_s :: "vname set \<Rightarrow> 'av st \<Rightarrow> nat" ("n\<^isub>s") where
|
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
409 |
"n\<^isub>s X S = (\<Sum>x\<in>X. n(fun S x))" |
| 49547 | 410 |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
411 |
lemma n_s_narrow_rep: |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
412 |
assumes "finite X" "S1 = S2 on -X" "\<forall>x. S2 x \<le> S1 x" "\<forall>x. S1 x \<triangle> S2 x \<le> S1 x" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
413 |
"S1 x \<noteq> S1 x \<triangle> S2 x" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
414 |
shows "(\<Sum>x\<in>X. n (S1 x \<triangle> S2 x)) < (\<Sum>x\<in>X. n (S1 x))" |
| 47613 | 415 |
proof- |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
416 |
have 1: "\<forall>x. n(S1 x \<triangle> S2 x) \<le> n(S1 x)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
417 |
by (metis assms(3) assms(4) eq_iff less_le_not_le n_narrow) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
418 |
have "x \<in> X" by (metis Compl_iff assms(2) assms(5) narrowid) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
419 |
hence 2: "\<exists>x\<in>X. n(S1 x \<triangle> S2 x) < n(S1 x)" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
420 |
by (metis assms(3-5) eq_iff less_le_not_le n_narrow) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
421 |
show ?thesis |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
422 |
apply(rule setsum_strict_mono_ex1[OF `finite X`]) using 1 2 by blast+ |
| 47613 | 423 |
qed |
424 |
||
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
425 |
lemma n_s_narrow: "finite X \<Longrightarrow> fun S1 = fun S2 on -X \<Longrightarrow> S2 \<le> S1 \<Longrightarrow> S1 \<triangle> S2 < S1 |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
426 |
\<Longrightarrow> n\<^isub>s X (S1 \<triangle> S2) < n\<^isub>s X S1" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
427 |
apply(auto simp add: less_st_def n_s_def) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
428 |
apply (transfer fixing: n) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
429 |
apply(auto simp add: less_eq_st_rep_iff eq_st_def fun_eq_iff n_s_narrow_rep) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
430 |
done |
| 47613 | 431 |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
432 |
definition n_o :: "vname set \<Rightarrow> 'av st option \<Rightarrow> nat" ("n\<^isub>o") where
|
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
433 |
"n\<^isub>o X opt = (case opt of None \<Rightarrow> 0 | Some S \<Rightarrow> n\<^isub>s X S + 1)" |
| 47613 | 434 |
|
435 |
lemma n_o_narrow: |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
436 |
"top_on (-X) S1 \<Longrightarrow> top_on (-X) S2 \<Longrightarrow> finite X |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
437 |
\<Longrightarrow> S2 \<le> S1 \<Longrightarrow> S1 \<triangle> S2 < S1 \<Longrightarrow> n\<^isub>o X (S1 \<triangle> S2) < n\<^isub>o X S1" |
| 47613 | 438 |
apply(induction S1 S2 rule: narrow_option.induct) |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
439 |
apply(auto simp: n_o_def n_s_narrow) |
| 47613 | 440 |
done |
441 |
||
| 49576 | 442 |
|
443 |
definition n_c :: "'av st option acom \<Rightarrow> nat" ("n\<^isub>c") where
|
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
444 |
"n\<^isub>c C = (let as = annos C in \<Sum>i<size as. n\<^isub>o (vars C) (as!i))" |
| 47613 | 445 |
|
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
446 |
lemma less_annos_iff: "(C1 < C2) = (C1 \<le> C2 \<and> |
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
447 |
(\<exists>i<length (annos C1). annos C1 ! i < annos C2 ! i))" |
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
448 |
by(metis (hide_lams, no_types) less_le_not_le le_iff_le_annos size_annos_same2) |
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
449 |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
450 |
lemma n_c_narrow: "strip C1 = strip C2 |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
451 |
\<Longrightarrow> top_on (- vars C1) C1 \<Longrightarrow> top_on (- vars C2) C2 |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
452 |
\<Longrightarrow> C2 \<le> C1 \<Longrightarrow> C1 \<triangle> C2 < C1 \<Longrightarrow> n\<^isub>c (C1 \<triangle> C2) < n\<^isub>c C1" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
453 |
apply(auto simp: n_c_def Let_def narrow_acom_def) |
| 47613 | 454 |
apply(subgoal_tac "length(annos C2) = length(annos C1)") |
455 |
prefer 2 apply (simp add: size_annos_same2) |
|
456 |
apply (auto) |
|
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
457 |
apply(simp add: less_annos_iff le_iff_le_annos) |
| 47613 | 458 |
apply(rule setsum_strict_mono_ex1) |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
459 |
apply (auto simp: vars_acom_def top_on_acom_def) |
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
460 |
apply (metis n_o_narrow nth_mem finite_cvars less_imp_le le_less order_refl) |
| 47613 | 461 |
apply(rule_tac x=i in bexI) |
462 |
prefer 2 apply simp |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
463 |
apply(rule n_o_narrow[where X = "vars(strip C2)"]) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
464 |
apply (simp_all) |
| 47613 | 465 |
done |
466 |
||
467 |
end |
|
468 |
||
469 |
||
470 |
lemma iter_widen_termination: |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
471 |
fixes m :: "'a::WN acom \<Rightarrow> nat" |
| 47613 | 472 |
assumes P_f: "\<And>C. P C \<Longrightarrow> P(f C)" |
473 |
and P_widen: "\<And>C1 C2. P C1 \<Longrightarrow> P C2 \<Longrightarrow> P(C1 \<nabla> C2)" |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
474 |
and m_widen: "\<And>C1 C2. P C1 \<Longrightarrow> P C2 \<Longrightarrow> ~ C2 \<le> C1 \<Longrightarrow> m(C1 \<nabla> C2) < m C1" |
| 47613 | 475 |
and "P C" shows "EX C'. iter_widen f C = Some C'" |
| 49547 | 476 |
proof(simp add: iter_widen_def, |
477 |
rule measure_while_option_Some[where P = P and f=m]) |
|
| 47613 | 478 |
show "P C" by(rule `P C`) |
479 |
next |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
480 |
fix C assume "P C" "\<not> f C \<le> C" thus "P (C \<nabla> f C) \<and> m (C \<nabla> f C) < m C" |
| 49547 | 481 |
by(simp add: P_f P_widen m_widen) |
| 47613 | 482 |
qed |
| 49496 | 483 |
|
| 47613 | 484 |
lemma iter_narrow_termination: |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
485 |
fixes n :: "'a::WN acom \<Rightarrow> nat" |
| 47613 | 486 |
assumes P_f: "\<And>C. P C \<Longrightarrow> P(f C)" |
487 |
and P_narrow: "\<And>C1 C2. P C1 \<Longrightarrow> P C2 \<Longrightarrow> P(C1 \<triangle> C2)" |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
488 |
and mono: "\<And>C1 C2. P C1 \<Longrightarrow> P C2 \<Longrightarrow> C1 \<le> C2 \<Longrightarrow> f C1 \<le> f C2" |
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
489 |
and n_narrow: "\<And>C1 C2. P C1 \<Longrightarrow> P C2 \<Longrightarrow> C2 \<le> C1 \<Longrightarrow> C1 \<triangle> C2 < C1 \<Longrightarrow> n(C1 \<triangle> C2) < n C1" |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
490 |
and init: "P C" "f C \<le> C" shows "EX C'. iter_narrow f C = Some C'" |
| 49547 | 491 |
proof(simp add: iter_narrow_def, |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
492 |
rule measure_while_option_Some[where f=n and P = "%C. P C \<and> f C \<le> C"]) |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
493 |
show "P C \<and> f C \<le> C" using init by blast |
| 47613 | 494 |
next |
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
495 |
fix C assume 1: "P C \<and> f C \<le> C" and 2: "C \<triangle> f C < C" |
| 47613 | 496 |
hence "P (C \<triangle> f C)" by(simp add: P_f P_narrow) |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
497 |
moreover then have "f (C \<triangle> f C) \<le> C \<triangle> f C" |
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
498 |
by (metis narrow1_acom narrow2_acom 1 mono order_trans) |
| 49547 | 499 |
moreover have "n (C \<triangle> f C) < n C" using 1 2 by(simp add: n_narrow P_f) |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
500 |
ultimately show "(P (C \<triangle> f C) \<and> f (C \<triangle> f C) \<le> C \<triangle> f C) \<and> n(C \<triangle> f C) < n C" |
| 49547 | 501 |
by blast |
| 47613 | 502 |
qed |
503 |
||
| 49547 | 504 |
locale Abs_Int2_measure = |
505 |
Abs_Int2 where \<gamma>=\<gamma> + Measure_WN where m=m |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
506 |
for \<gamma> :: "'av::{WN,bounded_lattice} \<Rightarrow> val set" and m :: "'av \<Rightarrow> nat"
|
| 49547 | 507 |
|
| 47613 | 508 |
|
509 |
subsubsection "Termination: Intervals" |
|
510 |
||
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
511 |
definition m_rep :: "eint2 \<Rightarrow> nat" where |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
512 |
"m_rep p = (if is_empty_rep p then 3 else |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
513 |
let (l,h) = p in (case l of Minf \<Rightarrow> 0 | _ \<Rightarrow> 1) + (case h of Pinf \<Rightarrow> 0 | _ \<Rightarrow> 1))" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
514 |
|
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
515 |
lift_definition m_ivl :: "ivl \<Rightarrow> nat" is m_rep |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
516 |
by(auto simp: m_rep_def eq_ivl_iff) |
| 47613 | 517 |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
518 |
lemma m_ivl_nice: "m_ivl[l\<dots>h] = (if [l\<dots>h] = \<bottom> then 3 else |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
519 |
(if l = Minf then 0 else 1) + (if h = Pinf then 0 else 1))" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
520 |
unfolding bot_ivl_def |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
521 |
by transfer (auto simp: m_rep_def eq_ivl_empty split: extended.split) |
| 47613 | 522 |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
523 |
lemma m_ivl_height: "m_ivl iv \<le> 3" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
524 |
by transfer (simp add: m_rep_def split: prod.split extended.split) |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
525 |
|
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
526 |
lemma m_ivl_anti_mono: "y \<le> x \<Longrightarrow> m_ivl x \<le> m_ivl y" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
527 |
by transfer |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
528 |
(auto simp: m_rep_def is_empty_rep_def \<gamma>_rep_cases le_iff_subset |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
529 |
split: prod.split extended.splits if_splits) |
| 47613 | 530 |
|
531 |
lemma m_ivl_widen: |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
532 |
"~ y \<le> x \<Longrightarrow> m_ivl(x \<nabla> y) < m_ivl x" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
533 |
by transfer |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
534 |
(auto simp: m_rep_def widen_rep_def is_empty_rep_def \<gamma>_rep_cases le_iff_subset |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
535 |
split: prod.split extended.splits if_splits) |
| 47613 | 536 |
|
537 |
definition n_ivl :: "ivl \<Rightarrow> nat" where |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
538 |
"n_ivl ivl = 3 - m_ivl ivl" |
| 47613 | 539 |
|
540 |
lemma n_ivl_narrow: |
|
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
541 |
"x \<triangle> y < x \<Longrightarrow> n_ivl(x \<triangle> y) < n_ivl x" |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
542 |
unfolding n_ivl_def |
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
543 |
apply(subst (asm) less_le_not_le) |
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
544 |
apply transfer |
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
545 |
by(auto simp add: m_rep_def narrow_rep_def is_empty_rep_def empty_rep_def \<gamma>_rep_cases le_iff_subset |
|
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
546 |
split: prod.splits if_splits extended.split) |
| 47613 | 547 |
|
548 |
||
549 |
interpretation Abs_Int2_measure |
|
| 51245 | 550 |
where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = "op +" |
| 47613 | 551 |
and test_num' = in_ivl |
552 |
and filter_plus' = filter_plus_ivl and filter_less' = filter_less_ivl |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
553 |
and m = m_ivl and n = n_ivl and h = 3 |
| 47613 | 554 |
proof |
| 51372 | 555 |
case goal2 thus ?case by(rule m_ivl_anti_mono) |
| 47613 | 556 |
next |
| 51372 | 557 |
case goal1 thus ?case by(rule m_ivl_height) |
| 47613 | 558 |
next |
| 49547 | 559 |
case goal3 thus ?case by(rule m_ivl_widen) |
| 47613 | 560 |
next |
|
51385
f193d44d4918
termination proof for narrowing: fewer assumptions
nipkow
parents:
51372
diff
changeset
|
561 |
case goal4 from goal4(2) show ?case by(rule n_ivl_narrow) |
| 49576 | 562 |
-- "note that the first assms is unnecessary for intervals" |
| 47613 | 563 |
qed |
564 |
||
565 |
lemma iter_winden_step_ivl_termination: |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
566 |
"\<exists>C. iter_widen (step_ivl \<top>) (bot c) = Some C" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
567 |
apply(rule iter_widen_termination[where m = "m_c" and P = "%C. strip C = c \<and> top_on (- vars C) C"]) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
568 |
apply (auto simp add: m_c_widen top_on_bot top_on_step'[simplified comp_def vars_acom_def] |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
569 |
vars_acom_def top_on_acom_widen) |
| 47613 | 570 |
done |
571 |
||
572 |
lemma iter_narrow_step_ivl_termination: |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
573 |
"top_on (- vars C0) C0 \<Longrightarrow> step_ivl \<top> C0 \<le> C0 \<Longrightarrow> |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
574 |
\<exists>C. iter_narrow (step_ivl \<top>) C0 = Some C" |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
575 |
apply(rule iter_narrow_termination[where n = "n_c" and P = "%C. strip C0 = strip C \<and> top_on (-vars C) C"]) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
576 |
apply(auto simp: top_on_step'[simplified comp_def vars_acom_def] |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
577 |
mono_step'_top n_c_narrow vars_acom_def top_on_acom_narrow) |
| 47613 | 578 |
done |
579 |
||
580 |
theorem AI_ivl'_termination: |
|
581 |
"\<exists>C. AI_ivl' c = Some C" |
|
582 |
apply(auto simp: AI_wn_def pfp_wn_def iter_winden_step_ivl_termination |
|
583 |
split: option.split) |
|
584 |
apply(rule iter_narrow_step_ivl_termination) |
|
|
51711
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
585 |
apply(rule conjunct2) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
586 |
apply(rule iter_widen_inv[where f = "step' \<top>" and P = "%C. c = strip C & top_on (- vars C) C"]) |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
587 |
apply(auto simp: top_on_acom_widen top_on_step'[simplified comp_def vars_acom_def] |
|
df3426139651
complete revision: finally got rid of annoying L-predicate
nipkow
parents:
51390
diff
changeset
|
588 |
iter_widen_pfp top_on_bot vars_acom_def) |
| 47613 | 589 |
done |
590 |
||
| 51390 | 591 |
(*unused_thms Abs_Int_init - *) |
| 47613 | 592 |
|
| 49578 | 593 |
subsubsection "Counterexamples" |
594 |
||
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
595 |
text{* Widening is increasing by assumption, but @{prop"x \<le> f x"} is not an invariant of widening.
|
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
596 |
It can already be lost after the first step: *} |
| 49578 | 597 |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
598 |
lemma assumes "!!x y::'a::WN. x \<le> y \<Longrightarrow> f x \<le> f y" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
599 |
and "x \<le> f x" and "\<not> f x \<le> x" shows "x \<nabla> f x \<le> f(x \<nabla> f x)" |
| 49578 | 600 |
nitpick[card = 3, expect = genuine, show_consts] |
601 |
(* |
|
602 |
1 < 2 < 3, |
|
603 |
f x = 2, |
|
604 |
x widen y = 3 -- guarantees termination with top=3 |
|
605 |
x = 1 |
|
606 |
Now f is mono, x <= f x, not f x <= x |
|
607 |
but x widen f x = 3, f 3 = 2, but not 3 <= 2 |
|
608 |
*) |
|
609 |
oops |
|
610 |
||
611 |
text{* Widening terminates but may converge more slowly than Kleene iteration.
|
|
612 |
In the following model, Kleene iteration goes from 0 to the least pfp |
|
613 |
in one step but widening takes 2 steps to reach a strictly larger pfp: *} |
|
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
614 |
lemma assumes "!!x y::'a::WN. x \<le> y \<Longrightarrow> f x \<le> f y" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
615 |
and "x \<le> f x" and "\<not> f x \<le> x" and "f(f x) \<le> f x" |
|
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51245
diff
changeset
|
616 |
shows "f(x \<nabla> f x) \<le> x \<nabla> f x" |
| 49578 | 617 |
nitpick[card = 4, expect = genuine, show_consts] |
618 |
(* |
|
619 |
||
620 |
0 < 1 < 2 < 3 |
|
621 |
f: 1 1 3 3 |
|
622 |
||
623 |
0 widen 1 = 2 |
|
624 |
2 widen 3 = 3 |
|
625 |
and x widen y arbitrary, eg 3, which guarantees termination |
|
626 |
||
627 |
Kleene: f(f 0) = f 1 = 1 <= 1 = f 1 |
|
628 |
||
629 |
but |
|
630 |
||
631 |
because not f 0 <= 0, we obtain 0 widen f 0 = 0 wide 1 = 2, |
|
632 |
which is again not a pfp: not f 2 = 3 <= 2 |
|
633 |
Another widening step yields 2 widen f 2 = 2 widen 3 = 3 |
|
634 |
*) |
|
|
49892
09956f7a00af
proper 'oops' to force sequential checking here, and avoid spurious *** Interrupt stemming from crash of forked outer syntax element;
wenzelm
parents:
49579
diff
changeset
|
635 |
oops |
| 49578 | 636 |
|
| 47613 | 637 |
end |