diff -r a4dc62a46ee4 -r 132634d24019 IMP/Com.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/IMP/Com.thy Fri Jan 19 12:49:36 1996 +0100 @@ -0,0 +1,112 @@ +(* Title: HOL/IMP/Com.thy + ID: $Id$ + Author: Heiko Loetzbeyer & Robert Sandner, TUM + Copyright 1994 TUM + +Arithmetic expressions, Boolean expressions, Commands + +And their Operational semantics +*) + +Com = Arith + + +(** Arithmetic expressions **) +types loc + state = "loc => nat" + n2n = "nat => nat" + n2n2n = "nat => nat => nat" + +arities loc :: term + +datatype + aexp = N (nat) + | X (loc) + | Op1 (n2n, aexp) + | Op2 (n2n2n, aexp, aexp) + +(** Evaluation of arithmetic expressions **) +consts evala :: "(aexp*state*nat)set" + "@evala" :: "[aexp,state,nat] => bool" ("<_,_>/ -a-> _" [0,0,50] 50) +translations + " -a-> n" == " : evala" +inductive "evala" + intrs + N " -a-> n" + X " -a-> s(x)" + Op1 " -a-> n ==> -a-> f(n)" + Op2 "[| -a-> n0; -a-> n1 |] + ==> -a-> f(n0,n1)" + +types n2n2b = "[nat,nat] => bool" + +(** Boolean expressions **) + +datatype + bexp = true + | false + | ROp (n2n2b, aexp, aexp) + | noti (bexp) + | andi (bexp,bexp) (infixl 60) + | ori (bexp,bexp) (infixl 60) + +(** Evaluation of boolean expressions **) +consts evalb :: "(bexp*state*bool)set" + "@evalb" :: "[bexp,state,bool] => bool" ("<_,_>/ -b-> _" [0,0,50] 50) + +translations + " -b-> b" == " : evalb" + +inductive "evalb" + intrs (*avoid clash with ML constructors true, false*) + tru " -b-> True" + fls " -b-> False" + ROp "[| -a-> n0; -a-> n1 |] + ==> -b-> f(n0,n1)" + noti " -b-> w ==> -b-> (~w)" + andi "[| -b-> w0; -b-> w1 |] + ==> -b-> (w0 & w1)" + ori "[| -b-> w0; -b-> w1 |] + ==> -b-> (w0 | w1)" + +(** Commands **) + +datatype + com = skip + | ":=" (loc,aexp) (infixl 60) + | semic (com,com) ("_; _" [60, 60] 10) + | whileC (bexp,com) ("while _ do _" 60) + | ifC (bexp, com, com) ("ifc _ then _ else _" 60) + +(** Execution of commands **) +consts evalc :: "(com*state*state)set" + "@evalc" :: "[com,state,state] => bool" ("<_,_>/ -c-> _" [0,0,50] 50) + "assign" :: "[state,nat,loc] => state" ("_[_'/_]" [95,0,0] 95) + +translations + " -c-> s" == " : evalc" + +rules + assign_def "s[m/x] == (%y. if(y=x,m,s(y)))" + +inductive "evalc" + intrs + skip " -c-> s" + + assign " -a-> m ==> -c-> s[m/x]" + + semi "[| -c-> s2; -c-> s1 |] + ==> -c-> s1" + + ifcTrue "[| -b-> True; -c-> s1 |] + ==> -c-> s1" + + ifcFalse "[| -b-> False; -c-> s1 |] + ==> -c-> s1" + + whileFalse " -b-> False ==> -c-> s" + + whileTrue "[| -b-> True; -c-> s2; + -c-> s1 |] + ==> -c-> s1 " + +end