src/HOL/IMP/Collecting_list.thy
changeset 48765 fb1ed5230abc
parent 48764 4fe0920d5049
child 48766 553ad5f99968
equal deleted inserted replaced
48764:4fe0920d5049 48765:fb1ed5230abc
     1 theory Collecting_list
       
     2 imports ACom
       
     3 begin
       
     4 
       
     5 subsection "Executable Collecting Semantics on lists"
       
     6 
       
     7 fun step :: "state list \<Rightarrow> state list acom \<Rightarrow> state list acom" where
       
     8 "step S (SKIP {P}) = (SKIP {S})" |
       
     9 "step S (x ::= e {P}) =
       
    10   (x ::= e {[s(x := aval e s). s \<leftarrow> S]})" |
       
    11 "step S (c1; c2) = step S c1; step (post c1) c2" |
       
    12 "step S (IF b THEN c1 ELSE c2 {P}) =
       
    13    IF b THEN step [s \<leftarrow> S. bval b s] c1 ELSE step [s\<leftarrow>S. \<not> bval b s] c2
       
    14    {post c1 @ post c2}" |
       
    15 "step S ({Inv} WHILE b DO c {P}) =
       
    16   {S @ post c} WHILE b DO (step [s\<leftarrow>Inv. bval b s] c) {[s\<leftarrow>Inv. \<not> bval b s]}"
       
    17 
       
    18 
       
    19 text{* Examples: *}
       
    20 
       
    21 definition "c = WHILE Less (V ''x'') (N 3)
       
    22                 DO ''x'' ::= Plus (V ''x'') (N 2)"
       
    23 definition "c0 = anno [] c"
       
    24 
       
    25 definition "show_acom xs = map_acom (map (show_state xs))"
       
    26 
       
    27 
       
    28 text{* Collecting semantics: *}
       
    29 
       
    30 value "show_acom [''x''] (((step [<>]) ^^ 6) c0)"
       
    31 
       
    32 
       
    33 text{* Small step semantics: *}
       
    34 
       
    35 value "show_acom [''x''] (((step []) ^^ 5) (step [<>] c0))"
       
    36 
       
    37 end