diff -r f4c64597fb02 -r 01fc1fc61384 src/HOL/IMP/Compiler.thy --- a/src/HOL/IMP/Compiler.thy Wed Oct 23 16:10:42 2002 +0200 +++ b/src/HOL/IMP/Compiler.thy Thu Oct 24 07:23:46 2002 +0200 @@ -11,11 +11,11 @@ consts compile :: "com \ instr list" primrec "compile \ = []" -"compile (x:==a) = [ASIN x a]" +"compile (x:==a) = [SET x a]" "compile (c1;c2) = compile c1 @ compile c2" "compile (\ b \ c1 \ c2) = [JMPF b (length(compile c1) + 1)] @ compile c1 @ - [JMPF (%x. False) (length(compile c2))] @ compile c2" + [JMPF (\x. False) (length(compile c2))] @ compile c2" "compile (\ b \ c) = [JMPF b (length(compile c) + 1)] @ compile c @ [JMPB (length(compile c)+1)]" @@ -68,12 +68,12 @@ constdefs forws :: "instr \ nat set" "forws instr == case instr of - ASIN x a \ {0} | + SET x a \ {0} | JMPF b n \ {0,n} | JMPB n \ {}" backws :: "instr \ nat set" "backws instr == case instr of - ASIN x a \ {} | + SET x a \ {} | JMPF b n \ {} | JMPB n \ {n}" @@ -263,7 +263,7 @@ proof cases assume b: "b s" then obtain m where m: "n = Suc m" - and 1: "\?C @ [?j2],[?j1],s\ -m\ \[],rev ?W,t\" + and "\?C @ [?j2],[?j1],s\ -m\ \[],rev ?W,t\" using H by fastsimp then obtain r n1 n2 where n1: "\?C,[],s\ -n1\ \[],rev ?C,r\" and n2: "\[?j2],rev ?C @ [?j1],r\ -n2\ \[],rev ?W,t\"