src/HOL/IMP/Hoare.thy
changeset 1447 bc2c0acbbf29
parent 1374 5e407f2a3323
child 1476 608483c2122a
equal deleted inserted replaced
1446:a8387e934fa7 1447:bc2c0acbbf29
     1 (*  Title: 	HOL/IMP/Hoare.thy
     1 (*  Title: 	HOL/IMP/Hoare.thy
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author: 	Tobias Nipkow
     3     Author: 	Tobias Nipkow
     4     Copyright   1995 TUM
     4     Copyright   1995 TUM
     5 
     5 
     6 Semantic embedding of Hoare logic
     6 Inductive definition of Hoare logic
     7 *)
     7 *)
     8 
     8 
     9 Hoare = Denotation +
     9 Hoare = Denotation +
       
    10 
       
    11 types assn = state => bool
       
    12 
    10 consts
    13 consts
       
    14   hoare :: "(assn * com * assn) set"
    11   spec :: [state=>bool,com,state=>bool] => bool
    15   spec :: [state=>bool,com,state=>bool] => bool
    12 (* syntax "@spec" :: [bool,com,bool] => bool *)
       
    13           ("{{(1_)}}/ (_)/ {{(1_)}}" 10)
       
    14 defs
    16 defs
    15   spec_def "spec P c Q == !s t. (s,t) : C(c) --> P s --> Q t"
    17   spec_def "spec P c Q == !s t. (s,t) : C(c) --> P s --> Q t"
    16 end
       
    17 (* Pretty-printing of assertions.
       
    18    Not very helpful as long as programs are not pretty-printed.
       
    19 ML
       
    20 
    18 
    21 local open Syntax
    19 syntax "@hoare" :: [bool,com,bool] => bool ("{{(1_)}}/ (_)/ {{(1_)}}" 10)
       
    20 translations "{{P}}c{{Q}}" == "(P,c,Q) : hoare"
    22 
    21 
    23 fun is_loc a = let val ch = hd(explode a)
    22 inductive "hoare"
    24                in ord "A" <= ord ch andalso ord ch <= ord "Z" end;
    23 intrs
    25 
    24   hoare_skip "{{P}}skip{{P}}"
    26 fun tr(s$t,i) = tr(s,i)$tr(t,i)
    25   hoare_ass  "{{%s.P(s[A a s/x])}} x:=a {{P}}"
    27   | tr(Abs(x,T,u),i) = Abs(x,T,tr(u,i+1))
    26   hoare_semi "[| {{P}}c{{Q}}; {{Q}}d{{R}} |] ==> {{P}} c;d {{R}}"
    28   | tr(t as Free(a,T),i) = if is_loc a then Bound(i) $ free(a) else t
    27   hoare_if   "[| {{%s. P s & B b s}}c{{Q}}; {{%s. P s & ~B b s}}d{{Q}} |] ==>
    29   | tr(t,_) = t;
    28               {{P}} ifc b then c else d {{Q}}"
    30 
    29   hoare_while "[| {{%s. P s & B b s}} c {{P}} |] ==>
    31 fun cond_tr(p) = Abs("",dummyT,tr(p,0))
    30 	       {{P}} while b do c {{%s. P s & ~B b s}}"
    32 
    31   hoare_conseq "[| !s. P' s --> P s; {{P}}c{{Q}}; !s. Q s --> Q' s |] ==>
    33 fun spec_tr[p,c,q] = const"spec" $ cond_tr p $ c $ cond_tr q;
    32 		{{P'}}c{{Q'}}"
    34 
       
    35 fun tr'(t as (Bound j $ (u as Free(a,_))),i) = if i=j then u else t
       
    36   | tr'(s$t,i) = tr'(s,i)$tr'(t,i)
       
    37   | tr'(Abs(x,T,u),i) = Abs(x,T,tr'(u,i+1))
       
    38   | tr'(t,_) = t;
       
    39 
       
    40 fun spec_tr'[Abs(_,_,p),c,Abs(_,_,q)] =
       
    41   const"@spec" $ tr'(p,0) $ c $ tr'(q,0);
       
    42 
       
    43 in
       
    44 
       
    45 val parse_translation = [("@spec", spec_tr)];
       
    46 val print_translation = [("spec", spec_tr')];
       
    47 
    33 
    48 end
    34 end
    49 *)