author | wenzelm |
Sun, 02 Mar 2014 22:43:20 +0100 | |
changeset 55849 | 3a2ad5ccc1c8 |
parent 55836 | 8093590e49e4 (diff) |
parent 55848 | 1bfe72d14630 (current diff) |
child 55850 | 7f229b0212fe |
--- a/src/HOL/IMP/Small_Step.thy Sun Mar 02 22:39:34 2014 +0100 +++ b/src/HOL/IMP/Small_Step.thy Sun Mar 02 22:43:20 2014 +0100 @@ -159,7 +159,7 @@ lemma small_to_big: "cs \<rightarrow>* (SKIP,t) \<Longrightarrow> cs \<Rightarrow> t" -apply (induction cs "(SKIP,t)" arbitrary: t rule: star.induct) +apply (induction cs "(SKIP,t)" rule: star.induct) apply (auto intro: small1_big_continue) done