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