| author | blanchet | 
| Tue, 10 Sep 2013 15:56:51 +0200 | |
| changeset 53503 | d2f21e305d0c | 
| parent 53015 | a1119cf551e8 | 
| child 55010 | 203b4f5482c3 | 
| permissions | -rw-r--r-- | 
| 47613 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 3 | theory Abs_Int1 | |
| 4 | imports Abs_State | |
| 5 | begin | |
| 6 | ||
| 7 | subsection "Computable Abstract Interpretation" | |
| 8 | ||
| 51390 | 9 | text{* Abstract interpretation over type @{text st} instead of functions. *}
 | 
| 47613 | 10 | |
| 52504 | 11 | context Gamma_semilattice | 
| 47613 | 12 | begin | 
| 13 | ||
| 14 | fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where | |
| 50896 | 15 | "aval' (N i) S = num' i" | | 
| 47613 | 16 | "aval' (V x) S = fun S x" | | 
| 17 | "aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" | |
| 18 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 19 | lemma aval'_correct: "s : \<gamma>\<^sub>s S \<Longrightarrow> aval a s : \<gamma>(aval' a S)" | 
| 47613 | 20 | by (induction a) (auto simp: gamma_num' gamma_plus' \<gamma>_st_def) | 
| 21 | ||
| 51390 | 22 | lemma gamma_Step_subcomm: fixes C1 C2 :: "'a::semilattice_sup acom" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 23 | assumes "!!x e S. f1 x e (\<gamma>\<^sub>o S) \<subseteq> \<gamma>\<^sub>o (f2 x e S)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 24 | "!!b S. g1 b (\<gamma>\<^sub>o S) \<subseteq> \<gamma>\<^sub>o (g2 b S)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 25 | shows "Step f1 g1 (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (Step f2 g2 S C)" | 
| 51390 | 26 | proof(induction C arbitrary: S) | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 27 | qed (auto simp: assms intro!: mono_gamma_o sup_ge1 sup_ge2) | 
| 51390 | 28 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 29 | lemma in_gamma_update: "\<lbrakk> s : \<gamma>\<^sub>s S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^sub>s(update S x a)" | 
| 51390 | 30 | by(simp add: \<gamma>_st_def) | 
| 31 | ||
| 47613 | 32 | end | 
| 33 | ||
| 51390 | 34 | |
| 52504 | 35 | locale Abs_Int = Gamma_semilattice where \<gamma>=\<gamma> | 
| 51826 | 36 | for \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set" | 
| 47613 | 37 | begin | 
| 38 | ||
| 51390 | 39 | definition "step' = Step | 
| 51389 | 40 | (\<lambda>x e S. case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S))) | 
| 41 | (\<lambda>b S. S)" | |
| 42 | ||
| 47613 | 43 | definition AI :: "com \<Rightarrow> 'av st option acom option" where | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 44 | "AI c = pfp (step' \<top>) (bot c)" | 
| 47613 | 45 | |
| 46 | ||
| 47 | lemma strip_step'[simp]: "strip(step' S C) = strip C" | |
| 51390 | 48 | by(simp add: step'_def) | 
| 47613 | 49 | |
| 50 | ||
| 51974 | 51 | text{* Correctness: *}
 | 
| 47613 | 52 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 53 | lemma step_step': "step (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' S C)" | 
| 51390 | 54 | unfolding step_def step'_def | 
| 55 | by(rule gamma_Step_subcomm) | |
| 51974 | 56 | (auto simp: intro!: aval'_correct in_gamma_update split: option.splits) | 
| 47613 | 57 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 58 | lemma AI_correct: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c C" | 
| 47613 | 59 | proof(simp add: CS_def AI_def) | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 60 | assume 1: "pfp (step' \<top>) (bot c) = Some C" | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 61 | have pfp': "step' \<top> C \<le> C" by(rule pfp_pfp[OF 1]) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 62 | have 2: "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c C" --"transfer the pfp'" | 
| 50986 | 63 | 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 | 64 | 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 | 65 | show "... \<le> \<gamma>\<^sub>c C" by (metis mono_gamma_c[OF pfp']) | 
| 47613 | 66 | qed | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 67 | 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 | 68 | 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 | 69 | 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 | 70 | thus "lfp c (step UNIV) \<le> \<gamma>\<^sub>c C" by simp | 
| 47613 | 71 | qed | 
| 72 | ||
| 73 | end | |
| 74 | ||
| 75 | ||
| 76 | subsubsection "Monotonicity" | |
| 77 | ||
| 78 | locale Abs_Int_mono = Abs_Int + | |
| 51359 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 nipkow parents: 
51036diff
changeset | 79 | assumes mono_plus': "a1 \<le> b1 \<Longrightarrow> a2 \<le> b2 \<Longrightarrow> plus' a1 a2 \<le> plus' b1 b2" | 
| 47613 | 80 | begin | 
| 81 | ||
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 82 | lemma mono_aval': "S1 \<le> S2 \<Longrightarrow> aval' e S1 \<le> aval' e S2" | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 83 | by(induction e) (auto simp: mono_plus' mono_fun) | 
| 47613 | 84 | |
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 85 | theorem mono_step': "S1 \<le> S2 \<Longrightarrow> C1 \<le> C2 \<Longrightarrow> step' S1 C1 \<le> step' S2 C2" | 
| 51390 | 86 | unfolding step'_def | 
| 87 | by(rule mono2_Step) (auto simp: mono_aval' split: option.split) | |
| 47613 | 88 | |
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 89 | lemma mono_step'_top: "C \<le> C' \<Longrightarrow> step' \<top> C \<le> step' \<top> C'" | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 90 | by (metis mono_step' order_refl) | 
| 47613 | 91 | |
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 92 | lemma AI_least_pfp: assumes "AI c = Some C" "step' \<top> C' \<le> C'" "strip C' = c" | 
| 51359 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 nipkow parents: 
51036diff
changeset | 93 | shows "C \<le> C'" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 94 | by(rule pfp_bot_least[OF _ _ assms(2,3) assms(1)[unfolded AI_def]]) | 
| 51390 | 95 | (simp_all add: mono_step'_top) | 
| 49464 | 96 | |
| 47613 | 97 | end | 
| 98 | ||
| 99 | ||
| 100 | subsubsection "Termination" | |
| 101 | ||
| 49547 | 102 | locale Measure1 = | 
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
52504diff
changeset | 103 | fixes m :: "'av::{order,order_top} \<Rightarrow> nat"
 | 
| 47613 | 104 | fixes h :: "nat" | 
| 105 | assumes h: "m x \<le> h" | |
| 106 | begin | |
| 107 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 108 | definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>s") where
 | 
| 51791 | 109 | "m_s S X = (\<Sum> x \<in> X. m(fun S x))" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 110 | |
| 51791 | 111 | lemma m_s_h: "finite X \<Longrightarrow> m_s S X \<le> h * card X" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 112 | by(simp add: m_s_def) (metis nat_mult_commute of_nat_id setsum_bounded[OF h]) | 
| 47613 | 113 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 114 | definition m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>o") where
 | 
| 51791 | 115 | "m_o opt X = (case opt of None \<Rightarrow> h * card X + 1 | Some S \<Rightarrow> m_s S X)" | 
| 49431 | 116 | |
| 51791 | 117 | lemma m_o_h: "finite X \<Longrightarrow> m_o opt X \<le> (h*card X + 1)" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 118 | by(auto simp add: m_o_def m_s_h le_SucI split: option.split dest:m_s_h) | 
| 47613 | 119 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 120 | definition m_c :: "'av st option acom \<Rightarrow> nat" ("m\<^sub>c") where
 | 
| 51791 | 121 | "m_c C = listsum (map (\<lambda>a. m_o a (vars C)) (annos C))" | 
| 49431 | 122 | |
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 123 | text{* Upper complexity bound: *}
 | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 124 | lemma m_c_h: "m_c C \<le> size(annos C) * (h * card(vars C) + 1)" | 
| 49431 | 125 | proof- | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 126 | let ?X = "vars C" let ?n = "card ?X" let ?a = "size(annos C)" | 
| 51791 | 127 | have "m_c C = (\<Sum>i<?a. m_o (annos C ! i) ?X)" | 
| 51786 | 128 | by(simp add: m_c_def listsum_setsum_nth atLeast0LessThan) | 
| 49431 | 129 | also have "\<dots> \<le> (\<Sum>i<?a. h * ?n + 1)" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 130 | apply(rule setsum_mono) using m_o_h[OF finite_Cvars] by simp | 
| 49431 | 131 | also have "\<dots> = ?a * (h * ?n + 1)" by simp | 
| 132 | finally show ?thesis . | |
| 133 | qed | |
| 134 | ||
| 49547 | 135 | end | 
| 136 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 137 | fun top_on_st :: "'a::order_top st \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>s") where
 | 
| 51785 | 138 | "top_on_st S X = (\<forall>x\<in>X. fun S x = \<top>)" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 139 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 140 | fun top_on_opt :: "'a::order_top st option \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>o") where
 | 
| 51785 | 141 | "top_on_opt (Some S) X = top_on_st S X" | | 
| 142 | "top_on_opt None X = True" | |
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 143 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52729diff
changeset | 144 | definition top_on_acom :: "'a::order_top st option acom \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>c") where
 | 
