| author | wenzelm | 
| Sun, 24 Feb 2013 14:14:07 +0100 | |
| changeset 51267 | c68c1b89a0f1 | 
| parent 50180 | c6626861c31a | 
| child 51436 | 790310525e97 | 
| permissions | -rw-r--r-- | 
| 45812 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 3 | theory Live_True | |
| 4 | imports "~~/src/HOL/Library/While_Combinator" Vars Big_Step | |
| 5 | begin | |
| 6 | ||
| 7 | subsection "True Liveness Analysis" | |
| 8 | ||
| 9 | fun L :: "com \<Rightarrow> vname set \<Rightarrow> vname set" where | |
| 10 | "L SKIP X = X" | | |
| 50009 | 11 | "L (x ::= a) X = (if x \<in> X then X - {x} \<union> vars a else X)" |
 | 
| 45812 | 12 | "L (c\<^isub>1; c\<^isub>2) X = (L c\<^isub>1 \<circ> L c\<^isub>2) X" | | 
| 13 | "L (IF b THEN c\<^isub>1 ELSE c\<^isub>2) X = vars b \<union> L c\<^isub>1 X \<union> L c\<^isub>2 X" | | |
| 50009 | 14 | "L (WHILE b DO c) X = lfp(\<lambda>Y. vars b \<union> X \<union> L c Y)" | 
| 45812 | 15 | |
| 16 | lemma L_mono: "mono (L c)" | |
| 17 | proof- | |
| 18 |   { fix X Y have "X \<subseteq> Y \<Longrightarrow> L c X \<subseteq> L c Y"
 | |
| 19 | proof(induction c arbitrary: X Y) | |
| 20 | case (While b c) | |
| 21 | show ?case | |
| 22 | proof(simp, rule lfp_mono) | |
| 23 | fix Z show "vars b \<union> X \<union> L c Z \<subseteq> vars b \<union> Y \<union> L c Z" | |
| 24 | using While by auto | |
| 25 | qed | |
| 26 | next | |
| 27 | case If thus ?case by(auto simp: subset_iff) | |
| 28 | qed auto | |
| 29 | } thus ?thesis by(rule monoI) | |
| 30 | qed | |
| 31 | ||
| 32 | lemma mono_union_L: | |
| 50009 | 33 | "mono (\<lambda>Y. X \<union> L c Y)" | 
| 45812 | 34 | by (metis (no_types) L_mono mono_def order_eq_iff set_eq_subset sup_mono) | 
| 35 | ||
| 36 | lemma L_While_unfold: | |
| 37 | "L (WHILE b DO c) X = vars b \<union> X \<union> L c (L (WHILE b DO c) X)" | |
| 38 | by(metis lfp_unfold[OF mono_union_L] L.simps(5)) | |
| 39 | ||
| 40 | ||
| 41 | subsection "Soundness" | |
| 42 | ||
| 43 | theorem L_sound: | |
| 44 | "(c,s) \<Rightarrow> s' \<Longrightarrow> s = t on L c X \<Longrightarrow> | |
| 45 | \<exists> t'. (c,t) \<Rightarrow> t' & s' = t' on X" | |
| 46 | proof (induction arbitrary: X t rule: big_step_induct) | |
| 47 | case Skip then show ?case by auto | |
| 48 | next | |
| 49 | case Assign then show ?case | |
| 50 | by (auto simp: ball_Un) | |
| 51 | next | |
| 47818 | 52 | case (Seq c1 s1 s2 c2 s3 X t1) | 
| 53 | from Seq.IH(1) Seq.prems obtain t2 where | |
| 45812 | 54 | t12: "(c1, t1) \<Rightarrow> t2" and s2t2: "s2 = t2 on L c2 X" | 
| 55 | by simp blast | |
| 47818 | 56 | from Seq.IH(2)[OF s2t2] obtain t3 where | 
| 45812 | 57 | t23: "(c2, t2) \<Rightarrow> t3" and s3t3: "s3 = t3 on X" | 
| 58 | by auto | |
| 59 | show ?case using t12 t23 s3t3 by auto | |
| 60 | next | |
| 61 | case (IfTrue b s c1 s' c2) | |
| 50009 | 62 | hence "s = t on vars b" and "s = t on L c1 X" by auto | 
| 45812 | 63 | from bval_eq_if_eq_on_vars[OF this(1)] IfTrue(1) have "bval b t" by simp | 
| 50009 | 64 | from IfTrue.IH[OF `s = t on L c1 X`] obtain t' where | 
| 45812 | 65 | "(c1, t) \<Rightarrow> t'" "s' = t' on X" by auto | 
| 66 | thus ?case using `bval b t` by auto | |
| 67 | next | |
| 68 | case (IfFalse b s c2 s' c1) | |
| 69 | hence "s = t on vars b" "s = t on L c2 X" by auto | |
| 70 | from bval_eq_if_eq_on_vars[OF this(1)] IfFalse(1) have "~bval b t" by simp | |
| 50009 | 71 | from IfFalse.IH[OF `s = t on L c2 X`] obtain t' where | 
| 45812 | 72 | "(c2, t) \<Rightarrow> t'" "s' = t' on X" by auto | 
| 73 | thus ?case using `~bval b t` by auto | |
| 74 | next | |
| 75 | case (WhileFalse b s c) | |
| 76 | hence "~ bval b t" | |
| 77 | by (metis L_While_unfold UnI1 bval_eq_if_eq_on_vars) | |
| 78 | thus ?case using WhileFalse.prems L_While_unfold[of b c X] by auto | |
| 79 | next | |
| 80 | case (WhileTrue b s1 c s2 s3 X t1) | |
| 81 | let ?w = "WHILE b DO c" | |
| 82 | from `bval b s1` WhileTrue.prems have "bval b t1" | |
| 83 | by (metis L_While_unfold UnI1 bval_eq_if_eq_on_vars) | |
| 84 | have "s1 = t1 on L c (L ?w X)" using L_While_unfold WhileTrue.prems | |
| 85 | by (blast) | |
| 86 | from WhileTrue.IH(1)[OF this] obtain t2 where | |
| 87 | "(c, t1) \<Rightarrow> t2" "s2 = t2 on L ?w X" by auto | |
| 88 | from WhileTrue.IH(2)[OF this(2)] obtain t3 where "(?w,t2) \<Rightarrow> t3" "s3 = t3 on X" | |
| 89 | by auto | |
| 90 | with `bval b t1` `(c, t1) \<Rightarrow> t2` show ?case by auto | |
| 91 | qed | |
| 92 | ||
| 50009 | 93 | |
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 94 | subsection "Executability" | 
| 45812 | 95 | |
| 96 | instantiation com :: vars | |
| 97 | begin | |
| 98 | ||
| 99 | fun vars_com :: "com \<Rightarrow> vname set" where | |
| 100 | "vars SKIP = {}" |
 | |
| 101 | "vars (x::=e) = vars e" | | |
| 102 | "vars (c\<^isub>1; c\<^isub>2) = vars c\<^isub>1 \<union> vars c\<^isub>2" | | |
| 103 | "vars (IF b THEN c\<^isub>1 ELSE c\<^isub>2) = vars b \<union> vars c\<^isub>1 \<union> vars c\<^isub>2" | | |
| 104 | "vars (WHILE b DO c) = vars b \<union> vars c" | |
| 105 | ||
| 106 | instance .. | |
| 107 | ||
| 108 | end | |
| 109 | ||
| 110 | lemma L_subset_vars: "L c X \<subseteq> vars c \<union> X" | |
| 111 | proof(induction c arbitrary: X) | |
| 112 | case (While b c) | |
| 50009 | 113 | have "lfp(\<lambda>Y. vars b \<union> X \<union> L c Y) \<subseteq> vars b \<union> vars c \<union> X" | 
| 45812 | 114 | using While.IH[of "vars b \<union> vars c \<union> X"] | 
| 115 | by (auto intro!: lfp_lowerbound) | |
| 116 | thus ?case by simp | |
| 117 | qed auto | |
| 118 | ||
| 119 | lemma afinite[simp]: "finite(vars(a::aexp))" | |
| 120 | by (induction a) auto | |
| 121 | ||
| 122 | lemma bfinite[simp]: "finite(vars(b::bexp))" | |
| 123 | by (induction b) auto | |
| 124 | ||
| 125 | lemma cfinite[simp]: "finite(vars(c::com))" | |
| 126 | by (induction c) auto | |
| 127 | ||
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 128 | |
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 129 | text{* Make @{const L} executable by replacing @{const lfp} with the @{const
 | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 130 | while} combinator from theory @{theory While_Combinator}. The @{const while}
 | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 131 | combinator obeys the recursion equation | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 132 | @{thm[display] While_Combinator.while_unfold[no_vars]}
 | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 133 | and is thus executable. *} | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 134 | |
| 45812 | 135 | lemma L_While: fixes b c X | 
| 136 | assumes "finite X" defines "f == \<lambda>A. vars b \<union> X \<union> L c A" | |
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 137 | shows "L (WHILE b DO c) X = while (\<lambda>A. f A \<noteq> A) f {}" (is "_ = ?r")
 | 
| 45812 | 138 | proof - | 
| 139 | let ?V = "vars b \<union> vars c \<union> X" | |
| 140 | have "lfp f = ?r" | |
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 141 | proof(rule lfp_while[where C = "?V"]) | 
| 45812 | 142 | show "mono f" by(simp add: f_def mono_union_L) | 
| 143 | next | |
| 144 | fix Y show "Y \<subseteq> ?V \<Longrightarrow> f Y \<subseteq> ?V" | |
| 145 | unfolding f_def using L_subset_vars[of c] by blast | |
| 146 | next | |
| 147 | show "finite ?V" using `finite X` by simp | |
| 148 | qed | |
| 149 | thus ?thesis by (simp add: f_def) | |
| 150 | qed | |
| 151 | ||
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 152 | lemma L_While_set: "L (WHILE b DO c) (set xs) = | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 153 | (let f = (\<lambda>A. vars b \<union> set xs \<union> L c A) | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 154 |    in while (\<lambda>A. f A \<noteq> A) f {})"
 | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 155 | by(simp add: L_While del: L.simps(5)) | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 156 | |
| 50180 | 157 | text{* Replace the equation for @{text "L (WHILE \<dots>)"} by the executable @{thm[source] L_While_set}: *}
 | 
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 158 | lemmas [code] = L.simps(1-4) L_While_set | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 159 | text{* Sorry, this syntax is odd. *}
 | 
| 45812 | 160 | |
| 50180 | 161 | text{* A test: *}
 | 
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 162 | lemma "(let b = Less (N 0) (V ''y''); c = ''y'' ::= V ''x''; ''x'' ::= V ''z'' | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 163 |   in L (WHILE b DO c) {''y''}) = {''x'', ''y'', ''z''}"
 | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 164 | by eval | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 165 | |
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 166 | |
| 50009 | 167 | subsection "Limiting the number of iterations" | 
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 168 | |
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 169 | text{* The final parameter is the default value: *}
 | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 170 | |
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 171 | fun iter :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
| 45812 | 172 | "iter f 0 p d = d" | | 
| 173 | "iter f (Suc n) p d = (if f p = p then p else iter f n (f p) d)" | |
| 174 | ||
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 175 | text{* A version of @{const L} with a bounded number of iterations (here: 2)
 | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 176 | in the WHILE case: *} | 
| 45812 | 177 | |
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 178 | fun Lb :: "com \<Rightarrow> vname set \<Rightarrow> vname set" where | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 179 | "Lb SKIP X = X" | | 
| 50009 | 180 | "Lb (x ::= a) X = (if x \<in> X then X - {x} \<union> vars a else X)" |
 | 
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 181 | "Lb (c\<^isub>1; c\<^isub>2) X = (Lb c\<^isub>1 \<circ> Lb c\<^isub>2) X" | | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 182 | "Lb (IF b THEN c\<^isub>1 ELSE c\<^isub>2) X = vars b \<union> Lb c\<^isub>1 X \<union> Lb c\<^isub>2 X" | | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 183 | "Lb (WHILE b DO c) X = iter (\<lambda>A. vars b \<union> X \<union> Lb c A) 2 {} (vars b \<union> vars c \<union> X)"
 | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 184 | |
| 50009 | 185 | text{* @{const Lb} (and @{const iter}) is not monotone! *}
 | 
| 186 | lemma "let w = WHILE Bc False DO (''x'' ::= V ''y''; ''z'' ::= V ''x'')
 | |
| 187 |   in \<not> (Lb w {''z''} \<subseteq> Lb w {''y'',''z''})"
 | |
| 188 | by eval | |
| 189 | ||
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 190 | lemma lfp_subset_iter: | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 191 | "\<lbrakk> mono f; !!X. f X \<subseteq> f' X; lfp f \<subseteq> D \<rbrakk> \<Longrightarrow> lfp f \<subseteq> iter f' n A D" | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 192 | proof(induction n arbitrary: A) | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 193 | case 0 thus ?case by simp | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 194 | next | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 195 | case Suc thus ?case by simp (metis lfp_lowerbound) | 
| 45812 | 196 | qed | 
| 197 | ||
| 50007 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 198 | lemma "L c X \<subseteq> Lb c X" | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 199 | proof(induction c arbitrary: X) | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 200 | case (While b c) | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 201 | let ?f = "\<lambda>A. vars b \<union> X \<union> L c A" | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 202 | let ?fb = "\<lambda>A. vars b \<union> X \<union> Lb c A" | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 203 | show ?case | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 204 | proof (simp, rule lfp_subset_iter[OF mono_union_L]) | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 205 | show "!!X. ?f X \<subseteq> ?fb X" using While.IH by blast | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 206 | show "lfp ?f \<subseteq> vars b \<union> vars c \<union> X" | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 207 | by (metis (full_types) L.simps(5) L_subset_vars vars_com.simps(5)) | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 208 | qed | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 209 | next | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 210 | case Seq thus ?case by simp (metis (full_types) L_mono monoD subset_trans) | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 211 | qed auto | 
| 
56f269baae76
executable true liveness analysis incl an approximating version
 nipkow parents: 
48256diff
changeset | 212 | |
| 45812 | 213 | end |