| author | wenzelm | 
| Thu, 09 Oct 2014 13:56:27 +0200 | |
| changeset 58640 | 37f852399a32 | 
| parent 55601 | b7f4da504b75 | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 44656 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 45110 
305f83b6da54
moved IMP/AbsInt stuff into subdirectory Abs_Int_Den
 nipkow parents: 
45015diff
changeset | 3 | header "Denotational Abstract Interpretation" | 
| 44656 | 4 | |
| 45111 | 5 | theory Abs_Int_den0_fun | 
| 55601 | 6 | imports "~~/src/Tools/Permanent_Interpretation" "../Big_Step" | 
| 44656 | 7 | begin | 
| 8 | ||
| 9 | subsection "Orderings" | |
| 10 | ||
| 11 | class preord = | |
| 12 | fixes le :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>" 50) | |
| 13 | assumes le_refl[simp]: "x \<sqsubseteq> x" | |
| 14 | and le_trans: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z" | |
| 15 | ||
| 16 | text{* Note: no antisymmetry. Allows implementations where some abstract
 | |
| 17 | element is implemented by two different values @{prop "x \<noteq> y"}
 | |
| 18 | such that @{prop"x \<sqsubseteq> y"} and @{prop"y \<sqsubseteq> x"}. Antisymmetry is not
 | |
| 19 | needed because we never compare elements for equality but only for @{text"\<sqsubseteq>"}.
 | |
| 20 | *} | |
| 21 | ||
| 22 | class SL_top = preord + | |
| 23 | fixes join :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65) | |
| 24 | fixes Top :: "'a" | |
| 25 | assumes join_ge1 [simp]: "x \<sqsubseteq> x \<squnion> y" | |
| 26 | and join_ge2 [simp]: "y \<sqsubseteq> x \<squnion> y" | |
| 27 | and join_least: "x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<squnion> y \<sqsubseteq> z" | |
| 28 | and top[simp]: "x \<sqsubseteq> Top" | |
| 29 | begin | |
| 30 | ||
| 31 | lemma join_le_iff[simp]: "x \<squnion> y \<sqsubseteq> z \<longleftrightarrow> x \<sqsubseteq> z \<and> y \<sqsubseteq> z" | |
| 32 | by (metis join_ge1 join_ge2 join_least le_trans) | |
| 33 | ||
| 44932 | 34 | fun iter :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
| 35 | "iter 0 f _ = Top" | | |
| 36 | "iter (Suc n) f x = (if f x \<sqsubseteq> x then x else iter n f (f x))" | |
| 44656 | 37 | |
| 44932 | 38 | lemma iter_pfp: "f(iter n f x) \<sqsubseteq> iter n f x" | 
| 45015 | 39 | apply (induction n arbitrary: x) | 
| 44656 | 40 | apply (simp) | 
| 41 | apply (simp) | |
| 42 | done | |
| 43 | ||
| 44944 | 44 | abbreviation iter' :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
| 45 | "iter' n f x0 == iter n (\<lambda>x. x0 \<squnion> f x) x0" | |
| 44923 | 46 | |
| 44944 | 47 | lemma iter'_pfp_above: | 
| 48 | "f(iter' n f x0) \<sqsubseteq> iter' n f x0" "x0 \<sqsubseteq> iter' n f x0" | |
| 49 | using iter_pfp[of "\<lambda>x. x0 \<squnion> f x"] by auto | |
| 44923 | 50 | |
| 51 | text{* So much for soundness. But how good an approximation of the post-fixed
 | |
| 44944 | 52 | point does @{const iter} yield? *}
 | 
| 44923 | 53 | |
| 44932 | 54 | lemma iter_funpow: "iter n f x \<noteq> Top \<Longrightarrow> \<exists>k. iter n f x = (f^^k) x" | 
| 45015 | 55 | apply(induction n arbitrary: x) | 
| 44923 | 56 | apply simp | 
| 44656 | 57 | apply (auto) | 
| 44932 | 58 | apply(metis funpow.simps(1) id_def) | 
| 44656 | 59 | by (metis funpow.simps(2) funpow_swap1 o_apply) | 
| 60 | ||
| 44944 | 61 | text{* For monotone @{text f}, @{term "iter f n x0"} yields the least
 | 
| 44923 | 62 | post-fixed point above @{text x0}, unless it yields @{text Top}. *}
 | 
