| 
12431
 | 
     1  | 
(*  Title:        HOL/IMP/Natural.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 "Natural Semantics of Commands"
  | 
| 
 | 
     9  | 
  | 
| 
 | 
    10  | 
theory Natural = Com:
  | 
| 
 | 
    11  | 
  | 
| 
 | 
    12  | 
subsection "Execution of commands"
  | 
| 
1700
 | 
    13  | 
  | 
| 
12546
 | 
    14  | 
consts  evalc   :: "(com \<times> state \<times> state) set" 
  | 
| 
 | 
    15  | 
syntax "_evalc" :: "[com,state,state] \<Rightarrow> bool" ("<_,_>/ -c-> _" [0,0,60] 60)
 | 
| 
1700
 | 
    16  | 
  | 
| 
12431
 | 
    17  | 
syntax (xsymbols)
  | 
| 
12546
 | 
    18  | 
  "_evalc" :: "[com,state,state] \<Rightarrow> bool" ("\<langle>_,_\<rangle>/ \<longrightarrow>\<^sub>c _" [0,0,60] 60)
 | 
| 
1700
 | 
    19  | 
  | 
| 
12431
 | 
    20  | 
text {*
 | 
| 
 | 
    21  | 
  We write @{text "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'"} for \emph{Statement @{text c}, started 
 | 
| 
 | 
    22  | 
  in state @{text s}, terminates in state @{text s'}}. Formally,
 | 
| 
 | 
    23  | 
  @{text "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'"} is just another form of saying \emph{the tuple
 | 
| 
 | 
    24  | 
  @{text "(c,s,s')"} is part of the relation @{text evalc}}:
 | 
| 
 | 
    25  | 
*}
  | 
| 
 | 
    26  | 
translations  "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" == "(c,s,s') \<in> evalc"
  | 
| 
1700
 | 
    27  | 
  | 
| 
12431
 | 
    28  | 
constdefs
  | 
| 
 | 
    29  | 
  update :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)" ("_/[_ ::= /_]" [900,0,0] 900)
 | 
| 
 | 
    30  | 
  "update == fun_upd"
  | 
| 
 | 
    31  | 
  | 
| 
 | 
    32  | 
syntax (xsymbols)
  | 
| 
 | 
    33  | 
  update :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)" ("_/[_ \<mapsto> /_]" [900,0,0] 900)
 | 
| 
 | 
    34  | 
  | 
| 
 | 
    35  | 
text {*
 | 
| 
 | 
    36  | 
  The big-step execution relation @{text evalc} is defined inductively:
 | 
| 
 | 
    37  | 
*}
  | 
| 
1789
 | 
    38  | 
inductive evalc
  | 
| 
12431
 | 
    39  | 
  intros 
  | 
| 
 | 
    40  | 
  Skip:    "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c s"
  | 
| 
 | 
    41  | 
  Assign:  "\<langle>x :== a,s\<rangle> \<longrightarrow>\<^sub>c s[x\<mapsto>a s]"
  | 
| 
 | 
    42  | 
  | 
| 
 | 
    43  | 
  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'"
  | 
| 
 | 
    44  | 
  | 
| 
 | 
    45  | 
  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'"
  | 
| 
 | 
    46  | 
  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'"
  | 
| 
 | 
    47  | 
  | 
| 
 | 
    48  | 
  WhileFalse: "\<not>b s \<Longrightarrow> \<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c s"
  | 
| 
 | 
    49  | 
  WhileTrue:  "b s \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'' \<Longrightarrow> \<langle>\<WHILE> b \<DO> c, s''\<rangle> \<longrightarrow>\<^sub>c s'  
  | 
| 
 | 
    50  | 
               \<Longrightarrow> \<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>c s'"
  | 
| 
 | 
    51  | 
  | 
| 
 | 
    52  | 
lemmas evalc.intros [intro] -- "use those rules in automatic proofs"
  | 
| 
 | 
    53  | 
  | 
| 
 | 
    54  | 
text {*
 | 
| 
 | 
    55  | 
The induction principle induced by this definition looks like this:
  | 
| 
 | 
    56  | 
  | 
| 
 | 
    57  | 
@{thm [display] evalc.induct [no_vars]}
 | 
| 
 | 
    58  | 
  | 
| 
 | 
    59  | 
  | 
| 
 | 
    60  | 
(@{text "\<And>"} and @{text "\<Longrightarrow>"} are Isabelle's 
 | 
| 
 | 
    61  | 
  meta symbols for @{text "\<forall>"} and @{text "\<longrightarrow>"})
 | 
| 
 | 
    62  | 
*}
  | 
| 
 | 
    63  | 
  | 
| 
 | 
    64  | 
  | 
| 
 | 
    65  | 
text {*
 | 
| 
 | 
    66  | 
  The rules of @{text evalc} are syntax directed, i.e.~for each
 | 
| 
 | 
    67  | 
  syntactic category there is always only one rule applicable. That
  | 
| 
 | 
    68  | 
  means we can use the rules in both directions. The proofs for this
  | 
| 
 | 
    69  | 
  are all the same: one direction is trivial, the other one is shown
  | 
| 
 | 
    70  | 
  by using the @{text evalc} rules backwards: 
 | 
| 
 | 
    71  | 
*}
  | 
| 
 | 
    72  | 
lemma skip: 
  | 
| 
 | 
    73  | 
  "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c s' = (s' = s)"
  | 
| 
 | 
    74  | 
  by (rule, erule evalc.elims) auto
  | 
| 
 | 
    75  | 
  | 
| 
 | 
    76  | 
lemma assign: 
  | 
| 
 | 
    77  | 
  "\<langle>x :== a,s\<rangle> \<longrightarrow>\<^sub>c s' = (s' = s[x\<mapsto>a s])" 
  | 
| 
 | 
    78  | 
  by (rule, erule evalc.elims) auto
  | 
| 
 | 
    79  | 
  | 
| 
 | 
    80  | 
lemma semi: 
  | 
| 
 | 
    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')"
  | 
| 
 | 
    82  | 
  by (rule, erule evalc.elims) auto
  | 
| 
1700
 | 
    83  | 
  | 
| 
12431
 | 
    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'" 
  | 
| 
 | 
    86  | 
  by (rule, erule evalc.elims) auto
  | 
| 
 | 
    87  | 
  | 
| 
 | 
    88  | 
lemma ifFalse: 
  | 
| 
 | 
    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'"
  | 
| 
 | 
    90  | 
  by (rule, erule evalc.elims) auto
  | 
| 
 | 
    91  | 
  | 
| 
 | 
    92  | 
lemma whileFalse: 
  | 
| 
 | 
    93  | 
  "\<not> b s \<Longrightarrow> \<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c s' = (s' = s)"
  | 
| 
 | 
    94  | 
  by (rule, erule evalc.elims) auto  
  | 
| 
 | 
    95  | 
  | 
| 
 | 
    96  | 
lemma whileTrue: 
  | 
| 
 | 
    97  | 
  "b s \<Longrightarrow> 
  | 
| 
 | 
    98  | 
  \<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>c s' = 
  | 
| 
 | 
    99  | 
  (\<exists>s''. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'' \<and> \<langle>\<WHILE> b \<DO> c, s''\<rangle> \<longrightarrow>\<^sub>c s')"
  | 
| 
 | 
   100  | 
  by (rule, erule evalc.elims) auto
  | 
| 
 | 
   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:
 | 
| 
 | 
   112  | 
*} 
  | 
| 
 | 
   113  | 
constdefs
  | 
| 
 | 
   114  | 
  equiv_c :: "com \<Rightarrow> com \<Rightarrow> bool" ("_ \<sim> _")
 | 
| 
 | 
   115  | 
  "c \<sim> c' \<equiv> \<forall>s s'. \<langle>c, s\<rangle> \<longrightarrow>\<^sub>c s' = \<langle>c', s\<rangle> \<longrightarrow>\<^sub>c s'"
  | 
| 
 | 
   116  | 
  | 
| 
 | 
   117  | 
text {*
 | 
| 
 | 
   118  | 
  Proof rules telling Isabelle to unfold the definition
  | 
| 
 | 
   119  | 
  if there is something to be proved about equivalent
  | 
| 
 | 
   120  | 
  statements: *} 
  | 
| 
 | 
   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"
  | 
| 
 | 
   141  | 
  -- "each side and from that construct a derivation tree for the other side"  
  | 
| 
 | 
   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:"
  | 
| 
 | 
   145  | 
    hence "\<not>b s \<Longrightarrow> s = s'" by simp
  | 
| 
 | 
   146  | 
    hence "\<not>b s \<Longrightarrow> \<langle>?if, s\<rangle> \<longrightarrow>\<^sub>c s'" by simp
  | 
| 
 | 
   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  | 
    }
  | 
