src/HOL/IMP/Natural.thy
author nipkow
Tue, 05 Jan 1999 17:27:59 +0100
changeset 6059 aa00e235ea27
parent 5102 8c782c25a11e
child 9241 f961c1fdff50
permissions -rw-r--r--
In Main: moved Bin to the left to preserve the solver in its simpset.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1700
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/IMP/Natural.thy
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow & Robert Sandner, TUM
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     4
    Copyright   1996 TUM
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     5
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     6
Natural semantics of commands
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     7
*)
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     8
5102
8c782c25a11e Adapted to new inductive definition package.
berghofe
parents: 4897
diff changeset
     9
Natural = Com + Inductive +
1700
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    10
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    11
(** Execution of commands **)
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    12
consts  evalc    :: "(com*state*state)set"
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    13
        "@evalc" :: [com,state,state] => bool ("<_,_>/ -c-> _" [0,0,50] 50)
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    14
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    15
translations  "<ce,sig> -c-> s" == "(ce,sig,s) : evalc"
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    16
4897
be11be0b6ea1 Changed [/] to [:=] and removed actual definition.
nipkow
parents: 1789
diff changeset
    17
consts
be11be0b6ea1 Changed [/] to [:=] and removed actual definition.
nipkow
parents: 1789
diff changeset
    18
  update  :: "('a => 'b) => 'a => 'b => ('a => 'b)"
be11be0b6ea1 Changed [/] to [:=] and removed actual definition.
nipkow
parents: 1789
diff changeset
    19
           ("_/[_/:=/_]" [900,0,0] 900)
be11be0b6ea1 Changed [/] to [:=] and removed actual definition.
nipkow
parents: 1789
diff changeset
    20
(* update is NOT defined, only declared!
be11be0b6ea1 Changed [/] to [:=] and removed actual definition.
nipkow
parents: 1789
diff changeset
    21
   Thus the whole theory is independent of its meaning!
be11be0b6ea1 Changed [/] to [:=] and removed actual definition.
nipkow
parents: 1789
diff changeset
    22
   If theory Update is included, proofs break.
be11be0b6ea1 Changed [/] to [:=] and removed actual definition.
nipkow
parents: 1789
diff changeset
    23
*)
1700
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    24
1789
aade046ec6d5 Quotes now optional around inductive set
paulson
parents: 1700
diff changeset
    25
inductive evalc
1700
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    26
  intrs
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    27
    Skip    "<SKIP,s> -c-> s"
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    28
4897
be11be0b6ea1 Changed [/] to [:=] and removed actual definition.
nipkow
parents: 1789
diff changeset
    29
    Assign  "<x := a,s> -c-> s[x:=a(s)]"
1700
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    30
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    31
    Semi    "[| <c0,s> -c-> s2; <c1,s2> -c-> s1 |] 
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    32
            ==> <c0 ; c1, s> -c-> s1"
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    33
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    34
    IfTrue "[| b s; <c0,s> -c-> s1 |] 
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    35
            ==> <IF b THEN c0 ELSE c1, s> -c-> s1"
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    36
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    37
    IfFalse "[| ~b s; <c1,s> -c-> s1 |] 
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    38
             ==> <IF b THEN c0 ELSE c1, s> -c-> s1"
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    39
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    40
    WhileFalse "~b s ==> <WHILE b DO c,s> -c-> s"
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    41
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    42
    WhileTrue  "[| b s;  <c,s> -c-> s2;  <WHILE b DO c, s2> -c-> s1 |] 
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    43
               ==> <WHILE b DO c, s> -c-> s1"
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    44
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    45
end