| 44656 | 63 | |
| 44923 | 64 | lemma iter_least_pfp: | 
| 44944 | 65 | assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" and "iter n f x0 \<noteq> Top" | 
| 66 | and "f p \<sqsubseteq> p" and "x0 \<sqsubseteq> p" shows "iter n f x0 \<sqsubseteq> p" | |
| 44656 | 67 | proof- | 
| 44944 | 68 | obtain k where "iter n f x0 = (f^^k) x0" | 
| 69 | using iter_funpow[OF `iter n f x0 \<noteq> Top`] by blast | |
| 44656 | 70 | moreover | 
| 44944 | 71 |   { fix n have "(f^^n) x0 \<sqsubseteq> p"
 | 
| 45015 | 72 | proof(induction n) | 
| 44656 | 73 | case 0 show ?case by(simp add: `x0 \<sqsubseteq> p`) | 
| 74 | next | |
| 75 | case (Suc n) thus ?case | |
| 76 | by (simp add: `x0 \<sqsubseteq> p`)(metis Suc assms(3) le_trans mono) | |
| 77 | qed | |
| 78 | } ultimately show ?thesis by simp | |
| 79 | qed | |
| 80 | ||
| 81 | end | |
| 82 | ||
| 83 | text{* The interface of abstract values: *}
 | |
| 84 | ||
| 85 | locale Rep = | |
| 86 | fixes rep :: "'a::SL_top \<Rightarrow> 'b set" | |
| 87 | assumes le_rep: "a \<sqsubseteq> b \<Longrightarrow> rep a \<subseteq> rep b" | |
| 88 | begin | |
| 89 | ||
| 90 | abbreviation in_rep (infix "<:" 50) where "x <: a == x : rep a" | |
| 91 | ||
| 92 | lemma in_rep_join: "x <: a1 \<or> x <: a2 \<Longrightarrow> x <: a1 \<squnion> a2" | |
| 93 | by (metis SL_top_class.join_ge1 SL_top_class.join_ge2 le_rep subsetD) | |
| 94 | ||
| 95 | end | |
| 96 | ||
| 97 | locale Val_abs = Rep rep | |
| 98 | for rep :: "'a::SL_top \<Rightarrow> val set" + | |
| 44932 | 99 | fixes num' :: "val \<Rightarrow> 'a" | 
| 100 | and plus' :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 44656 | 101 | assumes rep_num': "rep(num' n) = {n}"
 | 
| 102 | and rep_plus': "n1 <: a1 \<Longrightarrow> n2 <: a2 \<Longrightarrow> n1+n2 <: plus' a1 a2" | |
| 103 | ||
| 104 | ||
| 105 | instantiation "fun" :: (type, SL_top) SL_top | |
| 106 | begin | |
| 107 | ||
| 108 | definition "f \<sqsubseteq> g = (ALL x. f x \<sqsubseteq> g x)" | |
| 109 | definition "f \<squnion> g = (\<lambda>x. f x \<squnion> g x)" | |
| 110 | definition "Top = (\<lambda>x. Top)" | |
| 111 | ||
| 112 | lemma join_apply[simp]: | |
| 113 | "(f \<squnion> g) x = f x \<squnion> g x" | |
| 114 | by (simp add: join_fun_def) | |
| 115 | ||
| 116 | instance | |
| 117 | proof | |
| 118 | case goal2 thus ?case by (metis le_fun_def preord_class.le_trans) | |
| 119 | qed (simp_all add: le_fun_def Top_fun_def) | |
| 120 | ||
| 121 | end | |
| 122 | ||
| 123 | subsection "Abstract Interpretation Abstractly" | |
| 124 | ||
| 44932 | 125 | text{* Abstract interpretation over abstract values. Abstract states are
 | 
| 126 | simply functions. The post-fixed point finder is parameterized over. *} | |
| 127 | ||
| 45212 | 128 | type_synonym 'a st = "vname \<Rightarrow> 'a" | 
| 44656 | 129 | |
| 44932 | 130 | locale Abs_Int_Fun = Val_abs + | 
| 44944 | 131 | fixes pfp :: "('a st \<Rightarrow> 'a st) \<Rightarrow> 'a st \<Rightarrow> 'a st"
 | 