| 51785 | 145 | "top_on_acom C X = (\<forall>a \<in> set(annos C). top_on_opt a X)" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 146 | |
| 51785 | 147 | lemma top_on_top: "top_on_opt (\<top>::_ st option) X" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 148 | by(auto simp: top_option_def fun_top) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 149 | |
| 51785 | 150 | lemma top_on_bot: "top_on_acom (bot c) X" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 151 | by(auto simp add: top_on_acom_def bot_def) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 152 | |
| 51785 | 153 | lemma top_on_post: "top_on_acom C X \<Longrightarrow> top_on_opt (post C) X" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 154 | by(simp add: top_on_acom_def post_in_annos) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 155 | |
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 156 | lemma top_on_acom_simps: | 
| 51785 | 157 |   "top_on_acom (SKIP {Q}) X = top_on_opt Q X"
 | 
| 158 |   "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: 
51974diff
changeset | 159 | "top_on_acom (C1;;C2) X = (top_on_acom C1 X \<and> top_on_acom C2 X)" | 
| 51785 | 160 |   "top_on_acom (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) X =
 | 
| 161 | (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)" | |
| 162 |   "top_on_acom ({I} WHILE b DO {P} C {Q}) X =
 | |
| 163 | (top_on_opt I X \<and> top_on_acom C X \<and> top_on_opt P X \<and> top_on_opt Q X)" | |
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 164 | by(auto simp add: top_on_acom_def) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 165 | |
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 166 | lemma top_on_sup: | 
| 51785 | 167 | "top_on_opt o1 X \<Longrightarrow> top_on_opt o2 X \<Longrightarrow> top_on_opt (o1 \<squnion> o2 :: _ st option) X" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 168 | apply(induction o1 o2 rule: sup_option.induct) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 169 | apply(auto) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 170 | by transfer simp | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 171 | |
| 51826 | 172 | lemma top_on_Step: fixes C :: "('a::semilattice_sup_top)st option acom"
 | 