| 
 | 
   159  | 
    ultimately 
  | 
| 
 | 
   160  | 
    -- "both cases together give us what we want:"      
  | 
| 
 | 
   161  | 
    have "\<langle>?if, s\<rangle> \<longrightarrow>\<^sub>c s'" by blast
  | 
| 
 | 
   162  | 
  }
  | 
| 
 | 
   163  | 
  moreover
  | 
| 
 | 
   164  | 
  -- "now the other direction:"
  | 
| 
 | 
   165  | 
  { fix s s' assume if: "\<langle>?if, s\<rangle> \<longrightarrow>\<^sub>c s'"
 | 
| 
 | 
   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:"
 | 
| 
 | 
   168  | 
    hence "\<not>b s \<Longrightarrow> s = s'" by simp
  | 
| 
 | 
   169  | 
    hence "\<not>b s \<Longrightarrow> \<langle>?w, s\<rangle> \<longrightarrow>\<^sub>c s'" by simp
  | 
| 
 | 
   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"
 | 
| 
 | 
   174  | 
      with if have "\<langle>c; ?w, s\<rangle> \<longrightarrow>\<^sub>c s'" by (cases set: evalc) auto
  | 
| 
 | 
   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
  | 
| 
 | 
   180  | 
      have "\<langle>?w, s\<rangle> \<longrightarrow>\<^sub>c s'" by (rule WhileTrue) 
  | 
| 
 | 
   181  | 
    }
  | 
