| 12431 |      1 | (*  Title:        HOL/IMP/Natural.thy
 | 
|  |      2 |     ID:           $Id$
 | 
|  |      3 |     Author:       Tobias Nipkow & Robert Sandner, TUM
 | 
| 34055 |      4 |     Isar Version: Gerwin Klein, 2001; additional proofs by Lawrence Paulson
 | 
| 12431 |      5 |     Copyright     1996 TUM
 | 
| 1700 |      6 | *)
 | 
|  |      7 | 
 | 
| 12431 |      8 | header "Natural Semantics of Commands"
 | 
|  |      9 | 
 | 
| 16417 |     10 | theory Natural imports Com begin
 | 
| 12431 |     11 | 
 | 
|  |     12 | subsection "Execution of commands"
 | 
| 1700 |     13 | 
 | 
| 12431 |     14 | text {*
 | 
| 18372 |     15 |   We write @{text "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'"} for \emph{Statement @{text c}, started
 | 
| 12431 |     16 |   in state @{text s}, terminates in state @{text s'}}. Formally,
 | 
|  |     17 |   @{text "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'"} is just another form of saying \emph{the tuple
 | 
|  |     18 |   @{text "(c,s,s')"} is part of the relation @{text evalc}}:
 | 
|  |     19 | *}
 | 
| 1700 |     20 | 
 | 
| 27362 |     21 | definition
 | 
|  |     22 |   update :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)" ("_/[_ ::= /_]" [900,0,0] 900) where
 | 
|  |     23 |   "update = fun_upd"
 | 
| 12431 |     24 | 
 | 
| 27362 |     25 | notation (xsymbols)
 | 
|  |     26 |   update  ("_/[_ \<mapsto> /_]" [900,0,0] 900)
 | 
| 12431 |     27 | 
 | 
|  |     28 | text {*
 | 
|  |     29 |   The big-step execution relation @{text evalc} is defined inductively:
 | 
|  |     30 | *}
 | 
| 23746 |     31 | inductive
 | 
|  |     32 |   evalc :: "[com,state,state] \<Rightarrow> bool" ("\<langle>_,_\<rangle>/ \<longrightarrow>\<^sub>c _" [0,0,60] 60)
 | 
|  |     33 | where
 | 
| 12431 |     34 |   Skip:    "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c s"
 | 
| 23746 |     35 | | Assign:  "\<langle>x :== a,s\<rangle> \<longrightarrow>\<^sub>c s[x\<mapsto>a s]"
 | 
| 12431 |     36 | 
 | 
| 23746 |     37 | | Semi:    "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s'' \<Longrightarrow> \<langle>c1,s''\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>c0; c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
| 12431 |     38 | 
 | 
| 23746 |     39 | | IfTrue:  "b s \<Longrightarrow> \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
|  |     40 | | IfFalse: "\<not>b s \<Longrightarrow> \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
| 12431 |     41 | 
 | 
| 23746 |     42 | | WhileFalse: "\<not>b s \<Longrightarrow> \<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c s"
 | 
|  |     43 | | WhileTrue:  "b s \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'' \<Longrightarrow> \<langle>\<WHILE> b \<DO> c, s''\<rangle> \<longrightarrow>\<^sub>c s'
 | 
| 12431 |     44 |                \<Longrightarrow> \<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
|  |     45 | 
 | 
|  |     46 | lemmas evalc.intros [intro] -- "use those rules in automatic proofs"
 | 
|  |     47 | 
 | 
|  |     48 | text {*
 | 
|  |     49 | The induction principle induced by this definition looks like this:
 | 
|  |     50 | 
 | 
|  |     51 | @{thm [display] evalc.induct [no_vars]}
 | 
|  |     52 | 
 | 
|  |     53 | 
 | 
| 18372 |     54 | (@{text "\<And>"} and @{text "\<Longrightarrow>"} are Isabelle's
 | 
| 12431 |     55 |   meta symbols for @{text "\<forall>"} and @{text "\<longrightarrow>"})
 | 
|  |     56 | *}
 | 
|  |     57 | 
 | 
|  |     58 | text {*
 | 
|  |     59 |   The rules of @{text evalc} are syntax directed, i.e.~for each
 | 
|  |     60 |   syntactic category there is always only one rule applicable. That
 | 
| 34055 |     61 |   means we can use the rules in both directions.  This property is called rule inversion.
 | 
| 12431 |     62 | *}
 | 
| 34055 |     63 | inductive_cases skipE [elim!]:   "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
|  |     64 | inductive_cases semiE [elim!]:   "\<langle>c0; c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
|  |     65 | inductive_cases assignE [elim!]: "\<langle>x :== a,s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
|  |     66 | inductive_cases ifE [elim!]:     "\<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
|  |     67 | inductive_cases whileE [elim]:  "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
|  |     68 | 
 | 
|  |     69 | text {* The next proofs are all trivial by rule inversion.
 | 
|  |     70 | *}
 | 
|  |     71 | 
 | 
| 18372 |     72 | lemma skip:
 | 
| 12431 |     73 |   "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c s' = (s' = s)"
 | 
| 34055 |     74 |   by auto
 | 
| 12431 |     75 | 
 | 
| 18372 |     76 | lemma assign:
 | 
|  |     77 |   "\<langle>x :== a,s\<rangle> \<longrightarrow>\<^sub>c s' = (s' = s[x\<mapsto>a s])"
 | 
| 34055 |     78 |   by auto
 | 
| 12431 |     79 | 
 | 
| 18372 |     80 | lemma semi:
 | 
| 12431 |     81 |   "\<langle>c0; c1, s\<rangle> \<longrightarrow>\<^sub>c s' = (\<exists>s''. \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s'' \<and> \<langle>c1,s''\<rangle> \<longrightarrow>\<^sub>c s')"
 | 
| 34055 |     82 |   by auto
 | 
| 1700 |     83 | 
 | 
| 18372 |     84 | lemma ifTrue:
 | 
|  |     85 |   "b s \<Longrightarrow> \<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s' = \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
| 34055 |     86 |   by auto
 | 
| 12431 |     87 | 
 | 
| 18372 |     88 | lemma ifFalse:
 | 
| 12431 |     89 |   "\<not>b s \<Longrightarrow> \<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s' = \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
| 34055 |     90 |   by auto
 | 
| 12431 |     91 | 
 | 
| 18372 |     92 | lemma whileFalse:
 | 
| 12431 |     93 |   "\<not> b s \<Longrightarrow> \<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c s' = (s' = s)"
 | 
| 34055 |     94 |   by auto
 | 
| 12431 |     95 | 
 | 
| 18372 |     96 | lemma whileTrue:
 | 
|  |     97 |   "b s \<Longrightarrow>
 | 
|  |     98 |   \<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>c s' =
 | 
| 12431 |     99 |   (\<exists>s''. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'' \<and> \<langle>\<WHILE> b \<DO> c, s''\<rangle> \<longrightarrow>\<^sub>c s')"
 | 
| 34055 |    100 |   by auto
 | 
| 12431 |    101 | 
 | 
|  |    102 | text "Again, Isabelle may use these rules in automatic proofs:"
 | 
|  |    103 | lemmas evalc_cases [simp] = skip assign ifTrue ifFalse whileFalse semi whileTrue
 | 
|  |    104 | 
 | 
|  |    105 | subsection "Equivalence of statements"
 | 
| 1700 |    106 | 
 | 
| 12431 |    107 | text {*
 | 
|  |    108 |   We call two statements @{text c} and @{text c'} equivalent wrt.~the
 | 
|  |    109 |   big-step semantics when \emph{@{text c} started in @{text s} terminates
 | 
|  |    110 |   in @{text s'} iff @{text c'} started in the same @{text s} also terminates
 | 
|  |    111 |   in the same @{text s'}}. Formally:
 | 
| 18372 |    112 | *}
 | 
| 27362 |    113 | definition
 | 
|  |    114 |   equiv_c :: "com \<Rightarrow> com \<Rightarrow> bool" ("_ \<sim> _") where
 | 
|  |    115 |   "c \<sim> c' = (\<forall>s s'. \<langle>c, s\<rangle> \<longrightarrow>\<^sub>c s' = \<langle>c', s\<rangle> \<longrightarrow>\<^sub>c s')"
 | 
| 12431 |    116 | 
 | 
|  |    117 | text {*
 | 
|  |    118 |   Proof rules telling Isabelle to unfold the definition
 | 
|  |    119 |   if there is something to be proved about equivalent
 | 
| 18372 |    120 |   statements: *}
 | 
| 12431 |    121 | lemma equivI [intro!]:
 | 
|  |    122 |   "(\<And>s s'. \<langle>c, s\<rangle> \<longrightarrow>\<^sub>c s' = \<langle>c', s\<rangle> \<longrightarrow>\<^sub>c s') \<Longrightarrow> c \<sim> c'"
 | 
|  |    123 |   by (unfold equiv_c_def) blast
 | 
|  |    124 | 
 | 
|  |    125 | lemma equivD1:
 | 
|  |    126 |   "c \<sim> c' \<Longrightarrow> \<langle>c, s\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>c', s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
|  |    127 |   by (unfold equiv_c_def) blast
 | 
|  |    128 | 
 | 
|  |    129 | lemma equivD2:
 | 
|  |    130 |   "c \<sim> c' \<Longrightarrow> \<langle>c', s\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>c, s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
|  |    131 |   by (unfold equiv_c_def) blast
 | 
| 1700 |    132 | 
 | 
| 12431 |    133 | text {*
 | 
|  |    134 |   As an example, we show that loop unfolding is an equivalence
 | 
|  |    135 |   transformation on programs:
 | 
|  |    136 | *}
 | 
|  |    137 | lemma unfold_while:
 | 
|  |    138 |   "(\<WHILE> b \<DO> c) \<sim> (\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>)" (is "?w \<sim> ?if")
 | 
|  |    139 | proof -
 | 
|  |    140 |   -- "to show the equivalence, we look at the derivation tree for"
 | 
| 18372 |    141 |   -- "each side and from that construct a derivation tree for the other side"
 | 
| 12431 |    142 |   { fix s s' assume w: "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
|  |    143 |     -- "as a first thing we note that, if @{text b} is @{text False} in state @{text s},"
 | 
|  |    144 |     -- "then both statements do nothing:"
 | 
| 34055 |    145 |     hence "\<not>b s \<Longrightarrow> s = s'" by blast
 | 
|  |    146 |     hence "\<not>b s \<Longrightarrow> \<langle>?if, s\<rangle> \<longrightarrow>\<^sub>c s'" by blast
 | 
| 12431 |    147 |     moreover
 | 
|  |    148 |     -- "on the other hand, if @{text b} is @{text True} in state @{text s},"
 | 
|  |    149 |     -- {* then only the @{text WhileTrue} rule can have been used to derive @{text "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>c s'"} *}
 | 
|  |    150 |     { assume b: "b s"
 | 
|  |    151 |       with w obtain s'' where
 | 
|  |    152 |         "\<langle>c, s\<rangle> \<longrightarrow>\<^sub>c s''" and "\<langle>?w, s''\<rangle> \<longrightarrow>\<^sub>c s'" by (cases set: evalc) auto
 | 
|  |    153 |       -- "now we can build a derivation tree for the @{text \<IF>}"
 | 
|  |    154 |       -- "first, the body of the True-branch:"
 | 
|  |    155 |       hence "\<langle>c; ?w, s\<rangle> \<longrightarrow>\<^sub>c s'" by (rule Semi)
 | 
|  |    156 |       -- "then the whole @{text \<IF>}"
 | 
|  |    157 |       with b have "\<langle>?if, s\<rangle> \<longrightarrow>\<^sub>c s'" by (rule IfTrue)
 | 
|  |    158 |     }
 | 
| 18372 |    159 |     ultimately
 | 
|  |    160 |     -- "both cases together give us what we want:"
 | 
| 12431 |    161 |     have "\<langle>?if, s\<rangle> \<longrightarrow>\<^sub>c s'" by blast
 | 
|  |    162 |   }
 | 
|  |    163 |   moreover
 | 
|  |    164 |   -- "now the other direction:"
 | 
| 19796 |    165 |   { fix s s' assume "if": "\<langle>?if, s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
| 12431 |    166 |     -- "again, if @{text b} is @{text False} in state @{text s}, then the False-branch"
 | 
|  |    167 |     -- "of the @{text \<IF>} is executed, and both statements do nothing:"
 | 
| 34055 |    168 |     hence "\<not>b s \<Longrightarrow> s = s'" by blast
 | 
|  |    169 |     hence "\<not>b s \<Longrightarrow> \<langle>?w, s\<rangle> \<longrightarrow>\<^sub>c s'" by blast
 | 
| 12431 |    170 |     moreover
 | 
|  |    171 |     -- "on the other hand, if @{text b} is @{text True} in state @{text s},"
 | 
|  |    172 |     -- {* then this time only the @{text IfTrue} rule can have be used *}
 | 
|  |    173 |     { assume b: "b s"
 | 
| 19796 |    174 |       with "if" have "\<langle>c; ?w, s\<rangle> \<longrightarrow>\<^sub>c s'" by (cases set: evalc) auto
 | 
| 12431 |    175 |       -- "and for this, only the Semi-rule is applicable:"
 | 
|  |    176 |       then obtain s'' where
 | 
|  |    177 |         "\<langle>c, s\<rangle> \<longrightarrow>\<^sub>c s''" and "\<langle>?w, s''\<rangle> \<longrightarrow>\<^sub>c s'" by (cases set: evalc) auto
 | 
|  |    178 |       -- "with this information, we can build a derivation tree for the @{text \<WHILE>}"
 | 
|  |    179 |       with b
 | 
| 18372 |    180 |       have "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>c s'" by (rule WhileTrue)
 | 
| 12431 |    181 |     }
 | 
| 18372 |    182 |     ultimately
 | 
| 12431 |    183 |     -- "both cases together again give us what we want:"
 | 
|  |    184 |     have "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>c s'" by blast
 | 
|  |    185 |   }
 | 
|  |    186 |   ultimately
 | 
|  |    187 |   show ?thesis by blast
 | 
|  |    188 | qed
 | 
|  |    189 | 
 | 
| 34055 |    190 | text {*
 | 
|  |    191 |    Happily, such lengthy proofs are seldom necessary.  Isabelle can prove many such facts automatically.
 | 
|  |    192 | *}
 | 
|  |    193 | 
 | 
|  |    194 | lemma 
 | 
|  |    195 |   "(\<WHILE> b \<DO> c) \<sim> (\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>)"
 | 
|  |    196 | by blast
 | 
|  |    197 | 
 | 
|  |    198 | lemma triv_if:
 | 
|  |    199 |   "(\<IF> b \<THEN> c \<ELSE> c) \<sim> c"
 | 
|  |    200 | by blast
 | 
|  |    201 | 
 | 
|  |    202 | lemma commute_if:
 | 
|  |    203 |   "(\<IF> b1 \<THEN> (\<IF> b2 \<THEN> c11 \<ELSE> c12) \<ELSE> c2) 
 | 
|  |    204 |    \<sim> 
 | 
|  |    205 |    (\<IF> b2 \<THEN> (\<IF> b1 \<THEN> c11 \<ELSE> c2) \<ELSE> (\<IF> b1 \<THEN> c12 \<ELSE> c2))"
 | 
|  |    206 | by blast
 | 
|  |    207 | 
 | 
|  |    208 | lemma while_equiv:
 | 
|  |    209 |   "\<langle>c0, s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> c \<sim> c' \<Longrightarrow> (c0 = \<WHILE> b \<DO> c) \<Longrightarrow> \<langle>\<WHILE> b \<DO> c', s\<rangle> \<longrightarrow>\<^sub>c u" 
 | 
|  |    210 | by (induct rule: evalc.induct) (auto simp add: equiv_c_def) 
 | 
|  |    211 | 
 | 
|  |    212 | lemma equiv_while:
 | 
|  |    213 |   "c \<sim> c' \<Longrightarrow> (\<WHILE> b \<DO> c) \<sim> (\<WHILE> b \<DO> c')"
 | 
|  |    214 | by (simp add: equiv_c_def) (metis equiv_c_def while_equiv) 
 | 
|  |    215 | 
 | 
|  |    216 | 
 | 
|  |    217 | text {*
 | 
|  |    218 |     Program equivalence is an equivalence relation.
 | 
|  |    219 | *}
 | 
|  |    220 | 
 | 
|  |    221 | lemma equiv_refl:
 | 
|  |    222 |   "c \<sim> c"
 | 
|  |    223 | by blast
 | 
|  |    224 | 
 | 
|  |    225 | lemma equiv_sym:
 | 
|  |    226 |   "c1 \<sim> c2 \<Longrightarrow> c2 \<sim> c1"
 | 
|  |    227 | by (auto simp add: equiv_c_def) 
 | 
|  |    228 | 
 | 
|  |    229 | lemma equiv_trans:
 | 
|  |    230 |   "c1 \<sim> c2 \<Longrightarrow> c2 \<sim> c3 \<Longrightarrow> c1 \<sim> c3"
 | 
|  |    231 | by (auto simp add: equiv_c_def) 
 | 
|  |    232 | 
 | 
|  |    233 | text {*
 | 
|  |    234 |     Program constructions preserve equivalence.
 | 
|  |    235 | *}
 | 
|  |    236 | 
 | 
|  |    237 | lemma equiv_semi:
 | 
|  |    238 |   "c1 \<sim> c1' \<Longrightarrow> c2 \<sim> c2' \<Longrightarrow> (c1; c2) \<sim> (c1'; c2')"
 | 
|  |    239 | by (force simp add: equiv_c_def) 
 | 
|  |    240 | 
 | 
|  |    241 | lemma equiv_if:
 | 
|  |    242 |   "c1 \<sim> c1' \<Longrightarrow> c2 \<sim> c2' \<Longrightarrow> (\<IF> b \<THEN> c1 \<ELSE> c2) \<sim> (\<IF> b \<THEN> c1' \<ELSE> c2')"
 | 
|  |    243 | by (force simp add: equiv_c_def) 
 | 
|  |    244 | 
 | 
|  |    245 | lemma while_never: "\<langle>c, s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> c \<noteq> \<WHILE> (\<lambda>s. True) \<DO> c1"
 | 
|  |    246 | apply (induct rule: evalc.induct)
 | 
|  |    247 | apply auto
 | 
|  |    248 | done
 | 
|  |    249 | 
 | 
|  |    250 | lemma equiv_while_True:
 | 
|  |    251 |   "(\<WHILE> (\<lambda>s. True) \<DO> c1) \<sim> (\<WHILE> (\<lambda>s. True) \<DO> c2)" 
 | 
|  |    252 | by (blast dest: while_never) 
 | 
|  |    253 | 
 | 
| 12431 |    254 | 
 | 
|  |    255 | subsection "Execution is deterministic"
 | 
| 1700 |    256 | 
 | 
| 12431 |    257 | text {*
 | 
| 34055 |    258 | This proof is automatic.
 | 
|  |    259 | *}
 | 
|  |    260 | theorem "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = t"
 | 
|  |    261 | by (induct arbitrary: u rule: evalc.induct) blast+
 | 
|  |    262 | 
 | 
|  |    263 | 
 | 
|  |    264 | text {*
 | 
| 12431 |    265 | The following proof presents all the details:
 | 
|  |    266 | *}
 | 
| 18372 |    267 | theorem com_det:
 | 
|  |    268 |   assumes "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t" and "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
|  |    269 |   shows "u = t"
 | 
|  |    270 |   using prems
 | 
| 20503 |    271 | proof (induct arbitrary: u set: evalc)
 | 
| 18372 |    272 |   fix s u assume "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
| 34055 |    273 |   thus "u = s" by blast
 | 
| 18372 |    274 | next
 | 
|  |    275 |   fix a s x u assume "\<langle>x :== a,s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
| 34055 |    276 |   thus "u = s[x \<mapsto> a s]" by blast
 | 
| 18372 |    277 | next
 | 
|  |    278 |   fix c0 c1 s s1 s2 u
 | 
|  |    279 |   assume IH0: "\<And>u. \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s2"
 | 
|  |    280 |   assume IH1: "\<And>u. \<langle>c1,s2\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s1"
 | 
| 12431 |    281 | 
 | 
| 18372 |    282 |   assume "\<langle>c0;c1, s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
|  |    283 |   then obtain s' where
 | 
| 12431 |    284 |       c0: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s'" and
 | 
| 18372 |    285 |       c1: "\<langle>c1,s'\<rangle> \<longrightarrow>\<^sub>c u"
 | 
|  |    286 |     by auto
 | 
| 12431 |    287 | 
 | 
| 18372 |    288 |   from c0 IH0 have "s'=s2" by blast
 | 
|  |    289 |   with c1 IH1 show "u=s1" by blast
 | 
|  |    290 | next
 | 
|  |    291 |   fix b c0 c1 s s1 u
 | 
|  |    292 |   assume IH: "\<And>u. \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s1"
 | 
| 12431 |    293 | 
 | 
| 18372 |    294 |   assume "b s" and "\<langle>\<IF> b \<THEN> c0 \<ELSE> c1,s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
| 34055 |    295 |   hence "\<langle>c0, s\<rangle> \<longrightarrow>\<^sub>c u" by blast
 | 
| 18372 |    296 |   with IH show "u = s1" by blast
 | 
|  |    297 | next
 | 
|  |    298 |   fix b c0 c1 s s1 u
 | 
|  |    299 |   assume IH: "\<And>u. \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s1"
 | 
| 1700 |    300 | 
 | 
| 18372 |    301 |   assume "\<not>b s" and "\<langle>\<IF> b \<THEN> c0 \<ELSE> c1,s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
| 34055 |    302 |   hence "\<langle>c1, s\<rangle> \<longrightarrow>\<^sub>c u" by blast
 | 
| 18372 |    303 |   with IH show "u = s1" by blast
 | 
|  |    304 | next
 | 
|  |    305 |   fix b c s u
 | 
|  |    306 |   assume "\<not>b s" and "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
| 34055 |    307 |   thus "u = s" by blast
 | 
| 18372 |    308 | next
 | 
|  |    309 |   fix b c s s1 s2 u
 | 
|  |    310 |   assume "IH\<^sub>c": "\<And>u. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s2"
 | 
|  |    311 |   assume "IH\<^sub>w": "\<And>u. \<langle>\<WHILE> b \<DO> c,s2\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s1"
 | 
|  |    312 | 
 | 
|  |    313 |   assume "b s" and "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
|  |    314 |   then obtain s' where
 | 
| 12431 |    315 |       c: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" and
 | 
| 18372 |    316 |       w: "\<langle>\<WHILE> b \<DO> c,s'\<rangle> \<longrightarrow>\<^sub>c u"
 | 
|  |    317 |     by auto
 | 
|  |    318 | 
 | 
|  |    319 |   from c "IH\<^sub>c" have "s' = s2" by blast
 | 
|  |    320 |   with w "IH\<^sub>w" show "u = s1" by blast
 | 
| 12431 |    321 | qed
 | 
|  |    322 | 
 | 
| 1700 |    323 | 
 | 
| 12431 |    324 | text {*
 | 
|  |    325 |   This is the proof as you might present it in a lecture. The remaining
 | 
| 18372 |    326 |   cases are simple enough to be proved automatically:
 | 
|  |    327 | *}
 | 
|  |    328 | theorem
 | 
|  |    329 |   assumes "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t" and "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
|  |    330 |   shows "u = t"
 | 
|  |    331 |   using prems
 | 
| 20503 |    332 | proof (induct arbitrary: u)
 | 
| 18372 |    333 |   -- "the simple @{text \<SKIP>} case for demonstration:"
 | 
|  |    334 |   fix s u assume "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
| 34055 |    335 |   thus "u = s" by blast
 | 
| 18372 |    336 | next
 | 
|  |    337 |   -- "and the only really interesting case, @{text \<WHILE>}:"
 | 
|  |    338 |   fix b c s s1 s2 u
 | 
|  |    339 |   assume "IH\<^sub>c": "\<And>u. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s2"
 | 
|  |    340 |   assume "IH\<^sub>w": "\<And>u. \<langle>\<WHILE> b \<DO> c,s2\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s1"
 | 
|  |    341 | 
 | 
|  |    342 |   assume "b s" and "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c u"
 | 
|  |    343 |   then obtain s' where
 | 
| 12431 |    344 |       c: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" and
 | 
|  |    345 |       w: "\<langle>\<WHILE> b \<DO> c,s'\<rangle> \<longrightarrow>\<^sub>c u"
 | 
| 18372 |    346 |     by auto
 | 
|  |    347 | 
 | 
|  |    348 |   from c "IH\<^sub>c" have "s' = s2" by blast
 | 
|  |    349 |   with w "IH\<^sub>w" show "u = s1" by blast
 | 
| 34055 |    350 | qed blast+ -- "prove the rest automatically"
 | 
| 1700 |    351 | 
 | 
|  |    352 | end
 |