| 51785 | 173 | 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" | 
| 174 | "!!b S. top_on_opt S X \<Longrightarrow> vars b \<subseteq> -X \<Longrightarrow> top_on_opt (g b S) X" | |
| 175 | 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" | |
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 176 | proof(induction C arbitrary: S) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 177 | qed (auto simp: top_on_acom_simps vars_acom_def top_on_post top_on_sup assms) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 178 | |
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 179 | |
| 49547 | 180 | locale Measure = Measure1 + | 
| 51359 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 nipkow parents: 
51036diff
changeset | 181 | assumes m2: "x < y \<Longrightarrow> m x > m y" | 
| 49547 | 182 | begin | 
| 183 | ||
| 51372 | 184 | lemma m1: "x \<le> y \<Longrightarrow> m x \<ge> m y" | 
| 185 | by(auto simp: le_less m2) | |
| 186 | ||
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 187 | lemma m_s2_rep: assumes "finite(X)" and "S1 = S2 on -X" and "\<forall>x. S1 x \<le> S2 x" and "S1 \<noteq> S2" | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 188 | shows "(\<Sum>x\<in>X. m (S2 x)) < (\<Sum>x\<in>X. m (S1 x))" | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 189 | proof- | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 190 | from assms(3) have 1: "\<forall>x\<in>X. m(S1 x) \<ge> m(S2 x)" by (simp add: m1) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 191 | from assms(2,3,4) have "EX x:X. S1 x < S2 x" | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 192 | by(simp add: fun_eq_iff) (metis Compl_iff le_neq_trans) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 193 | hence 2: "\<exists>x\<in>X. m(S1 x) > m(S2 x)" by (metis m2) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 194 | from setsum_strict_mono_ex1[OF `finite X` 1 2] | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 195 | show "(\<Sum>x\<in>X. m (S2 x)) < (\<Sum>x\<in>X. m (S1 x))" . | 
| 49547 | 196 | qed | 
| 197 | ||
| 51791 | 198 | lemma m_s2: "finite(X) \<Longrightarrow> fun S1 = fun S2 on -X | 
| 199 | \<Longrightarrow> S1 < S2 \<Longrightarrow> m_s S1 X > m_s S2 X" | |
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 200 | apply(auto simp add: less_st_def m_s_def) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 201 | apply (transfer fixing: m) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 202 | apply(simp add: less_eq_st_rep_iff eq_st_def m_s2_rep) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 203 | done | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 204 | |
| 51785 | 205 | lemma m_o2: "finite X \<Longrightarrow> top_on_opt o1 (-X) \<Longrightarrow> top_on_opt o2 (-X) \<Longrightarrow> | 
| 51791 | 206 | o1 < o2 \<Longrightarrow> m_o o1 X > m_o o2 X" | 
| 51359 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 nipkow parents: 
51036diff
changeset | 207 | proof(induction o1 o2 rule: less_eq_option.induct) | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 208 | case 1 thus ?case by (auto simp: m_o_def m_s2 less_option_def) | 
| 49547 | 209 | next | 
| 51359 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 nipkow parents: 
51036diff
changeset | 210 | case 2 thus ?case by(auto simp: m_o_def less_option_def le_imp_less_Suc m_s_h) | 
| 49547 | 211 | next | 
| 51359 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 nipkow parents: 
51036diff
changeset | 212 | case 3 thus ?case by (auto simp: less_option_def) | 
| 49547 | 213 | qed | 
| 214 | ||
| 51785 | 215 | lemma m_o1: "finite X \<Longrightarrow> top_on_opt o1 (-X) \<Longrightarrow> top_on_opt o2 (-X) \<Longrightarrow> | 
| 51791 | 216 | o1 \<le> o2 \<Longrightarrow> m_o o1 X \<ge> m_o o2 X" | 
| 51372 | 217 | by(auto simp: le_less m_o2) | 
| 218 | ||
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 219 | |
| 51785 | 220 | lemma m_c2: "top_on_acom C1 (-vars C1) \<Longrightarrow> top_on_acom C2 (-vars C2) \<Longrightarrow> | 
| 51359 
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
 nipkow parents: 