| 
 | 
   182  | 
    ultimately 
  | 
| 
 | 
   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  | 
  | 
| 
 | 
   190  | 
  | 
| 
 | 
   191  | 
subsection "Execution is deterministic"
  | 
| 
1700
 | 
   192  | 
  | 
| 
12431
 | 
   193  | 
text {*
 | 
| 
 | 
   194  | 
The following proof presents all the details:
  | 
| 
 | 
   195  | 
*}
  | 
| 
 | 
   196  | 
theorem com_det: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t \<and> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u \<longrightarrow> u=t"
  | 
| 
 | 
   197  | 
proof clarify -- "transform the goal into canonical form"
  | 
| 
 | 
   198  | 
  assume "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"
  | 
| 
 | 
   199  | 
  thus "\<And>u. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u=t" 
  | 
| 
 | 
   200  | 
  proof (induct set: evalc)
  | 
| 
 | 
   201  | 
    fix s u assume "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c u"
  | 
| 
 | 
   202  | 
    thus "u = s" by simp
  | 
| 
 | 
   203  | 
  next
  | 
| 
 | 
   204  | 
    fix a s x u assume "\<langle>x :== a,s\<rangle> \<longrightarrow>\<^sub>c u"
  | 
| 
 | 
   205  | 
    thus "u = s[x \<mapsto> a s]" by simp      
  | 
| 
 | 
   206  | 
  next
  | 
| 
 | 
   207  | 
    fix c0 c1 s s1 s2 u
  | 
| 
 | 
   208  | 
    assume IH0: "\<And>u. \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s2"
  | 
| 
 | 
   209  | 
    assume IH1: "\<And>u. \<langle>c1,s2\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s1"
  | 
| 
 | 
   210  | 
  | 
| 
 | 
   211  | 
    assume "\<langle>c0;c1, s\<rangle> \<longrightarrow>\<^sub>c u"
  | 
| 
 | 
   212  | 
    then obtain s' where
  | 
| 
 | 
   213  | 
      c0: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s'" and
  | 
| 
 | 
   214  | 
      c1: "\<langle>c1,s'\<rangle> \<longrightarrow>\<^sub>c u" 
  | 
| 
 | 
   215  | 
      by auto
  | 
| 
 | 
   216  | 
  | 
| 
 | 
   217  | 
    from c0 IH0 have "s'=s2" by blast
  | 
| 
 | 
   218  | 
    with c1 IH1 show "u=s1" by blast
  | 
| 
 | 
   219  | 
  next
  | 
| 
 | 
   220  | 
    fix b c0 c1 s s1 u
  | 
| 
 | 
   221  | 
    assume IH: "\<And>u. \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s1" 
  | 
| 
 | 
   222  | 
  | 
| 
 | 
   223  | 
    assume "b s" and "\<langle>\<IF> b \<THEN> c0 \<ELSE> c1,s\<rangle> \<longrightarrow>\<^sub>c u"
  | 
| 
 | 
   224  | 
    hence "\<langle>c0, s\<rangle> \<longrightarrow>\<^sub>c u" by simp
  | 
| 
 | 
   225  | 
    with IH show "u = s1" by blast
  | 
| 
 | 
   226  | 
  next
  | 
| 
 | 
   227  | 
    fix b c0 c1 s s1 u
  | 
| 
 | 
   228  | 
    assume IH: "\<And>u. \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s1" 
  | 
