| 44932 |      1 | (* Author: Tobias Nipkow *)
 | 
|  |      2 | 
 | 
| 45111 |      3 | theory Abs_Int_den2
 | 
|  |      4 | imports Abs_Int_den1_ivl
 | 
| 44932 |      5 | begin
 | 
|  |      6 | 
 | 
| 45022 |      7 | context preord
 | 
|  |      8 | begin
 | 
|  |      9 | 
 | 
|  |     10 | definition mono where "mono f = (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y)"
 | 
|  |     11 | 
 | 
|  |     12 | lemma monoD: "mono f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" by(simp add: mono_def)
 | 
|  |     13 | 
 | 
|  |     14 | lemma mono_comp: "mono f \<Longrightarrow> mono g \<Longrightarrow> mono (g o f)"
 | 
|  |     15 | by(simp add: mono_def)
 | 
|  |     16 | 
 | 
|  |     17 | declare le_trans[trans]
 | 
|  |     18 | 
 | 
|  |     19 | end
 | 
|  |     20 | 
 | 
|  |     21 | 
 | 
| 44932 |     22 | subsection "Widening and Narrowing"
 | 
|  |     23 | 
 | 
| 45022 |     24 | text{* Jumping to the trivial post-fixed point @{const Top} in case @{text k}
 | 
|  |     25 | rounds of iteration did not reach a post-fixed point (as in @{const iter}) is
 | 
|  |     26 | a trivial widening step. We generalise this idea and complement it with
 | 
|  |     27 | narrowing, a process to regain precision.
 | 
|  |     28 | 
 | 
|  |     29 | Class @{text WN} makes some assumptions about the widening and narrowing
 | 
|  |     30 | operators. The assumptions serve two purposes. Together with a further
 | 
|  |     31 | assumption that certain chains become stationary, they permit to prove
 | 
|  |     32 | termination of the fixed point iteration, which we do not --- we limit the
 | 
|  |     33 | number of iterations as before. The second purpose of the narrowing
 | 
|  |     34 | assumptions is to prove that the narrowing iteration keeps on producing
 | 
|  |     35 | post-fixed points and that it goes down. However, this requires the function
 | 
|  |     36 | being iterated to be monotone. Unfortunately, abstract interpretation with
 | 
|  |     37 | widening is not monotone. Hence the (recursive) abstract interpretation of a
 | 
|  |     38 | loop body that again contains a loop may result in a non-monotone
 | 
| 45073 |     39 | function. Therefore our narrowing iteration needs to check at every step
 | 
| 45022 |     40 | that a post-fixed point is maintained, and we cannot prove that the precision
 | 
|  |     41 | increases. *}
 | 
| 44932 |     42 | 
 | 
|  |     43 | class WN = SL_top +
 | 
|  |     44 | fixes widen :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<nabla>" 65)
 | 
|  |     45 | assumes widen: "y \<sqsubseteq> x \<nabla> y"
 | 
|  |     46 | fixes narrow :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<triangle>" 65)
 | 
|  |     47 | assumes narrow1: "y \<sqsubseteq> x \<Longrightarrow> y \<sqsubseteq> x \<triangle> y"
 | 
|  |     48 | assumes narrow2: "y \<sqsubseteq> x \<Longrightarrow> x \<triangle> y \<sqsubseteq> x"
 | 
|  |     49 | begin
 | 
|  |     50 | 
 | 
|  |     51 | fun iter_up :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
|  |     52 | "iter_up f 0 x = Top" |
 | 
|  |     53 | "iter_up f (Suc n) x =
 | 
|  |     54 |   (let fx = f x in if fx \<sqsubseteq> x then x else iter_up f n (x \<nabla> fx))"
 | 
|  |     55 | 
 | 
|  |     56 | lemma iter_up_pfp: "f(iter_up f n x) \<sqsubseteq> iter_up f n x"
 | 
| 45015 |     57 | apply (induction n arbitrary: x)
 | 
| 44932 |     58 |  apply (simp)
 | 
|  |     59 | apply (simp add: Let_def)
 | 
|  |     60 | done
 | 
|  |     61 | 
 | 
|  |     62 | fun iter_down :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
|  |     63 | "iter_down f 0 x = x" |
 | 
|  |     64 | "iter_down f (Suc n) x =
 | 
|  |     65 |   (let y = x \<triangle> f x in if f y \<sqsubseteq> y then iter_down f n y else x)"
 | 
|  |     66 | 
 | 
|  |     67 | lemma iter_down_pfp: "f x \<sqsubseteq> x \<Longrightarrow> f(iter_down f n x) \<sqsubseteq> iter_down f n x"
 | 
| 45015 |     68 | apply (induction n arbitrary: x)
 | 
| 44932 |     69 |  apply (simp)
 | 
