src/HOL/IMP/Big_Step.thy
changeset 67071 a462583f0a37
parent 67019 7a3724078363
child 67406 23307fd33906
--- a/src/HOL/IMP/Big_Step.thy	Fri Nov 10 22:05:30 2017 +0100
+++ b/src/HOL/IMP/Big_Step.thy	Tue Nov 14 13:12:13 2017 +0100
@@ -206,13 +206,11 @@
       thus ?thesis using `\<not>bval b s` by blast
     next
       case IfTrue
-      with `(?iw, s) \<Rightarrow> t` have "(c;; ?w, s) \<Rightarrow> t" by auto
       -- "and for this, only the Seq-rule is applicable:"
-      then obtain s' where
+      from `(c;; ?w, s) \<Rightarrow> t` obtain s' where
         "(c, s) \<Rightarrow> s'" and "(?w, s') \<Rightarrow> t" by auto
-      -- "with this information, we can build a derivation tree for the @{text WHILE}"
-      with `bval b s`
-      show ?thesis by (rule WhileTrue)
+      -- "with this information, we can build a derivation tree for @{text WHILE}"
+      with `bval b s` show ?thesis by (rule WhileTrue)
     qed
   qed
   ultimately