| author | huffman | 
| Mon, 20 Apr 2009 17:38:25 -0700 | |
| changeset 30916 | a3d2128cac92 | 
| parent 27362 | a6dc1769fdda | 
| child 30952 | 7ab2716dd93b | 
| permissions | -rw-r--r-- | 
| 12431 | 1 | (* Title: HOL/IMP/Transition.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow & Robert Sandner, TUM | |
| 4 | Isar Version: Gerwin Klein, 2001 | |
| 5 | Copyright 1996 TUM | |
| 1700 | 6 | *) | 
| 7 | ||
| 12431 | 8 | header "Transition Semantics of Commands" | 
| 9 | ||
| 16417 | 10 | theory Transition imports Natural begin | 
| 12431 | 11 | |
| 12 | subsection "The transition relation" | |
| 1700 | 13 | |
| 12431 | 14 | text {*
 | 
| 15 |   We formalize the transition semantics as in \cite{Nielson}. This
 | |
| 16 | makes some of the rules a bit more intuitive, but also requires | |
| 17 | some more (internal) formal overhead. | |
| 18372 | 18 | |
| 19 | Since configurations that have terminated are written without | |
| 20 | a statement, the transition relation is not | |
| 12431 | 21 |   @{typ "((com \<times> state) \<times> (com \<times> state)) set"}
 | 
| 22 | but instead: | |
| 23746 | 23 |   @{typ "((com option \<times> state) \<times> (com option \<times> state)) set"}
 | 
| 4906 | 24 | |
| 18372 | 25 | Some syntactic sugar that we will use to hide the | 
| 12431 | 26 |   @{text option} part in configurations:
 | 
| 27 | *} | |
| 27362 | 28 | abbreviation | 
| 29 |   angle :: "[com, state] \<Rightarrow> com option \<times> state" ("<_,_>") where
 | |
| 30 | "<c,s> == (Some c, s)" | |
| 31 | abbreviation | |
| 32 |   angle2 :: "state \<Rightarrow> com option \<times> state"  ("<_>") where
 | |
| 33 | "<s> == (None, s)" | |
| 12431 | 34 | |
| 27362 | 35 | notation (xsymbols) | 
| 36 |   angle  ("\<langle>_,_\<rangle>") and
 | |
| 37 |   angle2  ("\<langle>_\<rangle>")
 | |
| 14565 | 38 | |
| 27362 | 39 | notation (HTML output) | 
| 40 |   angle  ("\<langle>_,_\<rangle>") and
 | |
| 41 |   angle2  ("\<langle>_\<rangle>")
 | |
| 12431 | 42 | |
| 43 | text {*
 | |
| 23746 | 44 | Now, finally, we are set to write down the rules for our small step semantics: | 
| 45 | *} | |
| 46 | inductive_set | |
| 47 | evalc1 :: "((com option \<times> state) \<times> (com option \<times> state)) set" | |
| 48 | and evalc1' :: "[(com option\<times>state),(com option\<times>state)] \<Rightarrow> bool" | |
| 49 |     ("_ \<longrightarrow>\<^sub>1 _" [60,60] 61)
 | |
| 50 | where | |
| 51 | "cs \<longrightarrow>\<^sub>1 cs' == (cs,cs') \<in> evalc1" | |
| 52 | | Skip: "\<langle>\<SKIP>, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s\<rangle>" | |
| 53 | | Assign: "\<langle>x :== a, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s[x \<mapsto> a s]\<rangle>" | |
| 54 | ||
| 55 | | Semi1: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s'\<rangle> \<Longrightarrow> \<langle>c0;c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1,s'\<rangle>" | |
| 56 | | Semi2: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c0',s'\<rangle> \<Longrightarrow> \<langle>c0;c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c0';c1,s'\<rangle>" | |
| 57 | ||
| 58 | | IfTrue: "b s \<Longrightarrow> \<langle>\<IF> b \<THEN> c1 \<ELSE> c2,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1,s\<rangle>" | |
| 59 | | IfFalse: "\<not>b s \<Longrightarrow> \<langle>\<IF> b \<THEN> c1 \<ELSE> c2,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c2,s\<rangle>" | |
| 60 | ||
| 61 | | While: "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>,s\<rangle>" | |
| 62 | ||
| 63 | lemmas [intro] = evalc1.intros -- "again, use these rules in automatic proofs" | |
| 64 | ||
| 65 | text {*
 | |
| 12431 | 66 | More syntactic sugar for the transition relation, and its | 
| 67 | iteration. | |
| 68 | *} | |
| 23746 | 69 | abbreviation | 
| 70 | evalcn :: "[(com option\<times>state),nat,(com option\<times>state)] \<Rightarrow> bool" | |
| 71 |     ("_ -_\<rightarrow>\<^sub>1 _" [60,60,60] 60)  where
 | |
| 72 | "cs -n\<rightarrow>\<^sub>1 cs' == (cs,cs') \<in> evalc1^n" | |
| 12431 | 73 | |
| 23746 | 74 | abbreviation | 
| 75 | evalc' :: "[(com option\<times>state),(com option\<times>state)] \<Rightarrow> bool" | |
| 76 |     ("_ \<longrightarrow>\<^sub>1\<^sup>* _" [60,60] 60)  where
 | |