|  |     70 | apply (simp add: Let_def)
 | 
|  |     71 | done
 | 
|  |     72 | 
 | 
| 44944 |     73 | definition iter' :: "nat \<Rightarrow> nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
|  |     74 | "iter' m n f x =
 | 
| 44932 |     75 |   (let f' = (\<lambda>y. x \<squnion> f y) in iter_down f' n (iter_up f' m x))"
 | 
|  |     76 | 
 | 
| 44944 |     77 | lemma iter'_pfp_above:
 | 
|  |     78 | shows "f(iter' m n f x0) \<sqsubseteq> iter' m n f x0"
 | 
|  |     79 | and "x0 \<sqsubseteq> iter' m n f x0"
 | 
| 44932 |     80 | using iter_up_pfp[of "\<lambda>x. x0 \<squnion> f x"] iter_down_pfp[of "\<lambda>x. x0 \<squnion> f x"]
 | 
| 44944 |     81 | by(auto simp add: iter'_def Let_def)
 | 
| 44932 |     82 | 
 | 
| 45022 |     83 | text{* This is how narrowing works on monotone functions: you just iterate. *}
 | 
|  |     84 | 
 | 
|  |     85 | abbreviation iter_down_mono :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
|  |     86 | "iter_down_mono f n x == ((\<lambda>x. x \<triangle> f x)^^n) x"
 | 
|  |     87 | 
 | 
|  |     88 | text{* Narrowing always yields a post-fixed point: *}
 | 
|  |     89 | 
 | 
|  |     90 | lemma iter_down_mono_pfp: assumes "mono f" and "f x0 \<sqsubseteq> x0" 
 | 
|  |     91 | defines "x n == iter_down_mono f n x0"
 | 
|  |     92 | shows "f(x n) \<sqsubseteq> x n"
 | 
|  |     93 | proof (induction n)
 | 
|  |     94 |   case 0 show ?case by (simp add: x_def assms(2))
 | 
|  |     95 | next
 | 
|  |     96 |   case (Suc n)
 | 
|  |     97 |   have "f (x (Suc n)) = f(x n \<triangle> f(x n))" by(simp add: x_def)
 | 
|  |     98 |   also have "\<dots> \<sqsubseteq> f(x n)" by(rule monoD[OF `mono f` narrow2[OF Suc]])
 | 
|  |     99 |   also have "\<dots> \<sqsubseteq> x n \<triangle> f(x n)" by(rule narrow1[OF Suc])
 | 
|  |    100 |   also have "\<dots> = x(Suc n)" by(simp add: x_def)
 | 
|  |    101 |   finally show ?case .
 | 
|  |    102 | qed
 | 
|  |    103 | 
 | 
|  |    104 | text{* Narrowing can only increase precision: *}
 | 
|  |    105 | 
 | 
|  |    106 | lemma iter_down_down: assumes "mono f" and "f x0 \<sqsubseteq> x0" 
 | 
|  |    107 | defines "x n == iter_down_mono f n x0"
 | 
|  |    108 | shows "x n \<sqsubseteq> x0"
 | 
|  |    109 | proof (induction n)
 | 
|  |    110 |   case 0 show ?case by(simp add: x_def)
 | 
|  |    111 | next
 | 
|  |    112 |   case (Suc n)
 | 
|  |    113 |   have "x(Suc n) = x n \<triangle> f(x n)" by(simp add: x_def)
 | 
|  |    114 |   also have "\<dots> \<sqsubseteq> x n" unfolding x_def
 | 
|  |    115 |     by(rule narrow2[OF iter_down_mono_pfp[OF assms(1), OF assms(2)]])
 | 
|  |    116 |   also have "\<dots> \<sqsubseteq> x0" by(rule Suc)
 | 
|  |    117 |   finally show ?case .
 | 
|  |    118 | qed
 | 
|  |    119 | 
 | 
|  |    120 | 
 | 
| 44932 |    121 | end
 | 
|  |    122 | 
 | 
|  |    123 | 
 | 
|  |    124 | instantiation ivl :: WN
 | 
|  |    125 | begin
 | 
|  |    126 | 
 | 
|  |    127 | definition "widen_ivl ivl1 ivl2 =
 | 
| 45019 |    128 |   ((*if is_empty ivl1 then ivl2 else if is_empty ivl2 then ivl1 else*)
 | 
| 44932 |    129 |      case (ivl1,ivl2) of (I l1 h1, I l2 h2) \<Rightarrow>
 | 
|  |    130 |        I (if le_option False l2 l1 \<and> l2 \<noteq> l1 then None else l2)
 | 
|  |    131 |          (if le_option True h1 h2 \<and> h1 \<noteq> h2 then None else h2))"
 | 
|  |    132 | 
 | 
|  |    133 | definition "narrow_ivl ivl1 ivl2 =
 | 
|  |    134 |   ((*if is_empty ivl1 \<or> is_empty ivl2 then empty else*)
 | 
|  |    135 |      case (ivl1,ivl2) of (I l1 h1, I l2 h2) \<Rightarrow>
 | 
|  |    136 |        I (if l1 = None then l2 else l1)
 | 
|  |    137 |          (if h1 = None then h2 else h1))"
 | 
|  |    138 | 
 | 
|  |    139 | instance
 | 
|  |    140 | proof qed
 | 
|  |    141 |   (auto simp add: widen_ivl_def narrow_ivl_def le_option_def le_ivl_def empty_def split: ivl.split option.split if_splits)
 | 
|  |    142 | 
 | 
|  |    143 | end
 | 
|  |    144 | 
 | 
| 45023 |    145 | instantiation astate :: (WN) WN
 | 
| 44932 |    146 | begin
 | 
|  |    147 | 
 | 
|  |    148 | definition "widen_astate F1 F2 =
 | 
|  |    149 |   FunDom (\<lambda>x. fun F1 x \<nabla> fun F2 x) (inter_list (dom F1) (dom F2))"
 | 
|  |    150 | 
 | 
|  |    151 | definition "narrow_astate F1 F2 =
 | 
|  |    152 |   FunDom (\<lambda>x. fun F1 x \<triangle> fun F2 x) (inter_list (dom F1) (dom F2))"
 | 
|  |    153 | 
 | 
|  |    154 | instance
 | 
|  |    155 | proof
 | 
|  |    156 |   case goal1 thus ?case
 | 
|  |    157 |     by(simp add: widen_astate_def le_astate_def lookup_def widen)
 | 
|  |    158 | next
 | 
|  |    159 |   case goal2 thus ?case
 | 
|  |    160 |     by(auto simp: narrow_astate_def le_astate_def lookup_def narrow1)
 | 
|  |    161 | next
 | 
|  |    162 |   case goal3 thus ?case
 | 
|  |    163 |     by(auto simp: narrow_astate_def le_astate_def lookup_def narrow2)
 | 
|  |    164 | qed
 | 
|  |    165 | 
 | 
|  |    166 | end
 | 
|  |    167 | 
 | 
| 45023 |    168 | instantiation up :: (WN) WN
 | 
| 44932 |    169 | begin
 | 
|  |    170 | 
 | 
|  |    171 | fun widen_up where
 | 
|  |    172 | "widen_up bot x = x" |
 | 
|  |    173 | "widen_up x bot = x" |
 | 
|  |    174 | "widen_up (Up x) (Up y) = Up(x \<nabla> y)"
 | 
|  |    175 | 
 | 
|  |    176 | fun narrow_up where
 | 
|  |    177 | "narrow_up bot x = bot" |
 | 
|  |    178 | "narrow_up x bot = bot" |
 | 
|  |    179 | "narrow_up (Up x) (Up y) = Up(x \<triangle> y)"
 | 
|  |    180 | 
 | 
|  |    181 | instance
 | 
|  |    182 | proof
 | 
|  |    183 |   case goal1 show ?case
 | 
|  |    184 |     by(induct x y rule: widen_up.induct) (simp_all add: widen)
 | 
|  |    185 | next
 | 
|  |    186 |   case goal2 thus ?case
 | 
|  |    187 |     by(induct x y rule: narrow_up.induct) (simp_all add: narrow1)
 | 
|  |    188 | next
 | 
|  |    189 |   case goal3 thus ?case
 | 
|  |    190 |     by(induct x y rule: narrow_up.induct) (simp_all add: narrow2)
 | 
|  |    191 | qed
 | 
|  |    192 | 
 | 
|  |    193 | end
 | 
|  |    194 | 
 | 
|  |    195 | interpretation
 | 
| 45023 |    196 |   Abs_Int1 rep_ivl num_ivl plus_ivl filter_plus_ivl filter_less_ivl "(iter' 3 2)"
 | 
| 44932 |    197 | defines afilter_ivl' is afilter
 | 
|  |    198 | and bfilter_ivl' is bfilter
 | 
|  |    199 | and AI_ivl' is AI
 | 
|  |    200 | and aval_ivl' is aval'
 | 
| 44944 |    201 | proof qed (auto simp: iter'_pfp_above)
 | 
| 44932 |    202 | 
 | 
|  |    203 | value [code] "list_up(AI_ivl' test3_ivl Top)"
 | 
|  |    204 | value [code] "list_up(AI_ivl' test4_ivl Top)"
 | 
| 45020 |    205 | value [code] "list_up(AI_ivl' test5_ivl Top)"
 | 
| 44932 |    206 | 
 | 
|  |    207 | end
 |