51036diff
changeset | 221 | C1 < C2 \<Longrightarrow> m_c C1 > m_c C2" | 
| 51786 | 222 | proof(auto simp add: le_iff_le_annos size_annos_same[of C1 C2] vars_acom_def less_acom_def) | 
| 49396 | 223 | let ?X = "vars(strip C2)" | 
| 51785 | 224 | assume top: "top_on_acom C1 (- vars(strip C2))" "top_on_acom C2 (- vars(strip C2))" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 225 | and strip_eq: "strip C1 = strip C2" | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 226 | and 0: "\<forall>i<size(annos C2). annos C1 ! i \<le> annos C2 ! i" | 
| 51791 | 227 | hence 1: "\<forall>i<size(annos C2). m_o (annos C1 ! i) ?X \<ge> m_o (annos C2 ! i) ?X" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 228 | apply (auto simp: all_set_conv_all_nth vars_acom_def top_on_acom_def) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 229 | by (metis finite_cvars m_o1 size_annos_same2) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 230 | fix i assume i: "i < size(annos C2)" "\<not> annos C2 ! i \<le> annos C1 ! i" | 
| 51785 | 231 | have topo1: "top_on_opt (annos C1 ! i) (- ?X)" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 232 | using i(1) top(1) by(simp add: top_on_acom_def size_annos_same[OF strip_eq]) | 
| 51785 | 233 | have topo2: "top_on_opt (annos C2 ! i) (- ?X)" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 234 | using i(1) top(2) by(simp add: top_on_acom_def size_annos_same[OF strip_eq]) | 
| 51791 | 235 | from i have "m_o (annos C1 ! i) ?X > m_o (annos C2 ! i) ?X" (is "?P i") | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 236 | by (metis 0 less_option_def m_o2[OF finite_cvars topo1] topo2) | 
| 49396 | 237 | hence 2: "\<exists>i < size(annos C2). ?P i" using `i < size(annos C2)` by blast | 
| 51791 | 238 | have "(\<Sum>i<size(annos C2). m_o (annos C2 ! i) ?X) | 
| 239 | < (\<Sum>i<size(annos C2). m_o (annos C1 ! i) ?X)" | |
| 47613 | 240 | apply(rule setsum_strict_mono_ex1) using 1 2 by (auto) | 
| 51786 | 241 | thus ?thesis | 
| 242 | by(simp add: m_c_def vars_acom_def strip_eq listsum_setsum_nth atLeast0LessThan size_annos_same[OF strip_eq]) | |
| 47613 | 243 | qed | 
| 244 | ||
| 49547 | 245 | end | 
| 246 | ||
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 247 | |
| 49547 | 248 | locale Abs_Int_measure = | 
| 249 | Abs_Int_mono where \<gamma>=\<gamma> + Measure where m=m | |
| 51826 | 250 | for \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set" and m :: "'av \<Rightarrow> nat" | 
| 49547 | 251 | begin | 
| 252 | ||
| 51785 | 253 | lemma top_on_step': "\<lbrakk> top_on_acom C (-vars C) \<rbrakk> \<Longrightarrow> top_on_acom (step' \<top> C) (-vars C)" | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 254 | unfolding step'_def | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 255 | by(rule top_on_Step) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 256 | (auto simp add: top_option_def fun_top split: option.splits) | 
| 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 257 | |
| 47613 | 258 | lemma AI_Some_measure: "\<exists>C. AI c = Some C" | 
| 259 | unfolding AI_def | |
| 51785 | 260 | apply(rule pfp_termination[where I = "\<lambda>C. top_on_acom C (- vars C)" and m="m_c"]) | 
| 51711 
df3426139651
complete revision: finally got rid of annoying L-predicate
 nipkow parents: 
51390diff
changeset | 261 | apply(simp_all add: m_c2 mono_step'_top bot_least top_on_bot) | 
| 51785 | 262 | using top_on_step' apply(auto simp add: vars_acom_def) | 
| 47613 | 263 | done | 
| 264 | ||
| 265 | end | |
| 266 | ||
| 267 | end |