| 
1700
 | 
   229  | 
  | 
| 
12431
 | 
   230  | 
    assume "\<not>b s" and "\<langle>\<IF> b \<THEN> c0 \<ELSE> c1,s\<rangle> \<longrightarrow>\<^sub>c u"
  | 
| 
 | 
   231  | 
    hence "\<langle>c1, s\<rangle> \<longrightarrow>\<^sub>c u" by simp
  | 
| 
 | 
   232  | 
    with IH show "u = s1" by blast
  | 
| 
 | 
   233  | 
  next
  | 
| 
 | 
   234  | 
    fix b c s u
  | 
| 
 | 
   235  | 
    assume "\<not>b s" and "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c u"
  | 
| 
 | 
   236  | 
    thus "u = s" by simp
  | 
| 
 | 
   237  | 
  next
  | 
| 
 | 
   238  | 
    fix b c s s1 s2 u
  | 
| 
 | 
   239  | 
    assume "IH\<^sub>c": "\<And>u. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s2"
  | 
| 
 | 
   240  | 
    assume "IH\<^sub>w": "\<And>u. \<langle>\<WHILE> b \<DO> c,s2\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s1"
  | 
| 
 | 
   241  | 
    
  | 
| 
 | 
   242  | 
    assume "b s" and "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c u"
  | 
| 
 | 
   243  | 
    then obtain s' where
  | 
| 
 | 
   244  | 
      c: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" and
  | 
| 
 | 
   245  | 
      w: "\<langle>\<WHILE> b \<DO> c,s'\<rangle> \<longrightarrow>\<^sub>c u" 
  | 
| 
 | 
   246  | 
      by auto
  | 
| 
 | 
   247  | 
    
  | 
| 
 | 
   248  | 
    from c "IH\<^sub>c" have "s' = s2" by blast
  | 
| 
 | 
   249  | 
    with w "IH\<^sub>w" show "u = s1" by blast
  | 
| 
 | 
   250  | 
  qed
  | 
| 
 | 
   251  | 
qed
  | 
| 
 | 
   252  | 
  | 
| 
1700
 | 
   253  | 
  | 
| 
12431
 | 
   254  | 
text {*
 | 
| 
 | 
   255  | 
  This is the proof as you might present it in a lecture. The remaining
  | 
| 
 | 
   256  | 
  cases are simple enough to be proved automatically: 
  | 
| 
 | 
   257  | 
*} 
  | 
| 
 | 
   258  | 
theorem "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t \<and> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u \<longrightarrow> u=t" 
  | 
| 
 | 
   259  | 
proof clarify
  | 
| 
 | 
   260  | 
  assume "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"
  | 
| 
 | 
   261  | 
  thus "\<And>u. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u=t" 
  | 
| 
 | 
   262  | 
  proof (induct set: evalc)
  | 
| 
 | 
   263  | 
    -- "the simple @{text \<SKIP>} case for demonstration:"
 | 
| 
 | 
   264  | 
    fix s u assume "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c u"
  | 
| 
 | 
   265  | 
    thus "u = s" by simp
  | 
| 
 | 
   266  | 
  next
  | 
| 
 | 
   267  | 
    -- "and the only really interesting case, @{text \<WHILE>}:"
 | 
| 
 | 
   268  | 
    fix b c s s1 s2 u
  | 
| 
 | 
   269  | 
    assume "IH\<^sub>c": "\<And>u. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s2"
  | 
| 
 | 
   270  | 
    assume "IH\<^sub>w": "\<And>u. \<langle>\<WHILE> b \<DO> c,s2\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> u = s1"
  | 
| 
 | 
   271  | 
    
  | 
| 
 | 
   272  | 
    assume "b s" and "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c u"
  | 
| 
 | 
   273  | 
    then obtain s' where
  | 
| 
 | 
   274  | 
      c: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" and
  | 
| 
 | 
   275  | 
      w: "\<langle>\<WHILE> b \<DO> c,s'\<rangle> \<longrightarrow>\<^sub>c u"
  | 
| 
 | 
   276  | 
      by auto
  | 
| 
 | 
   277  | 
    
  | 
| 
 | 
   278  | 
    from c "IH\<^sub>c" have "s' = s2" by blast
  | 
| 
 | 
   279  | 
    with w "IH\<^sub>w" show "u = s1" by blast
  | 
| 
 | 
   280  | 
  qed (best dest: evalc_cases [THEN iffD1])+ -- "prove the rest automatically"
  | 
| 
 | 
   281  | 
qed
  | 
| 
1700
 | 
   282  | 
  | 
| 
 | 
   283  | 
end
  |