| 1700 |      1 | (*  Title:      HOL/IMP/Natural.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Tobias Nipkow & Robert Sandner, TUM
 | 
|  |      4 |     Copyright   1996 TUM
 | 
|  |      5 | 
 | 
|  |      6 | Natural semantics of commands
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
| 5102 |      9 | Natural = Com + Inductive +
 | 
| 1700 |     10 | 
 | 
|  |     11 | (** Execution of commands **)
 | 
|  |     12 | consts  evalc    :: "(com*state*state)set"
 | 
|  |     13 |         "@evalc" :: [com,state,state] => bool ("<_,_>/ -c-> _" [0,0,50] 50)
 | 
|  |     14 | 
 | 
|  |     15 | translations  "<ce,sig> -c-> s" == "(ce,sig,s) : evalc"
 | 
|  |     16 | 
 | 
| 4897 |     17 | consts
 | 
|  |     18 |   update  :: "('a => 'b) => 'a => 'b => ('a => 'b)"
 | 
| 9241 |     19 |            ("_/[_/::=/_]" [900,0,0] 900)
 | 
| 4897 |     20 | (* update is NOT defined, only declared!
 | 
|  |     21 |    Thus the whole theory is independent of its meaning!
 | 
| 9241 |     22 |    If the definition (update == fun_upd) is included, proofs break.
 | 
| 4897 |     23 | *)
 | 
| 1700 |     24 | 
 | 
| 1789 |     25 | inductive evalc
 | 
| 1700 |     26 |   intrs
 | 
|  |     27 |     Skip    "<SKIP,s> -c-> s"
 | 
|  |     28 | 
 | 
| 9241 |     29 |     Assign  "<x :== a,s> -c-> s[x::=a s]"
 | 
| 1700 |     30 | 
 | 
|  |     31 |     Semi    "[| <c0,s> -c-> s2; <c1,s2> -c-> s1 |] 
 | 
|  |     32 |             ==> <c0 ; c1, s> -c-> s1"
 | 
|  |     33 | 
 | 
|  |     34 |     IfTrue "[| b s; <c0,s> -c-> s1 |] 
 | 
|  |     35 |             ==> <IF b THEN c0 ELSE c1, s> -c-> s1"
 | 
|  |     36 | 
 | 
|  |     37 |     IfFalse "[| ~b s; <c1,s> -c-> s1 |] 
 | 
|  |     38 |              ==> <IF b THEN c0 ELSE c1, s> -c-> s1"
 | 
|  |     39 | 
 | 
|  |     40 |     WhileFalse "~b s ==> <WHILE b DO c,s> -c-> s"
 | 
|  |     41 | 
 | 
|  |     42 |     WhileTrue  "[| b s;  <c,s> -c-> s2;  <WHILE b DO c, s2> -c-> s1 |] 
 | 
|  |     43 |                ==> <WHILE b DO c, s> -c-> s1"
 | 
|  |     44 | 
 | 
|  |     45 | end
 |