diff -r b505be6f029a -r 686fa0a0696e src/HOL/IMP/Def_Ass_Small.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/IMP/Def_Ass_Small.thy Mon Jun 06 16:29:38 2011 +0200 @@ -0,0 +1,26 @@ +(* Author: Tobias Nipkow *) + +theory Def_Ass_Small imports Star Com Def_Ass_Exp +begin + +subsection "Initialization-Sensitive Small Step Semantics" + +inductive + small_step :: "(com \ state) \ (com \ state) \ bool" (infix "\" 55) +where +Assign: "aval a s = Some i \ (x ::= a, s) \ (SKIP, s(x := Some i))" | + +Semi1: "(SKIP;c,s) \ (c,s)" | +Semi2: "(c\<^isub>1,s) \ (c\<^isub>1',s') \ (c\<^isub>1;c\<^isub>2,s) \ (c\<^isub>1';c\<^isub>2,s')" | + +IfTrue: "bval b s = Some True \ (IF b THEN c\<^isub>1 ELSE c\<^isub>2,s) \ (c\<^isub>1,s)" | +IfFalse: "bval b s = Some False \ (IF b THEN c\<^isub>1 ELSE c\<^isub>2,s) \ (c\<^isub>2,s)" | + +While: "(WHILE b DO c,s) \ (IF b THEN c; WHILE b DO c ELSE SKIP,s)" + +lemmas small_step_induct = small_step.induct[split_format(complete)] + +abbreviation small_steps :: "com * state \ com * state \ bool" (infix "\*" 55) +where "x \* y == star small_step x y" + +end