# HG changeset patch # User wenzelm # Date 1190827178 -7200 # Node ID dd9ea6b72eb9e6b0cf15f0194384134984c9591d # Parent fcf13a91cda2a428925d509421534f9b89bbb827 adapted variable order for inductive cases (determined by read_specification *before* expanding abbreviations); diff -r fcf13a91cda2 -r dd9ea6b72eb9 src/HOL/Bali/AxSound.thy --- a/src/HOL/Bali/AxSound.thy Wed Sep 26 19:18:01 2007 +0200 +++ b/src/HOL/Bali/AxSound.thy Wed Sep 26 19:19:38 2007 +0200 @@ -2017,7 +2017,7 @@ \\ (P'\=False\=\) v s' Z" (is "PROP ?Hyp n t s v s'") proof (induct) - case (Loop s0' e' n' b s1' c' s2' l' s3' Y' T E) + case (Loop s0' e' b n' s1' c' s2' l' s3' Y' T E) note while = `(\l'\ While(e') c'\\<^sub>s::term) = \l\ While(e) c\\<^sub>s` hence eqs: "l'=l" "e'=e" "c'=c" by simp_all note valid_A = `\t\A. G\n'\t` diff -r fcf13a91cda2 -r dd9ea6b72eb9 src/HOL/Bali/Evaln.thy --- a/src/HOL/Bali/Evaln.thy Wed Sep 26 19:18:01 2007 +0200 +++ b/src/HOL/Bali/Evaln.thy Wed Sep 26 19:19:38 2007 +0200 @@ -410,7 +410,7 @@ shows "G\s0 \t\\ (v,s1)" using evaln proof (induct) - case (Loop s0 e n b s1 c s2 l s3) + case (Loop s0 e b n s1 c s2 l s3) note `G\Norm s0 \e-\b\ s1` moreover have "if the_Bool b