| author | blanchet | 
| Tue, 10 Sep 2013 15:56:51 +0200 | |
| changeset 53503 | d2f21e305d0c | 
| parent 53015 | a1119cf551e8 | 
| child 54192 | a5eec4263b3a | 
| permissions | -rw-r--r-- | 
| 43141 | 1 | header "Small-Step Semantics of Commands" | 
| 2 | ||
| 3 | theory Small_Step imports Star Big_Step begin | |
| 4 | ||
| 5 | subsection "The transition relation" | |
| 6 | ||
| 49191 | 7 | text_raw{*\snip{SmallStepDef}{0}{2}{% *}
 | 
| 43141 | 8 | inductive | 
| 9 | small_step :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>" 55) | |
| 10 | where | |
| 11 | Assign: "(x ::= a, s) \<rightarrow> (SKIP, s(x := aval a s))" | | |
| 12 | ||
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 13 | Seq1: "(SKIP;;c\<^sub>2,s) \<rightarrow> (c\<^sub>2,s)" | | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 14 | Seq2: "(c\<^sub>1,s) \<rightarrow> (c\<^sub>1',s') \<Longrightarrow> (c\<^sub>1;;c\<^sub>2,s) \<rightarrow> (c\<^sub>1';;c\<^sub>2,s')" | | 
| 43141 | 15 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 16 | IfTrue: "bval b s \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2,s) \<rightarrow> (c\<^sub>1,s)" | | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52046diff
changeset | 17 | IfFalse: "\<not>bval b s \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2,s) \<rightarrow> (c\<^sub>2,s)" | | 
| 43141 | 18 | |
| 50054 | 19 | While: "(WHILE b DO c,s) \<rightarrow> | 
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
50054diff
changeset | 20 | (IF b THEN c;; WHILE b DO c ELSE SKIP,s)" | 
| 49191 | 21 | text_raw{*}%endsnip*}
 | 
| 43141 | 22 | |
| 23 | ||
| 50054 | 24 | abbreviation | 
| 25 | small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>*" 55) | |
| 43141 | 26 | where "x \<rightarrow>* y == star small_step x y" | 
| 27 | ||
| 28 | subsection{* Executability *}
 | |
| 29 | ||
| 30 | code_pred small_step . | |
| 31 | ||
| 32 | values "{(c',map t [''x'',''y'',''z'']) |c' t.
 | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
