diff -r 22e5804b135c -r 839e37b75d9a recursive.thy --- a/recursive.thy Sat Jan 12 01:05:01 2013 +0000 +++ b/recursive.thy Sat Jan 12 01:33:20 2013 +0000 @@ -4782,7 +4782,7 @@ done lemma [elim]: "\k < length ap; ap ! k = Inc n; - (a, b) \ set (abacus.tshift (abacus.tshift tinc_b (2 * n)) + (a, b) \ set (tshift (tshift tinc_b (2 * n)) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "b \ start_of (layout_of ap) (Suc k)") @@ -4794,7 +4794,7 @@ layout_of.simps length_of.simps startof_not0) done -lemma findnth_le[elim]: "(a, b) \ set (abacus.tshift (findnth n) (start_of (layout_of ap) k - Suc 0)) +lemma findnth_le[elim]: "(a, b) \ set (tshift (findnth n) (start_of (layout_of ap) k - Suc 0)) \ b \ Suc (start_of (layout_of ap) k + 2 * n)" apply(induct n, simp add: findnth.simps tshift.simps) apply(simp add: findnth.simps tshift_append, auto) @@ -4803,7 +4803,7 @@ lemma [elim]: "\k < length ap; ap ! k = Inc n; (a, b) \ - set (abacus.tshift (findnth n) (start_of (layout_of ap) k - Suc 0))\ + set (tshift (findnth n) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "b \ start_of (layout_of ap) (Suc k)") apply(subgoal_tac "start_of (layout_of ap) (Suc k) \ start_of (layout_of ap) (length ap) ") @@ -4831,7 +4831,7 @@ lemma [elim]: "\k < length ap; ap ! k = Dec n e; - (a, b) \ set (abacus.tshift (findnth n) (start_of (layout_of ap) k - Suc 0))\ + (a, b) \ set (tshift (findnth n) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "b \ start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) k + 2*n + 1 \ start_of (layout_of ap) (Suc k) \ @@ -4842,7 +4842,7 @@ done thm length_of.simps -lemma [elim]: "\k < length ap; ap ! k = Dec n e; (a, b) \ set (abacus.tshift (abacus.tshift tdec_b (2 * n)) +lemma [elim]: "\k < length ap; ap ! k = Dec n e; (a, b) \ set (tshift (tshift tdec_b (2 * n)) (start_of (layout_of ap) k - Suc 0))\ \ b \ start_of (layout_of ap) (length ap)" apply(subgoal_tac "2*n + start_of (layout_of ap) k + 16 \ start_of (layout_of ap) (length ap) \ start_of (layout_of ap) k > 0")