| 77 | "cs \<longrightarrow>\<^sub>1\<^sup>* cs' == (cs,cs') \<in> evalc1^*" | |
| 12431 | 78 | |
| 79 | (*<*) | |
| 80 | (* fixme: move to Relation_Power.thy *) | |
| 81 | lemma rel_pow_Suc_E2 [elim!]: | |
| 82 | "[| (x, z) \<in> R ^ Suc n; !!y. [| (x, y) \<in> R; (y, z) \<in> R ^ n |] ==> P |] ==> P" | |
| 18372 | 83 | by (blast dest: rel_pow_Suc_D2) | 
| 12431 | 84 | |
| 85 | lemma rtrancl_imp_rel_pow: "p \<in> R^* \<Longrightarrow> \<exists>n. p \<in> R^n" | |
| 18372 | 86 | proof (induct p) | 
| 87 | fix x y | |
| 88 | assume "(x, y) \<in> R\<^sup>*" | |
| 89 | thus "\<exists>n. (x, y) \<in> R^n" | |
| 12431 | 90 | proof induct | 
| 18372 | 91 | fix a have "(a, a) \<in> R^0" by simp | 
| 92 | thus "\<exists>n. (a, a) \<in> R ^ n" .. | |
| 12431 | 93 | next | 
| 18372 | 94 | fix a b c assume "\<exists>n. (a, b) \<in> R ^ n" | 
| 95 | then obtain n where "(a, b) \<in> R^n" .. | |
| 96 | moreover assume "(b, c) \<in> R" | |
| 97 | ultimately have "(a, c) \<in> R^(Suc n)" by auto | |
| 98 | thus "\<exists>n. (a, c) \<in> R^n" .. | |
| 12431 | 99 | qed | 
| 18372 | 100 | qed | 
| 101 | (*>*) | |
| 12431 | 102 | text {*
 | 
| 18372 | 103 | As for the big step semantics you can read these rules in a | 
| 12431 | 104 | syntax directed way: | 
| 105 | *} | |
| 18372 | 106 | lemma SKIP_1: "\<langle>\<SKIP>, s\<rangle> \<longrightarrow>\<^sub>1 y = (y = \<langle>s\<rangle>)" | 
| 23746 | 107 | by (induct y, rule, cases set: evalc1, auto) | 
| 12431 | 108 | |
| 109 | lemma Assign_1: "\<langle>x :== a, s\<rangle> \<longrightarrow>\<^sub>1 y = (y = \<langle>s[x \<mapsto> a s]\<rangle>)" | |
| 23746 | 110 | by (induct y, rule, cases set: evalc1, auto) | 
| 12431 | 111 | |
| 18372 | 112 | lemma Cond_1: | 
| 12431 | 113 | "\<langle>\<IF> b \<THEN> c1 \<ELSE> c2, s\<rangle> \<longrightarrow>\<^sub>1 y = ((b s \<longrightarrow> y = \<langle>c1, s\<rangle>) \<and> (\<not>b s \<longrightarrow> y = \<langle>c2, s\<rangle>))" | 
| 23746 | 114 | by (induct y, rule, cases set: evalc1, auto) | 
| 12431 | 115 | |
| 18372 | 116 | lemma While_1: | 
| 12434 | 117 | "\<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>1 y = (y = \<langle>\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>, s\<rangle>)" | 
| 23746 | 118 | by (induct y, rule, cases set: evalc1, auto) | 
| 12431 | 119 | |
| 120 | lemmas [simp] = SKIP_1 Assign_1 Cond_1 While_1 | |
| 121 | ||
| 122 | ||
| 123 | subsection "Examples" | |
| 124 | ||
| 18372 | 125 | lemma | 
| 12434 | 126 | "s x = 0 \<Longrightarrow> \<langle>\<WHILE> \<lambda>s. s x \<noteq> 1 \<DO> (x:== \<lambda>s. s x+1), s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s[x \<mapsto> 1]\<rangle>" | 
| 12431 | 127 | (is "_ \<Longrightarrow> \<langle>?w, _\<rangle> \<longrightarrow>\<^sub>1\<^sup>* _") | 
| 128 | proof - | |
| 12434 | 129 | let ?c = "x:== \<lambda>s. s x+1" | 
| 130 | let ?if = "\<IF> \<lambda>s. s x \<noteq> 1 \<THEN> ?c; ?w \<ELSE> \<SKIP>" | |
| 12431 | 131 | assume [simp]: "s x = 0" | 
| 132 | have "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>?if, s\<rangle>" .. | |
| 18372 | 133 | also have "\<langle>?if, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>?c; ?w, s\<rangle>" by simp | 
| 134 | also have "\<langle>?c; ?w, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>?w, s[x \<mapsto> 1]\<rangle>" by (rule Semi1) simp | |
| 12431 | 135 | also have "\<langle>?w, s[x \<mapsto> 1]\<rangle> \<longrightarrow>\<^sub>1 \<langle>?if, s[x \<mapsto> 1]\<rangle>" .. | 
| 136 | also have "\<langle>?if, s[x \<mapsto> 1]\<rangle> \<longrightarrow>\<^sub>1 \<langle>\<SKIP>, s[x \<mapsto> 1]\<rangle>" by (simp add: update_def) | |
| 137 | also have "\<langle>\<SKIP>, s[x \<mapsto> 1]\<rangle> \<longrightarrow>\<^sub>1 \<langle>s[x \<mapsto> 1]\<rangle>" .. | |
| 138 | finally show ?thesis .. | |
| 139 | qed | |
| 4906 | 140 | |
| 18372 | 141 | lemma | 
| 12434 | 142 | "s x = 2 \<Longrightarrow> \<langle>\<WHILE> \<lambda>s. s x \<noteq> 1 \<DO> (x:== \<lambda>s. s x+1), s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* s'" | 
| 12431 | 143 | (is "_ \<Longrightarrow> \<langle>?w, _\<rangle> \<longrightarrow>\<^sub>1\<^sup>* s'") | 
| 144 | proof - | |
| 145 | let ?c = "x:== \<lambda>s. s x+1" | |
| 18372 | 146 | let ?if = "\<IF> \<lambda>s. s x \<noteq> 1 \<THEN> ?c; ?w \<ELSE> \<SKIP>" | 
| 12431 | 147 | assume [simp]: "s x = 2" | 
| 148 | note update_def [simp] | |
| 149 | have "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>?if, s\<rangle>" .. | |
| 150 | also have "\<langle>?if, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>?c; ?w, s\<rangle>" by simp | |
| 18372 | 151 | also have "\<langle>?c; ?w, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>?w, s[x \<mapsto> 3]\<rangle>" by (rule Semi1) simp | 
| 12431 | 152 | also have "\<langle>?w, s[x \<mapsto> 3]\<rangle> \<longrightarrow>\<^sub>1 \<langle>?if, s[x \<mapsto> 3]\<rangle>" .. | 
| 153 | also have "\<langle>?if, s[x \<mapsto> 3]\<rangle> \<longrightarrow>\<^sub>1 \<langle>?c; ?w, s[x \<mapsto> 3]\<rangle>" by simp | |
| 18372 | 154 | also have "\<langle>?c; ?w, s[x \<mapsto> 3]\<rangle> \<longrightarrow>\<^sub>1 \<langle>?w, s[x \<mapsto> 4]\<rangle>" by (rule Semi1) simp | 
| 12431 | 155 | also have "\<langle>?w, s[x \<mapsto> 4]\<rangle> \<longrightarrow>\<^sub>1 \<langle>?if, s[x \<mapsto> 4]\<rangle>" .. | 
| 156 | also have "\<langle>?if, s[x \<mapsto> 4]\<rangle> \<longrightarrow>\<^sub>1 \<langle>?c; ?w, s[x \<mapsto> 4]\<rangle>" by simp | |
| 18372 | 157 | also have "\<langle>?c; ?w, s[x \<mapsto> 4]\<rangle> \<longrightarrow>\<^sub>1 \<langle>?w, s[x \<mapsto> 5]\<rangle>" by (rule Semi1) simp | 
| 12431 | 158 | oops | 
| 159 | ||
| 160 | subsection "Basic properties" | |
| 161 | ||
| 18372 | 162 | text {*
 | 
| 12431 | 163 |   There are no \emph{stuck} programs:
 | 
| 164 | *} | |
| 165 | lemma no_stuck: "\<exists>y. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>1 y" | |
| 166 | proof (induct c) | |
| 167 | -- "case Semi:" | |
| 18372 | 168 | fix c1 c2 assume "\<exists>y. \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>1 y" | 
| 12431 | 169 | then obtain y where "\<langle>c1,s\<rangle> \<longrightarrow>\<^sub>1 y" .. | 
| 170 | then obtain c1' s' where "\<langle>c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s'\<rangle> \<or> \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1',s'\<rangle>" | |
| 18372 | 171 | by (cases y, cases "fst y") auto | 
| 12431 | 172 | thus "\<exists>s'. \<langle>c1;c2,s\<rangle> \<longrightarrow>\<^sub>1 s'" by auto | 
| 173 | next | |
| 174 | -- "case If:" | |
| 175 | fix b c1 c2 assume "\<exists>y. \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>1 y" and "\<exists>y. \<langle>c2,s\<rangle> \<longrightarrow>\<^sub>1 y" | |
| 176 | thus "\<exists>y. \<langle>\<IF> b \<THEN> c1 \<ELSE> c2, s\<rangle> \<longrightarrow>\<^sub>1 y" by (cases "b s") auto | |
| 177 | qed auto -- "the rest is trivial" | |
| 178 | ||
| 179 | text {*
 | |
| 180 | If a configuration does not contain a statement, the program | |
| 181 | has terminated and there is no next configuration: | |
| 182 | *} | |
| 18372 | 183 | lemma stuck [elim!]: "\<langle>s\<rangle> \<longrightarrow>\<^sub>1 y \<Longrightarrow> P" | 
| 23746 | 184 | by (induct y, auto elim: evalc1.cases) | 
| 12434 | 185 | |
| 18372 | 186 | lemma evalc_None_retrancl [simp, dest!]: "\<langle>s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* s' \<Longrightarrow> s' = \<langle>s\<rangle>" | 
| 23746 | 187 | by (induct set: rtrancl) auto | 
| 12431 | 188 | |
| 189 | (*<*) | |
| 18372 | 190 | (* FIXME: relpow.simps don't work *) | 
| 12431 | 191 | lemmas [simp del] = relpow.simps | 
| 25862 | 192 | lemma rel_pow_0 [simp]: "!!R::('a*'a) set. R^0 = Id" by (simp add: relpow.simps)
 | 
| 193 | lemma rel_pow_Suc_0 [simp]: "!!R::('a*'a) set. R^(Suc 0) = R" by (simp add: relpow.simps)
 | |
| 194 | ||
| 12431 | 195 | (*>*) | 
| 18557 
60a0f9caa0a2
Provers/classical: stricter checks to ensure that supplied intro, dest and
 paulson parents: 
18447diff
changeset | 196 | lemma evalc1_None_0 [simp]: "\<langle>s\<rangle> -n\<rightarrow>\<^sub>1 y = (n = 0 \<and> y = \<langle>s\<rangle>)" | 
| 12431 | 197 | by (cases n) auto | 
| 4906 | 198 | |
| 18372 | 199 | lemma SKIP_n: "\<langle>\<SKIP>, s\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s'\<rangle> \<Longrightarrow> s' = s \<and> n=1" | 
| 12431 | 200 | by (cases n) auto | 
| 201 | ||
| 202 | subsection "Equivalence to natural semantics (after Nielson and Nielson)" | |
| 203 | ||
| 204 | text {*
 | |
| 205 | We first need two lemmas about semicolon statements: | |
| 206 | decomposition and composition. | |
| 207 | *} | |
| 208 | lemma semiD: | |
| 18372 | 209 | "\<langle>c1; c2, s\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle> \<Longrightarrow> | 
| 12431 | 210 | \<exists>i j s'. \<langle>c1, s\<rangle> -i\<rightarrow>\<^sub>1 \<langle>s'\<rangle> \<and> \<langle>c2, s'\<rangle> -j\<rightarrow>\<^sub>1 \<langle>s''\<rangle> \<and> n = i+j" | 
| 20503 | 211 | proof (induct n arbitrary: c1 c2 s s'') | 
| 18372 | 212 | case 0 | 
| 213 | then show ?case by simp | |
| 12431 | 214 | next | 
| 18372 | 215 | case (Suc n) | 
| 216 | ||
| 217 | from `\<langle>c1; c2, s\<rangle> -Suc n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>` | |
| 23746 | 218 | obtain co s''' where | 
| 219 | 1: "\<langle>c1; c2, s\<rangle> \<longrightarrow>\<^sub>1 (co, s''')" and | |
| 220 | n: "(co, s''') -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>" | |
| 221 | by auto | |
| 12431 | 222 | |
| 18372 | 223 | from 1 | 
| 224 | show "\<exists>i j s'. \<langle>c1, s\<rangle> -i\<rightarrow>\<^sub>1 \<langle>s'\<rangle> \<and> \<langle>c2, s'\<rangle> -j\<rightarrow>\<^sub>1 \<langle>s''\<rangle> \<and> Suc n = i+j" | |
| 225 | (is "\<exists>i j s'. ?Q i j s'") | |
| 226 | proof (cases set: evalc1) | |
| 227 | case Semi1 | |
| 228 | then obtain s' where | |
| 23746 | 229 | "co = Some c2" and "s''' = s'" and "\<langle>c1, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s'\<rangle>" | 
| 18372 | 230 | by auto | 
| 231 | with 1 n have "?Q 1 n s'" by simp | |
| 232 | thus ?thesis by blast | |
| 233 | next | |
| 234 | case Semi2 | |
| 235 | then obtain c1' s' where | |
| 23746 | 236 | "co = Some (c1'; c2)" "s''' = s'" and | 
| 12431 | 237 | c1: "\<langle>c1, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1', s'\<rangle>" | 
| 18372 | 238 | by auto | 
| 239 | with n have "\<langle>c1'; c2, s'\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>" by simp | |
| 240 | with Suc.hyps obtain i j s0 where | |
| 12431 | 241 | c1': "\<langle>c1',s'\<rangle> -i\<rightarrow>\<^sub>1 \<langle>s0\<rangle>" and | 
| 242 | c2: "\<langle>c2,s0\<rangle> -j\<rightarrow>\<^sub>1 \<langle>s''\<rangle>" and | |
| 243 | i: "n = i+j" | |
| 18372 | 244 | by fast | 
| 245 | ||
| 246 | from c1 c1' | |
| 247 | have "\<langle>c1,s\<rangle> -(i+1)\<rightarrow>\<^sub>1 \<langle>s0\<rangle>" by (auto intro: rel_pow_Suc_I2) | |
| 248 | with c2 i | |
| 249 | have "?Q (i+1) j s0" by simp | |
| 250 | thus ?thesis by blast | |
| 251 | qed auto -- "the remaining cases cannot occur" | |
| 12431 | 252 | qed | 
| 1700 | 253 | |
| 254 | ||
| 18372 | 255 | lemma semiI: | 
| 256 | "\<langle>c0,s\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle> \<Longrightarrow> \<langle>c1,s''\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle> \<Longrightarrow> \<langle>c0; c1, s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" | |
| 20503 | 257 | proof (induct n arbitrary: c0 s s'') | 
| 18372 | 258 | case 0 | 
| 259 | from `\<langle>c0,s\<rangle> -(0::nat)\<rightarrow>\<^sub>1 \<langle>s''\<rangle>` | |
| 260 | have False by simp | |
| 261 | thus ?case .. | |
| 12431 | 262 | next | 
| 18372 | 263 | case (Suc n) | 
| 264 | note c0 = `\<langle>c0,s\<rangle> -Suc n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>` | |
| 265 | note c1 = `\<langle>c1,s''\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>` | |
| 266 | note IH = `\<And>c0 s s''. | |
| 267 | \<langle>c0,s\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle> \<Longrightarrow> \<langle>c1,s''\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle> \<Longrightarrow> \<langle>c0; c1,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>` | |
| 268 | from c0 obtain y where | |
| 12431 | 269 | 1: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>1 y" and n: "y -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>" by blast | 
| 270 | from 1 obtain c0' s0' where | |
| 18372 | 271 | "y = \<langle>s0'\<rangle> \<or> y = \<langle>c0', s0'\<rangle>" | 
| 272 | by (cases y, cases "fst y") auto | |
| 12431 | 273 | moreover | 
| 274 |   { assume y: "y = \<langle>s0'\<rangle>"
 | |
| 275 | with n have "s'' = s0'" by simp | |
| 276 | with y 1 have "\<langle>c0; c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1, s''\<rangle>" by blast | |
| 277 | with c1 have "\<langle>c0; c1,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" by (blast intro: rtrancl_trans) | |
| 278 | } | |
| 279 | moreover | |
| 280 |   { assume y: "y = \<langle>c0', s0'\<rangle>"
 | |
| 281 | with n have "\<langle>c0', s0'\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>" by blast | |
| 282 | with IH c1 have "\<langle>c0'; c1,s0'\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" by blast | |
| 283 | moreover | |
| 284 | from y 1 have "\<langle>c0; c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c0'; c1,s0'\<rangle>" by blast | |
| 285 | hence "\<langle>c0; c1,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>c0'; c1,s0'\<rangle>" by blast | |
| 286 | ultimately | |
| 287 | have "\<langle>c0; c1,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" by (blast intro: rtrancl_trans) | |
| 288 | } | |
| 289 | ultimately | |
| 290 | show "\<langle>c0; c1,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" by blast | |
| 291 | qed | |
| 292 | ||
| 293 | text {*
 | |
| 294 | The easy direction of the equivalence proof: | |
| 295 | *} | |
| 18372 | 296 | lemma evalc_imp_evalc1: | 
| 297 | assumes "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" | |
| 298 | shows "\<langle>c, s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" | |
| 299 | using prems | |
| 300 | proof induct | |
| 301 | fix s show "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s\<rangle>" by auto | |
| 302 | next | |
| 303 | fix x a s show "\<langle>x :== a ,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s[x\<mapsto>a s]\<rangle>" by auto | |
| 304 | next | |
| 305 | fix c0 c1 s s'' s' | |
| 306 | assume "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s''\<rangle>" | |
| 307 | then obtain n where "\<langle>c0,s\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>" by (blast dest: rtrancl_imp_rel_pow) | |
| 308 | moreover | |
| 309 | assume "\<langle>c1,s''\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" | |
| 310 | ultimately | |
| 311 | show "\<langle>c0; c1,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" by (rule semiI) | |
| 312 | next | |
| 313 | fix s::state and b c0 c1 s' | |
| 314 | assume "b s" | |
| 315 | hence "\<langle>\<IF> b \<THEN> c0 \<ELSE> c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c0,s\<rangle>" by simp | |
| 316 | also assume "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" | |
| 317 | finally show "\<langle>\<IF> b \<THEN> c0 \<ELSE> c1,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" . | |
| 318 | next | |
| 319 | fix s::state and b c0 c1 s' | |
| 320 | assume "\<not>b s" | |
| 321 | hence "\<langle>\<IF> b \<THEN> c0 \<ELSE> c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1,s\<rangle>" by simp | |
| 322 | also assume "\<langle>c1,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" | |
| 323 | finally show "\<langle>\<IF> b \<THEN> c0 \<ELSE> c1,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" . | |
| 324 | next | |
| 325 | fix b c and s::state | |
| 326 | assume b: "\<not>b s" | |
| 327 | let ?if = "\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>" | |
| 328 | have "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>?if, s\<rangle>" by blast | |
| 329 | also have "\<langle>?if,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>\<SKIP>, s\<rangle>" by (simp add: b) | |
| 330 | also have "\<langle>\<SKIP>, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s\<rangle>" by blast | |
| 331 | finally show "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s\<rangle>" .. | |
| 332 | next | |
| 333 | fix b c s s'' s' | |
| 334 | let ?w = "\<WHILE> b \<DO> c" | |
| 335 | let ?if = "\<IF> b \<THEN> c; ?w \<ELSE> \<SKIP>" | |
| 336 | assume w: "\<langle>?w,s''\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" | |
| 337 | assume c: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s''\<rangle>" | |
| 338 | assume b: "b s" | |
| 339 | have "\<langle>?w,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>?if, s\<rangle>" by blast | |
| 340 | also have "\<langle>?if, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c; ?w, s\<rangle>" by (simp add: b) | |
| 341 | also | |
| 342 | from c obtain n where "\<langle>c,s\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>" by (blast dest: rtrancl_imp_rel_pow) | |
| 343 | with w have "\<langle>c; ?w,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" by - (rule semiI) | |
| 344 | finally show "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" .. | |
| 12431 | 345 | qed | 
| 346 | ||
| 347 | text {*
 | |
| 348 | Finally, the equivalence theorem: | |
| 349 | *} | |
| 350 | theorem evalc_equiv_evalc1: | |
| 351 | "\<langle>c, s\<rangle> \<longrightarrow>\<^sub>c s' = \<langle>c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" | |
| 352 | proof | |
| 353 | assume "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" | |
| 23373 | 354 | then show "\<langle>c, s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" by (rule evalc_imp_evalc1) | 
| 18372 | 355 | next | 
| 12431 | 356 | assume "\<langle>c, s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s'\<rangle>" | 
| 357 | then obtain n where "\<langle>c, s\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s'\<rangle>" by (blast dest: rtrancl_imp_rel_pow) | |
| 358 | moreover | |
| 18372 | 359 | have "\<langle>c, s\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s'\<rangle> \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" | 
| 20503 | 360 | proof (induct arbitrary: c s s' rule: less_induct) | 
| 12431 | 361 | fix n | 
| 18372 | 362 | assume IH: "\<And>m c s s'. m < n \<Longrightarrow> \<langle>c,s\<rangle> -m\<rightarrow>\<^sub>1 \<langle>s'\<rangle> \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" | 
| 12431 | 363 | fix c s s' | 
| 364 | assume c: "\<langle>c, s\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s'\<rangle>" | |
| 365 | then obtain m where n: "n = Suc m" by (cases n) auto | |
| 18372 | 366 | with c obtain y where | 
| 12431 | 367 | c': "\<langle>c, s\<rangle> \<longrightarrow>\<^sub>1 y" and m: "y -m\<rightarrow>\<^sub>1 \<langle>s'\<rangle>" by blast | 
| 18372 | 368 | show "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" | 
| 12431 | 369 | proof (cases c) | 
| 370 | case SKIP | |
| 371 | with c n show ?thesis by auto | |
| 18372 | 372 | next | 
| 12431 | 373 | case Assign | 
| 374 | with c n show ?thesis by auto | |
| 375 | next | |
| 376 | fix c1 c2 assume semi: "c = (c1; c2)" | |
| 377 | with c obtain i j s'' where | |
| 18372 | 378 | c1: "\<langle>c1, s\<rangle> -i\<rightarrow>\<^sub>1 \<langle>s''\<rangle>" and | 
| 379 | c2: "\<langle>c2, s''\<rangle> -j\<rightarrow>\<^sub>1 \<langle>s'\<rangle>" and | |
| 380 | ij: "n = i+j" | |
| 12431 | 381 | by (blast dest: semiD) | 
| 18372 | 382 | from c1 c2 obtain | 
| 12431 | 383 | "0 < i" and "0 < j" by (cases i, auto, cases j, auto) | 
| 384 | with ij obtain | |
| 385 | i: "i < n" and j: "j < n" by simp | |
| 18372 | 386 | from IH i c1 | 
| 387 | have "\<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c s''" . | |
| 12431 | 388 | moreover | 
| 18372 | 389 | from IH j c2 | 
| 390 | have "\<langle>c2,s''\<rangle> \<longrightarrow>\<^sub>c s'" . | |
| 12431 | 391 | moreover | 
| 392 | note semi | |
| 393 | ultimately | |
| 394 | show "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" by blast | |
| 395 | next | |
| 396 | fix b c1 c2 assume If: "c = \<IF> b \<THEN> c1 \<ELSE> c2" | |
| 397 |       { assume True: "b s = True"
 | |
| 398 | with If c n | |
| 18372 | 399 | have "\<langle>c1,s\<rangle> -m\<rightarrow>\<^sub>1 \<langle>s'\<rangle>" by auto | 
| 12431 | 400 | with n IH | 
| 401 | have "\<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c s'" by blast | |
| 402 | with If True | |
| 18372 | 403 | have "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" by simp | 
| 12431 | 404 | } | 
| 405 | moreover | |
| 406 |       { assume False: "b s = False"
 | |
| 407 | with If c n | |
| 18372 | 408 | have "\<langle>c2,s\<rangle> -m\<rightarrow>\<^sub>1 \<langle>s'\<rangle>" by auto | 
| 12431 | 409 | with n IH | 
| 410 | have "\<langle>c2,s\<rangle> \<longrightarrow>\<^sub>c s'" by blast | |
| 411 | with If False | |
| 18372 | 412 | have "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" by simp | 
| 12431 | 413 | } | 
| 414 | ultimately | |
| 415 | show "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" by (cases "b s") auto | |
| 416 | next | |
| 417 | fix b c' assume w: "c = \<WHILE> b \<DO> c'" | |
| 18372 | 418 | with c n | 
| 12431 | 419 | have "\<langle>\<IF> b \<THEN> c'; \<WHILE> b \<DO> c' \<ELSE> \<SKIP>,s\<rangle> -m\<rightarrow>\<^sub>1 \<langle>s'\<rangle>" | 
| 420 | (is "\<langle>?if,_\<rangle> -m\<rightarrow>\<^sub>1 _") by auto | |
| 421 | with n IH | |
| 422 | have "\<langle>\<IF> b \<THEN> c'; \<WHILE> b \<DO> c' \<ELSE> \<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c s'" by blast | |
| 423 | moreover note unfold_while [of b c'] | |
| 424 |       -- {* @{thm unfold_while [of b c']} *}
 | |
| 18372 | 425 | ultimately | 
| 12431 | 426 | have "\<langle>\<WHILE> b \<DO> c',s\<rangle> \<longrightarrow>\<^sub>c s'" by (blast dest: equivD2) | 
| 427 | with w show "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" by simp | |
| 428 | qed | |
| 429 | qed | |
| 430 | ultimately | |
| 431 | show "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" by blast | |
| 432 | qed | |
| 433 | ||
| 434 | ||
| 435 | subsection "Winskel's Proof" | |
| 436 | ||
| 437 | declare rel_pow_0_E [elim!] | |
| 438 | ||
| 439 | text {*
 | |
| 18372 | 440 |   Winskel's small step rules are a bit different \cite{Winskel};
 | 
| 12431 | 441 | we introduce their equivalents as derived rules: | 
| 442 | *} | |
| 443 | lemma whileFalse1 [intro]: | |
| 18372 | 444 | "\<not> b s \<Longrightarrow> \<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s\<rangle>" (is "_ \<Longrightarrow> \<langle>?w, s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s\<rangle>") | 
| 12431 | 445 | proof - | 
| 446 | assume "\<not>b s" | |
| 447 | have "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>\<IF> b \<THEN> c;?w \<ELSE> \<SKIP>, s\<rangle>" .. | |
| 23373 | 448 | also from `\<not>b s` have "\<langle>\<IF> b \<THEN> c;?w \<ELSE> \<SKIP>, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>\<SKIP>, s\<rangle>" .. | 
| 12431 | 449 | also have "\<langle>\<SKIP>, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s\<rangle>" .. | 
| 450 | finally show "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s\<rangle>" .. | |
| 451 | qed | |
| 452 | ||
| 453 | lemma whileTrue1 [intro]: | |
| 18372 | 454 | "b s \<Longrightarrow> \<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>c;\<WHILE> b \<DO> c, s\<rangle>" | 
| 12431 | 455 | (is "_ \<Longrightarrow> \<langle>?w, s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>c;?w,s\<rangle>") | 
| 456 | proof - | |
| 457 | assume "b s" | |
| 458 | have "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>\<IF> b \<THEN> c;?w \<ELSE> \<SKIP>, s\<rangle>" .. | |
| 23373 | 459 | also from `b s` have "\<langle>\<IF> b \<THEN> c;?w \<ELSE> \<SKIP>, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c;?w, s\<rangle>" .. | 
| 12431 | 460 | finally show "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>c;?w,s\<rangle>" .. | 
| 461 | qed | |
| 1700 | 462 | |
| 18372 | 463 | inductive_cases evalc1_SEs: | 
| 23746 | 464 | "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>1 (co, s')" | 
| 465 | "\<langle>x:==a,s\<rangle> \<longrightarrow>\<^sub>1 (co, s')" | |
| 466 | "\<langle>c1;c2, s\<rangle> \<longrightarrow>\<^sub>1 (co, s')" | |
| 467 | "\<langle>\<IF> b \<THEN> c1 \<ELSE> c2, s\<rangle> \<longrightarrow>\<^sub>1 (co, s')" | |
| 468 | "\<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>1 (co, s')" | |
| 12431 | 469 | |
| 23746 | 470 | inductive_cases evalc1_E: "\<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>1 (co, s')" | 
| 12431 | 471 | |
| 472 | declare evalc1_SEs [elim!] | |
| 473 | ||
| 474 | lemma evalc_impl_evalc1: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s1 \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s1\<rangle>" | |
| 18372 | 475 | apply (induct set: evalc) | 
| 12431 | 476 | |
| 18372 | 477 | -- SKIP | 
| 12431 | 478 | apply blast | 
| 479 | ||
| 18372 | 480 | -- ASSIGN | 
| 12431 | 481 | apply fast | 
| 482 | ||
| 18372 | 483 | -- SEMI | 
| 12431 | 484 | apply (fast dest: rtrancl_imp_UN_rel_pow intro: semiI) | 
| 485 | ||
| 18372 | 486 | -- IF | 
| 12566 
fe20540bcf93
renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
 nipkow parents: 
12546diff
changeset | 487 | apply (fast intro: converse_rtrancl_into_rtrancl) | 
| 
fe20540bcf93
renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
 nipkow parents: 
12546diff
changeset | 488 | apply (fast intro: converse_rtrancl_into_rtrancl) | 
| 12431 | 489 | |
| 18372 | 490 | -- WHILE | 
| 12431 | 491 | apply fast | 
| 12566 
fe20540bcf93
renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
 nipkow parents: 
12546diff
changeset | 492 | apply (fast dest: rtrancl_imp_UN_rel_pow intro: converse_rtrancl_into_rtrancl semiI) | 
| 12431 | 493 | |
| 494 | done | |
| 495 | ||
| 496 | ||
| 18372 | 497 | lemma lemma2: | 
| 498 | "\<langle>c;d,s\<rangle> -n\<rightarrow>\<^sub>1 \<langle>u\<rangle> \<Longrightarrow> \<exists>t m. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>t\<rangle> \<and> \<langle>d,t\<rangle> -m\<rightarrow>\<^sub>1 \<langle>u\<rangle> \<and> m \<le> n" | |
| 20503 | 499 | apply (induct n arbitrary: c d s u) | 
| 12431 | 500 | -- "case n = 0" | 
| 501 | apply fastsimp | |
| 502 | -- "induction step" | |
| 18372 | 503 | apply (fast intro!: le_SucI le_refl dest!: rel_pow_Suc_D2 | 
| 12566 
fe20540bcf93
renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
 nipkow parents: 
12546diff
changeset | 504 | elim!: rel_pow_imp_rtrancl converse_rtrancl_into_rtrancl) | 
| 12431 | 505 | done | 
| 506 | ||
| 18372 | 507 | lemma evalc1_impl_evalc: | 
| 508 | "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>t\<rangle> \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t" | |
| 20503 | 509 | apply (induct c arbitrary: s t) | 
| 12431 | 510 | apply (safe dest!: rtrancl_imp_UN_rel_pow) | 
| 511 | ||
| 512 | -- SKIP | |
| 513 | apply (simp add: SKIP_n) | |
| 514 | ||
| 18372 | 515 | -- ASSIGN | 
| 12431 | 516 | apply (fastsimp elim: rel_pow_E2) | 
| 517 | ||
| 518 | -- SEMI | |
| 519 | apply (fast dest!: rel_pow_imp_rtrancl lemma2) | |
| 520 | ||
| 18372 | 521 | -- IF | 
| 12431 | 522 | apply (erule rel_pow_E2) | 
| 523 | apply simp | |
| 524 | apply (fast dest!: rel_pow_imp_rtrancl) | |
| 525 | ||
| 526 | -- "WHILE, induction on the length of the computation" | |
| 527 | apply (rename_tac b c s t n) | |
| 528 | apply (erule_tac P = "?X -n\<rightarrow>\<^sub>1 ?Y" in rev_mp) | |
| 529 | apply (rule_tac x = "s" in spec) | |
| 18372 | 530 | apply (induct_tac n rule: nat_less_induct) | 
| 12431 | 531 | apply (intro strip) | 
| 532 | apply (erule rel_pow_E2) | |
| 533 | apply simp | |
| 23746 | 534 | apply (simp only: split_paired_all) | 
| 12431 | 535 | apply (erule evalc1_E) | 
| 536 | ||
| 537 | apply simp | |
| 538 | apply (case_tac "b x") | |
| 539 | -- WhileTrue | |
| 540 | apply (erule rel_pow_E2) | |
| 541 | apply simp | |
| 542 | apply (clarify dest!: lemma2) | |
| 18372 | 543 | apply atomize | 
| 12431 | 544 | apply (erule allE, erule allE, erule impE, assumption) | 
| 545 | apply (erule_tac x=mb in allE, erule impE, fastsimp) | |
| 546 | apply blast | |
| 18372 | 547 | -- WhileFalse | 
| 12431 | 548 | apply (erule rel_pow_E2) | 
| 549 | apply simp | |
| 550 | apply (simp add: SKIP_n) | |
| 551 | done | |
| 552 | ||
| 553 | ||
| 554 | text {* proof of the equivalence of evalc and evalc1 *}
 | |
| 555 | lemma evalc1_eq_evalc: "(\<langle>c, s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>t\<rangle>) = (\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t)" | |
| 18372 | 556 | by (fast elim!: evalc1_impl_evalc evalc_impl_evalc1) | 
| 12431 | 557 | |
| 558 | ||
| 559 | subsection "A proof without n" | |
| 560 | ||
| 561 | text {*
 | |
| 562 | The inductions are a bit awkward to write in this section, | |
| 563 |   because @{text None} as result statement in the small step
 | |
| 564 | semantics doesn't have a direct counterpart in the big step | |
| 18372 | 565 | semantics. | 
| 1700 | 566 | |
| 12431 | 567 |   Winskel's small step rule set (using the @{text "\<SKIP>"} statement
 | 
| 568 | to indicate termination) is better suited for this proof. | |
| 569 | *} | |
| 570 | ||
| 18372 | 571 | lemma my_lemma1: | 
| 572 | assumes "\<langle>c1,s1\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s2\<rangle>" | |
| 573 | and "\<langle>c2,s2\<rangle> \<longrightarrow>\<^sub>1\<^sup>* cs3" | |
| 574 | shows "\<langle>c1;c2,s1\<rangle> \<longrightarrow>\<^sub>1\<^sup>* cs3" | |
| 12431 | 575 | proof - | 
| 576 |   -- {* The induction rule needs @{text P} to be a function of @{term "Some c1"} *}
 | |
| 18372 | 577 | from prems | 
| 578 | have "\<langle>(\<lambda>c. if c = None then c2 else the c; c2) (Some c1),s1\<rangle> \<longrightarrow>\<^sub>1\<^sup>* cs3" | |
| 579 | apply (induct rule: converse_rtrancl_induct2) | |
| 12431 | 580 | apply simp | 
| 581 | apply (rename_tac c s') | |
| 582 | apply simp | |
| 583 | apply (rule conjI) | |
| 18372 | 584 | apply fast | 
| 12431 | 585 | apply clarify | 
| 586 | apply (case_tac c) | |
| 12566 
fe20540bcf93
renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
 nipkow parents: 
12546diff
changeset | 587 | apply (auto intro: converse_rtrancl_into_rtrancl) | 
| 12431 | 588 | done | 
| 18372 | 589 | then show ?thesis by simp | 
| 12431 | 590 | qed | 
| 591 | ||
| 13524 | 592 | lemma evalc_impl_evalc1': "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s1 \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>s1\<rangle>" | 
| 18372 | 593 | apply (induct set: evalc) | 
| 12431 | 594 | |
| 18372 | 595 | -- SKIP | 
| 12431 | 596 | apply fast | 
| 597 | ||
| 598 | -- ASSIGN | |
| 599 | apply fast | |
| 600 | ||
| 18372 | 601 | -- SEMI | 
| 12431 | 602 | apply (fast intro: my_lemma1) | 
| 603 | ||
| 604 | -- IF | |
| 12566 
fe20540bcf93
renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
 nipkow parents: 
12546diff
changeset | 605 | apply (fast intro: converse_rtrancl_into_rtrancl) | 
| 18372 | 606 | apply (fast intro: converse_rtrancl_into_rtrancl) | 
| 12431 | 607 | |
| 18372 | 608 | -- WHILE | 
| 12431 | 609 | apply fast | 
| 12566 
fe20540bcf93
renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
 nipkow parents: 
12546diff
changeset | 610 | apply (fast intro: converse_rtrancl_into_rtrancl my_lemma1) | 
| 12431 | 611 | |
| 612 | done | |
| 613 | ||
| 614 | text {*
 | |
| 615 | The opposite direction is based on a Coq proof done by Ranan Fraer and | |
| 616 | Yves Bertot. The following sketch is from an email by Ranan Fraer. | |
| 617 | ||
| 618 | \begin{verbatim}
 | |
| 619 | First we've broke it into 2 lemmas: | |
| 1700 | 620 | |
| 12431 | 621 | Lemma 1 | 
| 622 | ((c,s) --> (SKIP,t)) => (<c,s> -c-> t) | |
| 623 | ||
| 624 | This is a quick one, dealing with the cases skip, assignment | |
| 625 | and while_false. | |
| 626 | ||
| 627 | Lemma 2 | |
| 628 | ((c,s) -*-> (c',s')) /\ <c',s'> -c'-> t | |
| 18372 | 629 | => | 
| 12431 | 630 | <c,s> -c-> t | 
| 631 | ||
| 632 | This is proved by rule induction on the -*-> relation | |
| 18372 | 633 | and the induction step makes use of a third lemma: | 
| 12431 | 634 | |
| 635 | Lemma 3 | |
| 636 | ((c,s) --> (c',s')) /\ <c',s'> -c'-> t | |
| 18372 | 637 | => | 
| 12431 | 638 | <c,s> -c-> t | 
| 639 | ||
| 18372 | 640 | This captures the essence of the proof, as it shows that <c',s'> | 
| 12431 | 641 | behaves as the continuation of <c,s> with respect to the natural | 
| 642 | semantics. | |
| 643 | The proof of Lemma 3 goes by rule induction on the --> relation, | |
| 644 | dealing with the cases sequence1, sequence2, if_true, if_false and | |
| 645 | while_true. In particular in the case (sequence1) we make use again | |
| 646 | of Lemma 1. | |
| 647 | \end{verbatim}
 | |
| 648 | *} | |
| 649 | ||
| 650 | inductive_cases evalc1_term_cases: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s'\<rangle>" | |
| 651 | ||
| 18372 | 652 | lemma FB_lemma3: | 
| 653 | "(c,s) \<longrightarrow>\<^sub>1 (c',s') \<Longrightarrow> c \<noteq> None \<Longrightarrow> | |
| 654 | \<langle>if c'=None then \<SKIP> else the c',s'\<rangle> \<longrightarrow>\<^sub>c t \<Longrightarrow> \<langle>the c,s\<rangle> \<longrightarrow>\<^sub>c t" | |
| 20503 | 655 | by (induct arbitrary: t set: evalc1) | 
| 18372 | 656 | (auto elim!: evalc1_term_cases equivD2 [OF unfold_while]) | 
| 12431 | 657 | |
| 18372 | 658 | lemma FB_lemma2: | 
| 659 | "(c,s) \<longrightarrow>\<^sub>1\<^sup>* (c',s') \<Longrightarrow> c \<noteq> None \<Longrightarrow> | |
| 660 | \<langle>if c' = None then \<SKIP> else the c',s'\<rangle> \<longrightarrow>\<^sub>c t \<Longrightarrow> \<langle>the c,s\<rangle> \<longrightarrow>\<^sub>c t" | |
| 18447 | 661 | apply (induct rule: converse_rtrancl_induct2, force) | 
| 12434 | 662 | apply (fastsimp elim!: evalc1_term_cases intro: FB_lemma3) | 
| 12431 | 663 | done | 
| 664 | ||
| 13524 | 665 | lemma evalc1_impl_evalc': "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* \<langle>t\<rangle> \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t" | 
| 18372 | 666 | by (fastsimp dest: FB_lemma2) | 
| 1700 | 667 | |
| 668 | end |