| 132 | assumes pfp: "f(pfp f x) \<sqsubseteq> pfp f x" | |
| 133 | assumes above: "x \<sqsubseteq> pfp f x" | |
| 44656 | 134 | begin | 
| 135 | ||
| 45212 | 136 | fun aval' :: "aexp \<Rightarrow> 'a st \<Rightarrow> 'a" where | 
| 44656 | 137 | "aval' (N n) _ = num' n" | | 
| 138 | "aval' (V x) S = S x" | | |
| 139 | "aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" | |
| 140 | ||
| 141 | abbreviation fun_in_rep (infix "<:" 50) where | |
| 44923 | 142 | "f <: F == \<forall>x. f x <: F x" | 
| 44656 | 143 | |
| 144 | lemma fun_in_rep_le: "(s::state) <: S \<Longrightarrow> S \<sqsubseteq> T \<Longrightarrow> s <: T" | |
| 145 | by (metis le_fun_def le_rep subsetD) | |
| 146 | ||
| 147 | lemma aval'_sound: "s <: S \<Longrightarrow> aval a s <: aval' a S" | |
| 148 | by (induct a) (auto simp: rep_num' rep_plus') | |
| 149 | ||
| 45212 | 150 | fun AI :: "com \<Rightarrow> 'a st \<Rightarrow> 'a st" where | 
| 44656 | 151 | "AI SKIP S = S" | | 
| 152 | "AI (x ::= a) S = S(x := aval' a S)" | | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
48480diff
changeset | 153 | "AI (c1;;c2) S = AI c2 (AI c1 S)" | | 
| 44656 | 154 | "AI (IF b THEN c1 ELSE c2) S = (AI c1 S) \<squnion> (AI c2 S)" | | 
| 44944 | 155 | "AI (WHILE b DO c) S = pfp (AI c) S" | 
| 44656 | 156 | |
| 157 | lemma AI_sound: "(c,s) \<Rightarrow> t \<Longrightarrow> s <: S0 \<Longrightarrow> t <: AI c S0" | |
| 45015 | 158 | proof(induction c arbitrary: s t S0) | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44656diff
changeset | 159 | case SKIP thus ?case by fastforce | 
| 44656 | 160 | next | 
| 161 | case Assign thus ?case by (auto simp: aval'_sound) | |
| 162 | next | |
| 47818 | 163 | case Seq thus ?case by auto | 
| 44656 | 164 | next | 
| 165 | case If thus ?case by(auto simp: in_rep_join) | |
| 166 | next | |
| 167 | case (While b c) | |
| 44944 | 168 | let ?P = "pfp (AI c) S0" | 
| 44656 | 169 |   { fix s t have "(WHILE b DO c,s) \<Rightarrow> t \<Longrightarrow> s <: ?P \<Longrightarrow> t <: ?P"
 | 
| 45015 | 170 | proof(induction "WHILE b DO c" s t rule: big_step_induct) | 
| 44656 | 171 | case WhileFalse thus ?case by simp | 
| 172 | next | |
| 45015 | 173 | case WhileTrue thus ?case by(metis While.IH pfp fun_in_rep_le) | 
| 44656 | 174 | qed | 
| 175 | } | |
| 44932 | 176 | with fun_in_rep_le[OF `s <: S0` above] | 
| 44656 | 177 | show ?case by (metis While(2) AI.simps(5)) | 
| 178 | qed | |
| 179 | ||
| 180 | end | |
| 181 | ||
| 182 | ||
| 183 | text{* Problem: not executable because of the comparison of abstract states,
 | |
| 184 | i.e. functions, in the post-fixedpoint computation. Need to implement | |
| 185 | abstract states concretely. *} | |
| 186 | ||
| 187 | end |