50054diff
changeset | 33 |    (''x'' ::= V ''z'';; ''y'' ::= V ''x'',
 | 
| 44036 | 34 | <''x'' := 3, ''y'' := 7, ''z'' := 5>) \<rightarrow>* (c',t)}" | 
| 43141 | 35 | |
| 36 | ||
| 37 | subsection{* Proof infrastructure *}
 | |
| 38 | ||
| 39 | subsubsection{* Induction rules *}
 | |
| 40 | ||
| 41 | text{* The default induction rule @{thm[source] small_step.induct} only works
 | |
| 42 | for lemmas of the form @{text"a \<rightarrow> b \<Longrightarrow> \<dots>"} where @{text a} and @{text b} are
 | |
| 43 | not already pairs @{text"(DUMMY,DUMMY)"}. We can generate a suitable variant
 | |
| 44 | of @{thm[source] small_step.induct} for pairs by ``splitting'' the arguments
 | |
| 45 | @{text"\<rightarrow>"} into pairs: *}
 | |
| 46 | lemmas small_step_induct = small_step.induct[split_format(complete)] | |
| 47 | ||
| 48 | ||
| 49 | subsubsection{* Proof automation *}
 | |
| 50 | ||
| 51 | declare small_step.intros[simp,intro] | |
| 52 | ||
| 53 | text{* Rule inversion: *}
 | |
| 54 | ||
| 55 | inductive_cases SkipE[elim!]: "(SKIP,s) \<rightarrow> ct" | |
| 56 | thm SkipE | |
| 57 | inductive_cases AssignE[elim!]: "(x::=a,s) \<rightarrow> ct" | |
| 58 | thm AssignE | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
50054diff
changeset | 59 | inductive_cases SeqE[elim]: "(c1;;c2,s) \<rightarrow> ct" | 
| 47818 | 60 | thm SeqE | 
| 43141 | 61 | inductive_cases IfE[elim!]: "(IF b THEN c1 ELSE c2,s) \<rightarrow> ct" | 
| 62 | inductive_cases WhileE[elim]: "(WHILE b DO c, s) \<rightarrow> ct" | |
| 63 | ||
| 64 | ||
| 65 | text{* A simple property: *}
 | |
| 66 | lemma deterministic: | |
| 67 | "cs \<rightarrow> cs' \<Longrightarrow> cs \<rightarrow> cs'' \<Longrightarrow> cs'' = cs'" | |
| 45015 | 68 | apply(induction arbitrary: cs'' rule: small_step.induct) | 
| 43141 | 69 | apply blast+ | 
| 70 | done | |
| 71 | ||
| 72 | ||
| 73 | subsection "Equivalence with big-step semantics" | |
| 74 | ||
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
50054diff
changeset | 75 | lemma star_seq2: "(c1,s) \<rightarrow>* (c1',s') \<Longrightarrow> (c1;;c2,s) \<rightarrow>* (c1';;c2,s')" | 
| 45015 | 76 | proof(induction rule: star_induct) | 
| 43141 | 77 | case refl thus ?case by simp | 
| 78 | next | |
| 79 | case step | |
| 47818 | 80 | thus ?case by (metis Seq2 star.step) | 
| 43141 | 81 | qed | 
| 82 | ||
| 47818 | 83 | lemma seq_comp: | 
| 43141 | 84 | "\<lbrakk> (c1,s1) \<rightarrow>* (SKIP,s2); (c2,s2) \<rightarrow>* (SKIP,s3) \<rbrakk> | 
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
50054diff
changeset | 85 | \<Longrightarrow> (c1;;c2, s1) \<rightarrow>* (SKIP,s3)" | 
| 47818 | 86 | by(blast intro: star.step star_seq2 star_trans) | 
| 43141 | 87 | |
| 88 | text{* The following proof corresponds to one on the board where one would
 | |
| 45218 | 89 | show chains of @{text "\<rightarrow>"} and @{text "\<rightarrow>*"} steps. *}
 | 
| 43141 | 90 | |
| 91 | lemma big_to_small: | |
| 92 | "cs \<Rightarrow> t \<Longrightarrow> cs \<rightarrow>* (SKIP,t)" | |
| 45015 | 93 | proof (induction rule: big_step.induct) | 
| 43141 | 94 | fix s show "(SKIP,s) \<rightarrow>* (SKIP,s)" by simp | 
| 95 | next | |
| 96 | fix x a s show "(x ::= a,s) \<rightarrow>* (SKIP, s(x := aval a s))" by auto | |
| 97 | next | |
| 98 | fix c1 c2 s1 s2 s3 | |
| 99 | assume "(c1,s1) \<rightarrow>* (SKIP,s2)" and "(c2,s2) \<rightarrow>* (SKIP,s3)" | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
50054diff
changeset | 100 | thus "(c1;;c2, s1) \<rightarrow>* (SKIP,s3)" by (rule seq_comp) | 
| 43141 | 101 | next | 
| 102 | fix s::state and b c0 c1 t | |
| 103 | assume "bval b s" | |
| 104 | hence "(IF b THEN c0 ELSE c1,s) \<rightarrow> (c0,s)" by simp | |
| 45218 | 105 | moreover assume "(c0,s) \<rightarrow>* (SKIP,t)" | 
| 106 | ultimately | |
| 107 | show "(IF b THEN c0 ELSE c1,s) \<rightarrow>* (SKIP,t)" by (metis star.simps) | |
| 43141 | 108 | next | 
| 109 | fix s::state and b c0 c1 t | |
| 110 | assume "\<not>bval b s" | |
| 111 | hence "(IF b THEN c0 ELSE c1,s) \<rightarrow> (c1,s)" by simp | |
| 45218 | 112 | moreover assume "(c1,s) \<rightarrow>* (SKIP,t)" | 
| 113 | ultimately | |
| 114 | show "(IF b THEN c0 ELSE c1,s) \<rightarrow>* (SKIP,t)" by (metis star.simps) | |
| 43141 | 115 | next | 
| 116 | fix b c and s::state | |
| 117 | assume b: "\<not>bval b s" | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
50054diff
changeset | 118 | let ?if = "IF b THEN c;; WHILE b DO c ELSE SKIP" | 
| 43141 | 119 | have "(WHILE b DO c,s) \<rightarrow> (?if, s)" by blast | 
| 45218 | 120 | moreover have "(?if,s) \<rightarrow> (SKIP, s)" by (simp add: b) | 
| 45265 | 121 | ultimately show "(WHILE b DO c,s) \<rightarrow>* (SKIP,s)" by(metis star.refl star.step) | 
| 43141 | 122 | next | 
| 123 | fix b c s s' t | |
| 124 | let ?w = "WHILE b DO c" | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
50054diff
changeset | 125 | let ?if = "IF b THEN c;; ?w ELSE SKIP" | 
| 43141 | 126 | assume w: "(?w,s') \<rightarrow>* (SKIP,t)" | 
| 127 | assume c: "(c,s) \<rightarrow>* (SKIP,s')" | |
| 128 | assume b: "bval b s" | |
| 129 | have "(?w,s) \<rightarrow> (?if, s)" by blast | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
50054diff
changeset | 130 | moreover have "(?if, s) \<rightarrow> (c;; ?w, s)" by (simp add: b) | 
| 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
50054diff
changeset | 131 | moreover have "(c;; ?w,s) \<rightarrow>* (SKIP,t)" by(rule seq_comp[OF c w]) | 
| 45218 | 132 | ultimately show "(WHILE b DO c,s) \<rightarrow>* (SKIP,t)" by (metis star.simps) | 
| 43141 | 133 | qed | 
| 134 | ||
| 135 | text{* Each case of the induction can be proved automatically: *}
 | |
| 136 | lemma "cs \<Rightarrow> t \<Longrightarrow> cs \<rightarrow>* (SKIP,t)" | |
| 45015 | 137 | proof (induction rule: big_step.induct) | 
| 43141 | 138 | case Skip show ?case by blast | 
| 139 | next | |
| 140 | case Assign show ?case by blast | |
| 141 | next | |
| 47818 | 142 | case Seq thus ?case by (blast intro: seq_comp) | 
| 43141 | 143 | next | 
| 45265 | 144 | case IfTrue thus ?case by (blast intro: star.step) | 
| 43141 | 145 | next | 
| 45265 | 146 | case IfFalse thus ?case by (blast intro: star.step) | 
| 43141 | 147 | next | 
| 148 | case WhileFalse thus ?case | |
| 45265 | 149 | by (metis star.step star_step1 small_step.IfFalse small_step.While) | 
| 43141 | 150 | next | 
| 151 | case WhileTrue | |
| 152 | thus ?case | |
| 47818 | 153 | by(metis While seq_comp small_step.IfTrue star.step[of small_step]) | 
| 43141 | 154 | qed | 
| 155 | ||
| 156 | lemma small1_big_continue: | |
| 157 | "cs \<rightarrow> cs' \<Longrightarrow> cs' \<Rightarrow> t \<Longrightarrow> cs \<Rightarrow> t" | |
| 45015 | 158 | apply (induction arbitrary: t rule: small_step.induct) | 
| 43141 | 159 | apply auto | 
| 160 | done | |
| 161 | ||
| 162 | lemma small_big_continue: | |
| 163 | "cs \<rightarrow>* cs' \<Longrightarrow> cs' \<Rightarrow> t \<Longrightarrow> cs \<Rightarrow> t" | |
| 45015 | 164 | apply (induction rule: star.induct) | 
| 43141 | 165 | apply (auto intro: small1_big_continue) | 
| 166 | done | |
| 167 | ||
| 168 | lemma small_to_big: "cs \<rightarrow>* (SKIP,t) \<Longrightarrow> cs \<Rightarrow> t" | |
| 169 | by (metis small_big_continue Skip) | |
| 170 | ||
| 171 | text {*
 | |
| 172 | Finally, the equivalence theorem: | |
| 173 | *} | |
| 174 | theorem big_iff_small: | |
| 175 | "cs \<Rightarrow> t = cs \<rightarrow>* (SKIP,t)" | |
| 176 | by(metis big_to_small small_to_big) | |
| 177 | ||
| 178 | ||
| 179 | subsection "Final configurations and infinite reductions" | |
| 180 | ||
| 181 | definition "final cs \<longleftrightarrow> \<not>(EX cs'. cs \<rightarrow> cs')" | |
| 182 | ||
| 183 | lemma finalD: "final (c,s) \<Longrightarrow> c = SKIP" | |
| 184 | apply(simp add: final_def) | |
| 45015 | 185 | apply(induction c) | 
| 43141 | 186 | apply blast+ | 
| 187 | done | |
| 188 | ||
| 189 | lemma final_iff_SKIP: "final (c,s) = (c = SKIP)" | |
| 190 | by (metis SkipE finalD final_def) | |
| 191 | ||
| 192 | text{* Now we can show that @{text"\<Rightarrow>"} yields a final state iff @{text"\<rightarrow>"}
 | |
| 193 | terminates: *} | |
| 194 | ||
| 195 | lemma big_iff_small_termination: | |
| 196 | "(EX t. cs \<Rightarrow> t) \<longleftrightarrow> (EX cs'. cs \<rightarrow>* cs' \<and> final cs')" | |
| 197 | by(simp add: big_iff_small final_iff_SKIP) | |
| 198 | ||
| 199 | text{* This is the same as saying that the absence of a big step result is
 | |
| 200 | equivalent with absence of a terminating small step sequence, i.e.\ with | |
| 201 | nontermination.  Since @{text"\<rightarrow>"} is determininistic, there is no difference
 | |
| 202 | between may and must terminate. *} | |
| 203 | ||
| 204 | end |